From 846b74275511bd89c2f3abe19245133050d2199c Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Thu, 12 Jan 2017 20:11:01 +0100 Subject: [constrexpr] Make patterns use Loc.located for location information This is first of a series of patches, converting `constrexpr` pattern data type from ad-hoc location handling to `Loc.located`. Along Coq, we can find two different coding styles for handling objects with location information: one style uses `'a Loc.located`, whereas other data structures directly embed `Loc.t` in their constructors. Handling all located objects uniformly would be very convenient, and would allow optimizing certain cases, in particular making located smarter when there is no location information, as it is the case for all terms coming from the kernel. `git grep 'Loc.t \*'` gives an overview of the remaining work to do. We've also added an experimental API for `located` to the `Loc` module, `Loc.tag` should be used to add locations objects, making it explicit in the code when a "located" object is created. --- intf/constrexpr.mli | 43 +++++++++++++++++++++++-------------------- 1 file changed, 23 insertions(+), 20 deletions(-) (limited to 'intf') diff --git a/intf/constrexpr.mli b/intf/constrexpr.mli index 49bafadc8..c1de0ce24 100644 --- a/intf/constrexpr.mli +++ b/intf/constrexpr.mli @@ -36,31 +36,33 @@ type prim_token = | Numeral of Bigint.bigint (** representation of integer literals that appear in Coq scripts. *) | String of string -type raw_cases_pattern_expr = - | RCPatAlias of Loc.t * raw_cases_pattern_expr * Id.t - | RCPatCstr of Loc.t * Globnames.global_reference +type raw_cases_pattern_expr_r = + | RCPatAlias of raw_cases_pattern_expr * Id.t + | RCPatCstr of Globnames.global_reference * raw_cases_pattern_expr list * raw_cases_pattern_expr list (** [CPatCstr (_, c, l1, l2)] represents ((@c l1) l2) *) - | RCPatAtom of Loc.t * Id.t option - | RCPatOr of Loc.t * raw_cases_pattern_expr list + | RCPatAtom of Id.t option + | RCPatOr of raw_cases_pattern_expr list +and raw_cases_pattern_expr = raw_cases_pattern_expr_r Loc.located type instance_expr = Misctypes.glob_level list -type cases_pattern_expr = - | CPatAlias of Loc.t * cases_pattern_expr * Id.t - | CPatCstr of Loc.t * reference +type cases_pattern_expr_r = + | CPatAlias of cases_pattern_expr * Id.t + | CPatCstr of reference * cases_pattern_expr list option * cases_pattern_expr list (** [CPatCstr (_, c, Some l1, l2)] represents (@c l1) l2 *) - | CPatAtom of Loc.t * reference option - | CPatOr of Loc.t * cases_pattern_expr list - | CPatNotation of Loc.t * notation * cases_pattern_notation_substitution + | CPatAtom of reference option + | CPatOr of cases_pattern_expr list + | CPatNotation of notation * cases_pattern_notation_substitution * cases_pattern_expr list (** CPatNotation (_, n, l1 ,l2) represents (notation n applied with substitution l1) applied to arguments l2 *) - | CPatPrim of Loc.t * prim_token - | CPatRecord of Loc.t * (reference * cases_pattern_expr) list - | CPatDelimiters of Loc.t * string * cases_pattern_expr - | CPatCast of Loc.t * cases_pattern_expr * constr_expr + | CPatPrim of prim_token + | CPatRecord of (reference * cases_pattern_expr) list + | CPatDelimiters of string * cases_pattern_expr + | CPatCast of cases_pattern_expr * constr_expr +and cases_pattern_expr = cases_pattern_expr_r located and cases_pattern_notation_substitution = cases_pattern_expr list * (** for constr subterms *) @@ -104,7 +106,7 @@ and case_expr = constr_expr (* expression that is being matched * cases_pattern_expr option (* in-clause *) and branch_expr = - Loc.t * cases_pattern_expr list located list * constr_expr + (cases_pattern_expr list located list * constr_expr) Loc.located and binder_expr = Name.t located list * binder_kind * constr_expr @@ -144,7 +146,8 @@ type with_declaration_ast = | CWith_Module of Id.t list located * qualid located | CWith_Definition of Id.t list located * constr_expr -type module_ast = - | CMident of qualid located - | CMapply of Loc.t * module_ast * module_ast - | CMwith of Loc.t * module_ast * with_declaration_ast +type module_ast_r = + | CMident of qualid + | CMapply of module_ast * module_ast + | CMwith of module_ast * with_declaration_ast +and module_ast = module_ast_r located -- cgit v1.2.3 From 6d9e008ffd81bbe927e3442fb0c37269ed25b21f Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Sat, 14 Jan 2017 01:27:40 +0100 Subject: [location] Use Loc.located for constr_expr. This is the second patch, which is a bit more invasive. We reasoning is similar to the previous patch. Code is not as clean as it could as we would need to convert `glob_constr` to located too, then a few parts could just map the location. --- ide/texmacspp.ml | 58 ++--- interp/constrexpr_ops.ml | 126 +++++------ interp/constrexpr_ops.mli | 4 +- interp/constrextern.ml | 78 +++---- interp/constrintern.ml | 110 ++++----- interp/implicit_quantifiers.ml | 22 +- interp/topconstr.ml | 94 ++++---- intf/constrexpr.mli | 60 ++--- lib/loc.ml | 5 + lib/loc.mli | 5 + parsing/egramcoq.ml | 18 +- parsing/g_constr.ml4 | 111 ++++----- parsing/g_proofs.ml4 | 2 +- parsing/g_vernac.ml4 | 24 +- plugins/decl_mode/g_decl_mode.ml4 | 387 ++++++++++++++++++++++++++++++++ plugins/funind/glob_term_to_relation.ml | 31 ++- plugins/funind/indfun.ml | 105 +++++---- plugins/funind/merge.ml | 4 +- plugins/ltac/g_ltac.ml4 | 8 +- plugins/ltac/g_tactic.ml4 | 12 +- plugins/ltac/pptactic.ml | 4 +- plugins/ltac/rewrite.ml | 12 +- plugins/ltac/tacintern.ml | 10 +- plugins/ltac/tacinterp.ml | 2 +- plugins/ssrmatching/ssrmatching.ml4 | 32 +-- printing/ppconstr.ml | 117 +++++----- printing/ppvernac.ml | 4 +- vernac/classes.ml | 8 +- vernac/command.ml | 19 +- vernac/metasyntax.ml | 4 +- vernac/record.ml | 7 +- 31 files changed, 933 insertions(+), 550 deletions(-) create mode 100644 plugins/decl_mode/g_decl_mode.ml4 (limited to 'intf') 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 *) +(* + 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 -- cgit v1.2.3 From ad3aab9415b98a247a6cbce05752632c3c42391c Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Mon, 16 Jan 2017 13:02:55 +0100 Subject: [location] Move Glob_term.cases_pattern to located. We continue the uniformization pass. No big news here, trying to be minimally invasive. --- interp/constrextern.ml | 26 ++++++------ interp/constrintern.ml | 32 +++++++-------- interp/notation.ml | 2 +- interp/notation_ops.ml | 67 ++++++++++++++++--------------- intf/glob_term.mli | 7 ++-- lib/loc.ml | 3 +- lib/loc.mli | 1 + plugins/funind/glob_term_to_relation.ml | 15 +++---- plugins/funind/glob_termops.ml | 59 +++++++++++++-------------- plugins/funind/recdef.ml | 6 +-- pretyping/cases.ml | 71 +++++++++++++++++---------------- pretyping/coercion.ml | 5 ++- pretyping/detyping.ml | 22 +++++----- pretyping/glob_ops.ml | 40 +++++++++---------- pretyping/patternops.ml | 10 ++--- 15 files changed, 188 insertions(+), 178 deletions(-) (limited to 'intf') diff --git a/interp/constrextern.ml b/interp/constrextern.ml index 255de8500..b3059f5d0 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -182,7 +182,7 @@ let add_patt_for_params ind l = let add_cpatt_for_params ind l = if !Flags.in_debugger then l else - Util.List.addn (Inductiveops.inductive_nparamdecls ind) (PatVar (Loc.ghost,Anonymous)) l + Util.List.addn (Inductiveops.inductive_nparamdecls ind) (Loc.tag @@ PatVar Anonymous) l let drop_implicits_in_patt cst nb_expl args = let impl_st = (implicits_of_global cst) in @@ -291,7 +291,7 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = (* pboutill: There are letins in pat which is incompatible with notations and not explicit application. *) match pat with - | PatCstr(loc,cstrsp,args,na) + | loc, PatCstr(cstrsp,args,na) when !Flags.in_debugger||Inductiveops.constructor_has_local_defs cstrsp -> let c = extern_reference loc Id.Set.empty (ConstructRef cstrsp) in let args = List.map (extern_cases_pattern_in_scope scopes vars) args in @@ -312,9 +312,9 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = (uninterp_cases_pattern_notations pat) with No_match -> match pat with - | PatVar (loc,Name id) -> Loc.tag ~loc @@ CPatAtom (Some (Ident (loc,id))) - | PatVar (loc,Anonymous) -> Loc.tag ~loc @@ CPatAtom None - | PatCstr(loc,cstrsp,args,na) -> + | loc, PatVar (Name id) -> Loc.tag ~loc @@ CPatAtom (Some (Ident (loc,id))) + | loc, PatVar (Anonymous) -> Loc.tag ~loc @@ CPatAtom None + | loc, PatCstr(cstrsp,args,na) -> let args = List.map (extern_cases_pattern_in_scope scopes vars) args in let p = try @@ -391,20 +391,20 @@ and apply_notation_to_pattern loc gr ((subst,substlist),(nb_to_drop,more_args)) in assert (List.is_empty substlist); mkPat loc qid (List.rev_append l1 l2') -and extern_notation_pattern (tmp_scope,scopes as allscopes) vars t = function +and extern_notation_pattern (tmp_scope,scopes as allscopes) vars (loc, t) = function | [] -> raise No_match | (keyrule,pat,n as _rule)::rules -> try if List.mem keyrule !print_non_active_notations then raise No_match; match t with - | PatCstr (loc,cstr,_,na) -> + | PatCstr (cstr,_,na) -> let p = apply_notation_to_pattern loc (ConstructRef cstr) - (match_notation_constr_cases_pattern t pat) allscopes vars keyrule in + (match_notation_constr_cases_pattern (loc, t) pat) allscopes vars keyrule in insert_pat_alias loc p na - | PatVar (loc,Anonymous) -> Loc.tag ~loc @@ CPatAtom None - | PatVar (loc,Name id) -> Loc.tag ~loc @@ CPatAtom (Some (Ident (loc,id))) + | PatVar Anonymous -> Loc.tag ~loc @@ CPatAtom None + | PatVar (Name id) -> Loc.tag ~loc @@ CPatAtom (Some (Ident (loc,id))) with - No_match -> extern_notation_pattern allscopes vars t rules + No_match -> extern_notation_pattern allscopes vars (loc, t) rules let rec extern_notation_ind_pattern allscopes vars ind args = function | [] -> raise No_match @@ -750,7 +750,7 @@ let rec extern inctx scopes vars r = (sub_extern false scopes vars tm, na', Option.map (fun (loc,ind,nal) -> - let args = List.map (fun x -> PatVar (Loc.ghost, x)) nal in + let args = List.map (fun x -> Loc.tag @@ PatVar x) nal in let fullargs = add_cpatt_for_params ind args in extern_ind_pattern_in_scope scopes vars ind fullargs ) x)) @@ -1008,7 +1008,7 @@ let extern_closed_glob ?lax goal_concl_style env sigma t = let any_any_branch = (* | _ => _ *) - (loc,[],[PatVar (loc,Anonymous)],GHole (loc,Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None)) + (loc,[],[Loc.tag ~loc @@ PatVar Anonymous],GHole (loc,Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None)) let rec glob_of_pat env sigma = function | PRef ref -> GRef (loc,ref,None) diff --git a/interp/constrintern.ml b/interp/constrintern.ml index 4af89e1ef..4960d7332 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -591,14 +591,14 @@ let rec subordinate_letins letins = function letins,[] let terms_of_binders bl = - let rec term_of_pat = function - | 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 rec term_of_pat pt = Loc.map_with_loc (fun ~loc -> function + | PatVar (Name id) -> CRef (Ident (loc,id), None) + | PatVar (Anonymous) -> error "Cannot turn \"_\" into a term." + | PatCstr (c,l,_) -> let r = Qualid (loc,qualid_of_path (path_of_global (ConstructRef c))) in let hole = Loc.tag ~loc @@ CHole (None,Misctypes.IntroAnonymous,None) in let params = List.make (Inductiveops.inductive_nparams (fst c)) hole in - Loc.tag ~loc @@ CAppExpl ((None,r,None),params @ List.map term_of_pat l) in + CAppExpl ((None,r,None),params @ List.map term_of_pat l)) pt in let rec extract_variables = function | GLocalAssum (loc,Name id,_,_)::l -> (Loc.tag ~loc @@ CRef (Ident (loc,id), None)) :: extract_variables l | GLocalDef (loc,Name id,_,_,_)::l -> extract_variables l @@ -1015,8 +1015,8 @@ let chop_params_pattern loc ind args with_letin = else Inductiveops.inductive_nparams ind in assert (nparams <= List.length args); let params,args = List.chop nparams args in - List.iter (function PatVar(_,Anonymous) -> () - | PatVar (loc',_) | PatCstr(loc',_,_,_) -> error_parameter_not_implicit ~loc:loc') params; + List.iter (function _, PatVar Anonymous -> () + | loc', PatVar _ | loc', PatCstr(_,_,_) -> error_parameter_not_implicit ~loc:loc') params; args let find_constructor loc add_params ref = @@ -1036,7 +1036,7 @@ let find_constructor loc add_params ref = then Inductiveops.inductive_nparamdecls ind else Inductiveops.inductive_nparams ind in - List.make nb ([], [(Id.Map.empty, PatVar(Loc.ghost,Anonymous))]) + List.make nb ([], [(Id.Map.empty, Loc.tag @@ PatVar Anonymous)]) | None -> [] let find_pattern_variable = function @@ -1358,7 +1358,7 @@ let rec intern_pat genv aliases pat = let idslpl2 = List.map (intern_pat genv empty_alias) pl2 in let (ids',pll) = product_of_cases_patterns aliases.alias_ids (idslpl1@idslpl2) in let pl' = List.map (fun (asubst,pl) -> - (asubst, PatCstr (loc,c,chop_params_pattern loc (fst c) pl with_letin,alias_of aliases))) pll in + (asubst, Loc.tag ~loc @@ PatCstr (c,chop_params_pattern loc (fst c) pl with_letin,alias_of aliases))) pll in ids',pl' in match pat with | loc, RCPatAlias (p, id) -> @@ -1378,10 +1378,10 @@ let rec intern_pat genv aliases pat = intern_cstr_with_all_args loc c with_letin idslpl1 (expl_pl@pl2) | loc, RCPatAtom (Some id) -> let aliases = merge_aliases aliases id in - (aliases.alias_ids,[aliases.alias_map, PatVar (loc, alias_of aliases)]) + (aliases.alias_ids,[aliases.alias_map, Loc.tag ~loc @@ PatVar (alias_of aliases)]) | loc, RCPatAtom (None) -> let { alias_ids = ids; alias_map = asubst; } = aliases in - (ids, [asubst, PatVar (loc, alias_of aliases)]) + (ids, [asubst, Loc.tag ~loc @@ PatVar (alias_of aliases)]) | loc, RCPatOr pl -> assert (not (List.is_empty pl)); let pl' = List.map (intern_pat genv aliases) pl in @@ -1678,14 +1678,14 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let stripped_match_from_in = let rec aux = function | [] -> [] - | (_,PatVar _) :: q -> aux q + | (_, (_loc, PatVar _)) :: q -> aux q | l -> l in aux match_from_in in let rtnpo = match stripped_match_from_in with | [] -> Option.map (intern_type env') rtnpo (* Only PatVar in "in" clauses *) | l -> (* Build a return predicate by expansion of the patterns of the "in" clause *) - let thevars,thepats = List.split l in + let thevars, thepats = List.split l in let sub_rtn = (* Some (GSort (Loc.ghost,GType None)) *) None in let sub_tms = List.map (fun id -> GVar (Loc.ghost,id),(Name id,None)) thevars (* "match v1,..,vn" *) in let main_sub_eqn = @@ -1695,7 +1695,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = rtnpo) (* "=> P" if there were a return predicate P, and "=> _" otherwise *) in let catch_all_sub_eqn = if List.for_all (irrefutable globalenv) thepats then [] else - [Loc.ghost,[],List.make (List.length thepats) (PatVar(Loc.ghost,Anonymous)), (* "|_,..,_" *) + [Loc.ghost,[],List.make (List.length thepats) (Loc.tag @@ PatVar Anonymous), (* "|_,..,_" *) GHole(Loc.ghost,Evar_kinds.ImpossibleCase,Misctypes.IntroAnonymous,None)] (* "=> _" *) in Some (GCases(Loc.ghost,Term.RegularStyle,sub_rtn,sub_tms,main_sub_eqn::catch_all_sub_eqn)) in @@ -1817,14 +1817,14 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let rec canonize_args case_rel_ctxt arg_pats forbidden_names match_acc var_acc = let add_name l = function | _,Anonymous -> l - | loc,(Name y as x) -> (y,PatVar(loc,x)) :: l in + | loc,(Name y as x) -> (y, Loc.tag ~loc @@ PatVar x) :: l in match case_rel_ctxt,arg_pats with (* LetIn in the rel_context *) | LocalDef _ :: t, l when not with_letin -> canonize_args t l forbidden_names match_acc ((Loc.ghost,Anonymous)::var_acc) | [],[] -> (add_name match_acc na, var_acc) - | _::t,PatVar (loc,x)::tt -> + | _::t, (loc, PatVar x)::tt -> canonize_args t tt forbidden_names (add_name match_acc (loc,x)) ((loc,x)::var_acc) | (LocalAssum (cano_name,ty) | LocalDef (cano_name,_,ty)) :: t, c::tt -> diff --git a/interp/notation.ml b/interp/notation.ml index 04711808b..aef089299 100644 --- a/interp/notation.ml +++ b/interp/notation.ml @@ -273,7 +273,7 @@ let glob_constr_keys = function | _ -> [Oth] let cases_pattern_key = function - | PatCstr (_,ref,_,_) -> RefKey (canonical_gr (ConstructRef ref)) + | _, PatCstr (ref,_,_) -> RefKey (canonical_gr (ConstructRef ref)) | _ -> Oth let notation_constr_key = function (* Rem: NApp(NRef ref,[]) stands for @ref *) diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml index 8b4fadb5a..29f42d0e9 100644 --- a/interp/notation_ops.ml +++ b/interp/notation_ops.ml @@ -117,13 +117,14 @@ let name_to_ident = function let to_id g e id = let e,na = g e (Name id) in e,name_to_ident na -let rec cases_pattern_fold_map loc g e = function - | PatVar (_,na) -> - let e',na' = g e na in e', PatVar (loc,na') - | PatCstr (_,cstr,patl,na) -> +let rec cases_pattern_fold_map loc g e = Loc.with_unloc (function + | PatVar na -> + let e',na' = g e na in e', Loc.tag ~loc @@ PatVar na' + | PatCstr (cstr,patl,na) -> let e',na' = g e na in let e',patl' = List.fold_map (cases_pattern_fold_map loc g) e patl in - e', PatCstr (loc,cstr,patl',na') + e', Loc.tag ~loc @@ PatCstr (cstr,patl',na') + ) let subst_binder_type_vars l = function | Evar_kinds.BinderType (Name id) -> @@ -450,14 +451,15 @@ let notation_constr_of_constr avoiding t = } in notation_constr_of_glob_constr nenv t -let rec subst_pat subst pat = +let rec subst_pat subst (loc, pat) = match pat with - | PatVar _ -> pat - | PatCstr (loc,((kn,i),j),cpl,n) -> + | PatVar _ -> (loc, pat) + | PatCstr (((kn,i),j),cpl,n) -> let kn' = subst_mind subst kn and cpl' = List.smartmap (subst_pat subst) cpl in + Loc.tag ~loc @@ if kn' == kn && cpl' == cpl then pat else - PatCstr (loc,((kn',i),j),cpl',n) + PatCstr (((kn',i),j),cpl',n) let rec subst_notation_constr subst bound raw = match raw with @@ -662,11 +664,11 @@ let add_bindinglist_env (terms,onlybinders,termlists,binderlists) x bl = (terms,onlybinders,termlists,(x,bl)::binderlists) let rec pat_binder_of_term = function - | GVar (loc, id) -> PatVar (loc, Name id) + | GVar (loc, id) -> Loc.tag ~loc @@ PatVar (Name id) | GApp (loc, GRef (_,ConstructRef cstr,_), l) -> let nparams = Inductiveops.inductive_nparams (fst cstr) in let _,l = List.chop nparams l in - PatCstr (loc, cstr, List.map pat_binder_of_term l, Anonymous) + Loc.tag ~loc @@ PatCstr (cstr, List.map pat_binder_of_term l, Anonymous) | _ -> raise No_match let bind_term_env alp (terms,onlybinders,termlists,binderlists as sigma) var v = @@ -738,16 +740,17 @@ let bind_binding_env alp (terms,onlybinders,termlists,binderlists as sigma) var else (fst alp,(id1,id2)::snd alp),sigma with Not_found -> alp, add_binding_env alp sigma var v -let rec map_cases_pattern_name_left f = function - | PatVar (loc,na) -> PatVar (loc,f na) - | PatCstr (loc,c,l,na) -> PatCstr (loc,c,List.map_left (map_cases_pattern_name_left f) l,f na) +let rec map_cases_pattern_name_left f = Loc.map (function + | PatVar na -> PatVar (f na) + | PatCstr (c,l,na) -> PatCstr (c,List.map_left (map_cases_pattern_name_left f) l,f na) + ) let rec fold_cases_pattern_eq f x p p' = match p, p' with - | PatVar (loc,na), PatVar (_,na') -> let x,na = f x na na' in x, PatVar (loc,na) - | PatCstr (loc,c,l,na), PatCstr (_,c',l',na') when eq_constructor c c' -> + | (loc, PatVar na), (_, PatVar na') -> let x,na = f x na na' in x, Loc.tag ~loc @@ PatVar na + | (loc, PatCstr (c,l,na)), (_, PatCstr (c',l',na')) when eq_constructor c c' -> let x,l = fold_cases_pattern_list_eq f x l l' in let x,na = f x na na' in - x, PatCstr (loc,c,l,na) + x, Loc.tag ~loc @@ PatCstr (c,l,na) | _ -> failwith "Not equal" and fold_cases_pattern_list_eq f x pl pl' = match pl, pl' with @@ -758,9 +761,9 @@ and fold_cases_pattern_list_eq f x pl pl' = match pl, pl' with x, p :: pl | _ -> assert false -let rec cases_pattern_eq p1 p2 = match p1, p2 with -| PatVar (_, na1), PatVar (_, na2) -> Name.equal na1 na2 -| PatCstr (_, c1, pl1, na1), PatCstr (_, c2, pl2, na2) -> +let rec cases_pattern_eq (_,p1) (_,p2) = match p1, p2 with +| PatVar na1, PatVar na2 -> Name.equal na1 na2 +| PatCstr (c1, pl1, na1), PatCstr (c2, pl2, na2) -> eq_constructor c1 c2 && List.equal cases_pattern_eq pl1 pl2 && Name.equal na1 na2 | _ -> false @@ -878,10 +881,10 @@ let match_names metas (alp,sigma) na1 na2 = match (na1,na2) with | (Anonymous,Anonymous) -> alp,sigma | _ -> raise No_match -let rec match_cases_pattern_binders metas acc pat1 pat2 = - match (pat1,pat2) with - | PatVar (_,na1), PatVar (_,na2) -> match_names metas acc na1 na2 - | PatCstr (_,c1,patl1,na1), PatCstr (_,c2,patl2,na2) +let rec match_cases_pattern_binders metas acc (_, pat1) (_, pat2) = + match pat1, pat2 with + | PatVar na1, PatVar na2 -> match_names metas acc na1 na2 + | PatCstr (c1,patl1,na1), PatCstr (c2,patl2,na2) when eq_constructor c1 c2 && Int.equal (List.length patl1) (List.length patl2) -> List.fold_left2 (match_cases_pattern_binders metas) (match_names metas acc na1 na2) patl1 patl2 @@ -1173,7 +1176,7 @@ let match_notation_constr u c (metas,pat) = let add_patterns_for_params ind l = let mib,_ = Global.lookup_inductive ind in let nparams = mib.Declarations.mind_nparams in - Util.List.addn nparams (PatVar (Loc.ghost,Anonymous)) l + Util.List.addn nparams (Loc.tag @@ PatVar Anonymous) l let bind_env_cases_pattern (terms,x,termlists,y as sigma) var v = try @@ -1197,13 +1200,13 @@ let match_cases_pattern_list match_fun metas sigma rest x y iter termin lassoc = let l,(terms,onlybinders,termlists,binderlists as sigma) = aux sigma [] rest in (terms,onlybinders,(x,if lassoc then l else List.rev l)::termlists, binderlists) -let rec match_cases_pattern metas (terms,(),termlists,() as sigma) a1 a2 = - match (a1,a2) with - | r1, NVar id2 when Id.List.mem_assoc id2 metas -> (bind_env_cases_pattern sigma id2 r1),(0,[]) - | PatVar (_,Anonymous), NHole _ -> sigma,(0,[]) - | PatCstr (loc,(ind,_ as r1),largs,_), NRef (ConstructRef r2) when eq_constructor r1 r2 -> +let rec match_cases_pattern metas (terms,(),termlists,() as sigma) (loc, a1) a2 = + match a1, a2 with + | r1, NVar id2 when Id.List.mem_assoc id2 metas -> (bind_env_cases_pattern sigma id2 (loc, r1)),(0,[]) + | PatVar Anonymous, NHole _ -> sigma,(0,[]) + | PatCstr ((ind,_ as r1),largs,_), NRef (ConstructRef r2) when eq_constructor r1 r2 -> sigma,(0,add_patterns_for_params (fst r1) largs) - | PatCstr (loc,(ind,_ as r1),args1,_), NApp (NRef (ConstructRef r2),l2) + | PatCstr ((ind,_ as r1),args1,_), NApp (NRef (ConstructRef r2),l2) when eq_constructor r1 r2 -> let l1 = add_patterns_for_params (fst r1) args1 in let le2 = List.length l2 in @@ -1215,7 +1218,7 @@ let rec match_cases_pattern metas (terms,(),termlists,() as sigma) a1 a2 = (List.fold_left2 (match_cases_pattern_no_more_args metas) sigma l1' l2),(le2,more_args) | r1, NList (x,y,iter,termin,lassoc) -> (match_cases_pattern_list (match_cases_pattern_no_more_args) - metas (terms,(),termlists,()) r1 x y iter termin lassoc),(0,[]) + metas (terms,(),termlists,()) (loc, r1) x y iter termin lassoc),(0,[]) | _ -> raise No_match and match_cases_pattern_no_more_args metas sigma a1 a2 = diff --git a/intf/glob_term.mli b/intf/glob_term.mli index ced5a8b44..5e771245c 100644 --- a/intf/glob_term.mli +++ b/intf/glob_term.mli @@ -24,10 +24,11 @@ type existential_name = Id.t (** The kind of patterns that occurs in "match ... with ... end" locs here refers to the ident's location, not whole pat *) -type cases_pattern = - | PatVar of Loc.t * Name.t - | PatCstr of Loc.t * constructor * cases_pattern list * Name.t +type cases_pattern_r = + | PatVar of Name.t + | PatCstr of constructor * cases_pattern list * Name.t (** [PatCstr(p,C,l,x)] = "|'C' 'l' as 'x'" *) +and cases_pattern = cases_pattern_r Loc.located (** Representation of an internalized (or in other words globalized) term. *) type glob_constr = diff --git a/lib/loc.ml b/lib/loc.ml index 8ae8fe25d..8d7432ff4 100644 --- a/lib/loc.ml +++ b/lib/loc.ml @@ -63,7 +63,8 @@ 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 with_loc f (loc, x) = f ~loc x +let with_unloc f (_,x) = f x let map f (l,x) = (l, f x) let map_with_loc f (loc, x) = (loc, f ~loc x) diff --git a/lib/loc.mli b/lib/loc.mli index 7fc8efaa8..3f484bc4c 100644 --- a/lib/loc.mli +++ b/lib/loc.mli @@ -62,6 +62,7 @@ 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 with_unloc : ('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 diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml index 4b942c989..85d465a4b 100644 --- a/plugins/funind/glob_term_to_relation.ml +++ b/plugins/funind/glob_term_to_relation.ml @@ -345,12 +345,12 @@ let raw_push_named (na,raw_value,raw_typ) env = let add_pat_variables pat typ env : Environ.env = - let rec add_pat_variables env pat typ : Environ.env = + let rec add_pat_variables env (loc, pat) typ : Environ.env = observe (str "new rel env := " ++ Printer.pr_rel_context_of env (Evd.from_env env)); match pat with - | PatVar(_,na) -> Environ.push_rel (RelDecl.LocalAssum (na,typ)) env - | PatCstr(_,c,patl,na) -> + | PatVar na -> Environ.push_rel (RelDecl.LocalAssum (na,typ)) env + | PatCstr(c,patl,na) -> let Inductiveops.IndType(indf,indargs) = try Inductiveops.find_rectype env (Evd.from_env env) (EConstr.of_constr typ) with Not_found -> assert false @@ -398,11 +398,11 @@ let add_pat_variables pat typ env : Environ.env = -let rec pattern_to_term_and_type env typ = function - | PatVar(loc,Anonymous) -> assert false - | PatVar(loc,Name id) -> +let rec pattern_to_term_and_type env typ = Loc.with_unloc (function + | PatVar Anonymous -> assert false + | PatVar (Name id) -> mkGVar id - | PatCstr(loc,constr,patternl,_) -> + | PatCstr(constr,patternl,_) -> let cst_narg = Inductiveops.constructor_nallargs_env (Global.env ()) @@ -430,6 +430,7 @@ let rec pattern_to_term_and_type env typ = function mkGApp(mkGRef(ConstructRef constr), implicit_args@patl_as_term ) + ) (* [build_entry_lc funnames avoid rt] construct the list (in fact a build_entry_return) of constructors corresponding to [rt] when replacing calls to [funnames] by calls to the diff --git a/plugins/funind/glob_termops.ml b/plugins/funind/glob_termops.ml index 99f50437b..51ca8c471 100644 --- a/plugins/funind/glob_termops.ml +++ b/plugins/funind/glob_termops.ml @@ -194,19 +194,19 @@ let change_vars = -let rec alpha_pat excluded pat = +let rec alpha_pat excluded (loc, pat) = match pat with - | PatVar(loc,Anonymous) -> + | PatVar Anonymous -> let new_id = Indfun_common.fresh_id excluded "_x" in - PatVar(loc,Name new_id),(new_id::excluded),Id.Map.empty - | PatVar(loc,Name id) -> + (Loc.tag ~loc @@ PatVar(Name new_id)),(new_id::excluded),Id.Map.empty + | PatVar(Name id) -> if Id.List.mem id excluded then let new_id = Namegen.next_ident_away id excluded in - PatVar(loc,Name new_id),(new_id::excluded), + (Loc.tag ~loc @@ PatVar(Name new_id)),(new_id::excluded), (Id.Map.add id new_id Id.Map.empty) - else pat,excluded,Id.Map.empty - | PatCstr(loc,constr,patl,na) -> + else (Loc.tag ~loc pat),excluded,Id.Map.empty + | PatCstr(constr,patl,na) -> let new_na,new_excluded,map = match na with | Name id when Id.List.mem id excluded -> @@ -223,7 +223,7 @@ let rec alpha_pat excluded pat = ([],new_excluded,map) patl in - PatCstr(loc,constr,List.rev new_patl,new_na),new_excluded,new_map + (Loc.tag ~loc @@ PatCstr(constr,List.rev new_patl,new_na)),new_excluded,new_map let alpha_patl excluded patl = let patl,new_excluded,map = @@ -241,12 +241,12 @@ let alpha_patl excluded patl = let raw_get_pattern_id pat acc = - let rec get_pattern_id pat = + let rec get_pattern_id (loc, pat) = match pat with - | PatVar(loc,Anonymous) -> assert false - | PatVar(loc,Name id) -> + | PatVar(Anonymous) -> assert false + | PatVar(Name id) -> [id] - | PatCstr(loc,constr,patternl,_) -> + | PatCstr(constr,patternl,_) -> List.fold_right (fun pat idl -> let idl' = get_pattern_id pat in @@ -425,11 +425,11 @@ let is_free_in id = -let rec pattern_to_term = function - | PatVar(loc,Anonymous) -> assert false - | PatVar(loc,Name id) -> +let rec pattern_to_term pt = Loc.with_unloc (function + | PatVar Anonymous -> assert false + | PatVar(Name id) -> mkGVar id - | PatCstr(loc,constr,patternl,_) -> + | PatCstr(constr,patternl,_) -> let cst_narg = Inductiveops.constructor_nallargs_env (Global.env ()) @@ -448,7 +448,7 @@ let rec pattern_to_term = function mkGApp(mkGRef(Globnames.ConstructRef constr), implicit_args@patl_as_term ) - + ) pt let replace_var_by_term x_id term = @@ -533,8 +533,8 @@ let rec are_unifiable_aux = function | [] -> () | eq::eqs -> match eq with - | PatVar _,_ | _,PatVar _ -> are_unifiable_aux eqs - | PatCstr(_,constructor1,cpl1,_),PatCstr(_,constructor2,cpl2,_) -> + | (_,PatVar _),_ | _,(_,PatVar _) -> are_unifiable_aux eqs + | (_,PatCstr(constructor1,cpl1,_)),(_,PatCstr(constructor2,cpl2,_)) -> if not (eq_constructor constructor2 constructor1) then raise NotUnifiable else @@ -555,8 +555,8 @@ let rec eq_cases_pattern_aux = function | [] -> () | eq::eqs -> match eq with - | PatVar _,PatVar _ -> eq_cases_pattern_aux eqs - | PatCstr(_,constructor1,cpl1,_),PatCstr(_,constructor2,cpl2,_) -> + | (_,PatVar _),(_,PatVar _) -> eq_cases_pattern_aux eqs + | (_,PatCstr(constructor1,cpl1,_)),(_,PatCstr(constructor2,cpl2,_)) -> if not (eq_constructor constructor2 constructor1) then raise NotUnifiable else @@ -576,10 +576,11 @@ let eq_cases_pattern pat1 pat2 = let ids_of_pat = - let rec ids_of_pat ids = function - | PatVar(_,Anonymous) -> ids - | PatVar(_,Name id) -> Id.Set.add id ids - | PatCstr(_,_,patl,_) -> List.fold_left ids_of_pat ids patl + let rec ids_of_pat ids = Loc.with_unloc (function + | PatVar Anonymous -> ids + | PatVar(Name id) -> Id.Set.add id ids + | PatCstr(_,patl,_) -> List.fold_left ids_of_pat ids patl + ) in ids_of_pat Id.Set.empty @@ -679,12 +680,12 @@ let zeta_normalize = let expand_as = - let rec add_as map pat = + let rec add_as map (loc, pat) = match pat with | PatVar _ -> map - | PatCstr(_,_,patl,Name id) -> - Id.Map.add id (pattern_to_term pat) (List.fold_left add_as map patl) - | PatCstr(_,_,patl,_) -> List.fold_left add_as map patl + | PatCstr(_,patl,Name id) -> + Id.Map.add id (pattern_to_term (loc, pat)) (List.fold_left add_as map patl) + | PatCstr(_,patl,_) -> List.fold_left add_as map patl in let rec expand_as map rt = match rt with diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 1e405d2c9..ee7b33227 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -193,10 +193,10 @@ let (value_f:Constr.constr list -> global_reference -> Constr.constr) = (d0,RegularStyle,None, [GApp(d0, GRef(d0,fterm,None), List.rev_map (fun x_id -> GVar(d0, x_id)) rev_x_id_l), (Anonymous,None)], - [d0, [v_id], [PatCstr(d0,(destIndRef + [d0, [v_id], [d0,PatCstr((destIndRef (delayed_force coq_sig_ref),1), - [PatVar(d0, Name v_id); - PatVar(d0, Anonymous)], + [d0, PatVar(Name v_id); + d0, PatVar(Anonymous)], Anonymous)], GVar(d0,v_id)]) in diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 6bc2a4f94..ce4610e3b 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -95,7 +95,7 @@ let msg_may_need_inversion () = (* Utils *) let make_anonymous_patvars n = - List.make n (PatVar (Loc.ghost,Anonymous)) + List.make n (Loc.tag @@ PatVar Anonymous) (* We have x1:t1...xn:tn,xi':ti,y1..yk |- c and re-generalize over xi:ti to get x1:t1...xn:tn,xi':ti,y1..yk |- c[xi:=xi'] *) @@ -178,7 +178,7 @@ and build_glob_pattern args = function | Top -> args | MakeConstructor (pci, rh) -> glob_pattern_of_partial_history - [PatCstr (Loc.ghost, pci, args, Anonymous)] rh + [Loc.tag @@ PatCstr (pci, args, Anonymous)] rh let complete_history = glob_pattern_of_partial_history [] @@ -188,12 +188,12 @@ let pop_history_pattern = function | Continuation (0, l, Top) -> Result (List.rev l) | Continuation (0, l, MakeConstructor (pci, rh)) -> - feed_history (PatCstr (Loc.ghost,pci,List.rev l,Anonymous)) rh + feed_history (Loc.tag @@ PatCstr (pci,List.rev l,Anonymous)) rh | _ -> anomaly (Pp.str "Constructor not yet filled with its arguments") let pop_history h = - feed_history (PatVar (Loc.ghost, Anonymous)) h + feed_history (Loc.tag @@ PatVar Anonymous) h (* Builds a continuation expecting [n] arguments and building [ci] applied to this [n] arguments *) @@ -273,8 +273,8 @@ type 'a pattern_matching_problem = let rec find_row_ind = function [] -> None - | PatVar _ :: l -> find_row_ind l - | PatCstr(loc,c,_,_) :: _ -> Some (loc,c) + | (_, PatVar _) :: l -> find_row_ind l + | (loc, PatCstr(c,_,_)) :: _ -> Some (loc,c) let inductive_template evdref env tmloc ind = let indu = evd_comb1 (Evd.fresh_inductive_instance env) evdref ind in @@ -427,9 +427,10 @@ let current_pattern eqn = | pat::_ -> pat | [] -> anomaly (Pp.str "Empty list of patterns") -let alias_of_pat = function - | PatVar (_,name) -> name - | PatCstr(_,_,_,name) -> name +let alias_of_pat = Loc.with_loc (fun ~loc -> function + | PatVar name -> name + | PatCstr(_,_,name) -> name + ) let remove_current_pattern eqn = match eqn.patterns with @@ -472,13 +473,13 @@ let rec adjust_local_defs loc = function | (pat :: pats, LocalAssum _ :: decls) -> pat :: adjust_local_defs loc (pats,decls) | (pats, LocalDef _ :: decls) -> - PatVar (loc, Anonymous) :: adjust_local_defs loc (pats,decls) + (Loc.tag ~loc @@ PatVar Anonymous) :: adjust_local_defs loc (pats,decls) | [], [] -> [] | _ -> raise NotAdjustable let check_and_adjust_constructor env ind cstrs = function - | PatVar _ as pat -> pat - | PatCstr (loc,((_,i) as cstr),args,alias) as pat -> + | _, PatVar _ as pat -> pat + | loc, PatCstr (((_,i) as cstr),args,alias) as pat -> (* Check it is constructor of the right type *) let ind' = inductive_of_constructor cstr in if eq_ind ind' ind then @@ -489,7 +490,7 @@ let check_and_adjust_constructor env ind cstrs = function else try let args' = adjust_local_defs loc (args, List.rev ci.cs_args) - in PatCstr (loc, cstr, args', alias) + in Loc.tag ~loc @@ PatCstr (cstr, args', alias) with NotAdjustable -> error_wrong_numarg_constructor ~loc env cstr nb_args_constr else @@ -502,8 +503,8 @@ let check_and_adjust_constructor env ind cstrs = function let check_all_variables env sigma typ mat = List.iter (fun eqn -> match current_pattern eqn with - | PatVar (_,id) -> () - | PatCstr (loc,cstr_sp,_,_) -> + | _, PatVar id -> () + | loc, PatCstr (cstr_sp,_,_) -> error_bad_pattern ~loc env sigma cstr_sp typ) mat @@ -529,8 +530,8 @@ let occur_in_rhs na rhs = | Name id -> Id.List.mem id rhs.rhs_vars let is_dep_patt_in eqn = function - | PatVar (_,name) -> Flags.is_program_mode () || occur_in_rhs name eqn.rhs - | PatCstr _ -> true + | _, PatVar name -> Flags.is_program_mode () || occur_in_rhs name eqn.rhs + | _, PatCstr _ -> true let mk_dep_patt_row (pats,_,eqn) = List.map (is_dep_patt_in eqn) pats @@ -750,7 +751,7 @@ let recover_and_adjust_alias_names names sign = | x::names, LocalAssum (_,t)::sign -> (x, LocalAssum (alias_of_pat x,t)) :: aux (names,sign) | names, (LocalDef (na,_,_) as decl)::sign -> - (PatVar (Loc.ghost,na), decl) :: aux (names,sign) + (Loc.tag @@ PatVar na, decl) :: aux (names,sign) | _ -> assert false in List.split (aux (names,sign)) @@ -967,7 +968,7 @@ let use_unit_judge evd = evd', j let add_assert_false_case pb tomatch = - let pats = List.map (fun _ -> PatVar (Loc.ghost,Anonymous)) tomatch in + let pats = List.map (fun _ -> Loc.tag @@ PatVar Anonymous) tomatch in let aliasnames = List.map_filter (function Alias _ | NonDepAlias -> Some Anonymous | _ -> None) tomatch in @@ -1165,8 +1166,8 @@ let postprocess_dependencies evd tocheck brs tomatch pred deps cs = (* Sorting equations by constructor *) let rec irrefutable env = function - | PatVar (_,name) -> true - | PatCstr (_,cstr,args,_) -> + | _, PatVar name -> true + | _, PatCstr (cstr,args,_) -> let ind = inductive_of_constructor cstr in let (_,mip) = Inductive.lookup_mind_specif env ind in let one_constr = Int.equal (Array.length mip.mind_user_lc) 1 in @@ -1187,14 +1188,14 @@ let group_equations pb ind current cstrs mat = let rest = remove_current_pattern eqn in let pat = current_pattern eqn in match check_and_adjust_constructor pb.env ind cstrs pat with - | PatVar (_,name) -> + | _, PatVar name -> (* This is a default clause that we expand *) for i=1 to Array.length cstrs do let args = make_anonymous_patvars cstrs.(i-1).cs_nargs in brs.(i-1) <- (args, name, rest) :: brs.(i-1) done; if !only_default == None then only_default := Some true - | PatCstr (loc,((_,i)),args,name) -> + | loc, PatCstr (((_,i)),args,name) -> (* This is a regular clause *) only_default := Some false; brs.(i-1) <- (args, name, rest) :: brs.(i-1)) mat () in @@ -1718,16 +1719,16 @@ let build_tycon loc env tycon_env s subst tycon extenv evdref t = let build_inversion_problem loc env sigma tms t = let make_patvar t (subst,avoid) = let id = next_name_away (named_hd env sigma t Anonymous) avoid in - PatVar (Loc.ghost,Name id), ((id,t)::subst, id::avoid) in + Loc.tag @@ PatVar (Name id), ((id,t)::subst, id::avoid) in let rec reveal_pattern t (subst,avoid as acc) = match EConstr.kind sigma (whd_all env sigma t) with - | Construct (cstr,u) -> PatCstr (Loc.ghost,cstr,[],Anonymous), acc + | Construct (cstr,u) -> Loc.tag (PatCstr (cstr,[],Anonymous)), acc | App (f,v) when isConstruct sigma f -> let cstr,u = destConstruct sigma f in let n = constructor_nrealargs_env env cstr in let l = List.lastn n (Array.to_list v) in let l,acc = List.fold_map' reveal_pattern l acc in - PatCstr (Loc.ghost,cstr,l,Anonymous), acc + Loc.tag (PatCstr (cstr,l,Anonymous)), acc | _ -> make_patvar t acc in let rec aux n env acc_sign tms acc = match tms with @@ -1803,7 +1804,7 @@ let build_inversion_problem loc env sigma tms t = (* No need for a catch all clause *) [] else - [ { patterns = List.map (fun _ -> PatVar (Loc.ghost,Anonymous)) patl; + [ { patterns = List.map (fun _ -> Loc.tag @@ PatVar Anonymous) patl; alias_stack = []; eqn_loc = Loc.ghost; used = ref false; @@ -2063,18 +2064,18 @@ let hole = Misctypes.IntroAnonymous, None) let constr_of_pat env evdref arsign pat avoid = - let rec typ env (ty, realargs) pat avoid = + let rec typ env (ty, realargs) (loc, pat) avoid = match pat with - | PatVar (l,name) -> + | PatVar name -> let name, avoid = match name with Name n -> name, avoid | Anonymous -> let previd, id = prime avoid (Name (Id.of_string "wildcard")) in Name id, id :: avoid in - (PatVar (l, name), [LocalAssum (name, ty)] @ realargs, mkRel 1, ty, + ((Loc.tag ~loc @@ PatVar name), [LocalAssum (name, ty)] @ realargs, mkRel 1, ty, (List.map (fun x -> mkRel 1) realargs), 1, avoid) - | PatCstr (l,((_, i) as cstr),args,alias) -> + | PatCstr (((_, i) as cstr),args,alias) -> let cind = inductive_of_constructor cstr in let IndType (indf, _) = try find_rectype env ( !evdref) (lift (-(List.length realargs)) ty) @@ -2083,7 +2084,7 @@ let constr_of_pat env evdref arsign pat avoid = in let (ind,u), params = dest_ind_family indf in let params = List.map EConstr.of_constr params in - if not (eq_ind ind cind) then error_bad_constructor ~loc:l env cstr ind; + if not (eq_ind ind cind) then error_bad_constructor ~loc env cstr ind; let cstrs = get_constructors env indf in let ci = cstrs.(i-1) in let nb_args_constr = ci.cs_nargs in @@ -2103,7 +2104,7 @@ let constr_of_pat env evdref arsign pat avoid = in let args = List.rev args in let patargs = List.rev patargs in - let pat' = PatCstr (l, cstr, patargs, alias) in + let pat' = Loc.tag ~loc @@ PatCstr (cstr, patargs, alias) in let cstr = mkConstructU (on_snd EInstance.make ci.cs_cstr) in let app = applist (cstr, List.map (lift (List.length sign)) params) in let app = applist (app, args) in @@ -2169,11 +2170,11 @@ let vars_of_ctx sigma ctx = ctx (Id.of_string "vars_of_ctx_error", []) in List.rev y -let rec is_included x y = +let rec is_included (loc_x, x) (loc_y, y) = match x, y with | PatVar _, _ -> true | _, PatVar _ -> true - | PatCstr (l, (_, i), args, alias), PatCstr (l', (_, i'), args', alias') -> + | PatCstr ((_, i), args, alias), PatCstr ((_, i'), args', alias') -> if Int.equal i i' then List.for_all2 is_included args args' else false diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index 542db7fdf..b2c1d0116 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -78,8 +78,9 @@ let apply_coercion_args env evd check isproj argl funj = let apply_pattern_coercion loc pat p = List.fold_left (fun pat (co,n) -> - let f i = if i pat::hd,rest) lines) clauses) | _ -> - let pat = PatVar(dl,update_name sigma na rhs) in + let pat = Loc.tag @@ PatVar(update_name sigma na rhs) in let mat = align_tree nal isgoal rhs sigma in List.map (fun (hd,rest) -> pat::hd,rest) mat @@ -644,17 +644,17 @@ and detype_eqns flags avoid env sigma ci computable constructs consnargsl bl = and detype_eqn (lax,isgoal as flags) avoid env sigma constr construct_nargs branch = let make_pat x avoid env b body ty ids = if force_wildcard () && noccurn sigma 1 b then - PatVar (dl,Anonymous),avoid,(add_name Anonymous body ty env),ids + Loc.tag @@ PatVar (Anonymous),avoid,(add_name Anonymous body ty env),ids else let flag = if isgoal then RenamingForGoal else RenamingForCasesPattern (fst env,b) in let na,avoid' = compute_displayed_name_in sigma flag avoid x b in - PatVar (dl,na),avoid',(add_name na body ty env),add_vname ids na + Loc.tag @@ PatVar na,avoid',(add_name na body ty env),add_vname ids na in let rec buildrec ids patlist avoid env l b = match EConstr.kind sigma b, l with | _, [] -> (dl, Id.Set.elements ids, - [PatCstr(dl, constr, List.rev patlist,Anonymous)], + [Loc.tag ~loc:dl @@ PatCstr(constr, List.rev patlist,Anonymous)], detype flags avoid env sigma b) | Lambda (x,t,b), false::l -> let pat,new_avoid,new_env,new_ids = make_pat x avoid env b None t ids in @@ -668,7 +668,7 @@ and detype_eqn (lax,isgoal as flags) avoid env sigma constr construct_nargs bran buildrec ids patlist avoid env l c | _, true::l -> - let pat = PatVar (dl,Anonymous) in + let pat = Loc.tag ~loc:dl @@ PatVar Anonymous in buildrec ids (pat::patlist) avoid env l b | _, false::l -> @@ -793,14 +793,14 @@ let detype_closed_glob ?lax isgoal avoid env sigma t = (**********************************************************************) (* Module substitution: relies on detyping *) -let rec subst_cases_pattern subst pat = +let rec subst_cases_pattern subst (loc, pat) = Loc.tag ~loc @@ match pat with | PatVar _ -> pat - | PatCstr (loc,((kn,i),j),cpl,n) -> + | PatCstr (((kn,i),j),cpl,n) -> let kn' = subst_mind subst kn and cpl' = List.smartmap (subst_cases_pattern subst) cpl in if kn' == kn && cpl' == cpl then pat else - PatCstr (loc,((kn',i),j),cpl',n) + PatCstr (((kn',i),j),cpl',n) let (f_subst_genarg, subst_genarg_hook) = Hook.make () @@ -910,8 +910,8 @@ let rec subst_glob_constr subst raw = let simple_cases_matrix_of_branches ind brs = List.map (fun (i,n,b) -> let nal,c = it_destRLambda_or_LetIn_names n b in - let mkPatVar na = PatVar (Loc.ghost,na) in - let p = PatCstr (Loc.ghost,(ind,i+1),List.map mkPatVar nal,Anonymous) in + let mkPatVar na = Loc.tag @@ PatVar na in + let p = Loc.tag @@ PatCstr ((ind,i+1),List.map mkPatVar nal,Anonymous) in let map name = try Some (Nameops.out_name name) with Failure _ -> None in let ids = List.map_filter map nal in (Loc.ghost,ids,[p],c)) diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml index ebbfa195f..27ceabf4e 100644 --- a/pretyping/glob_ops.ml +++ b/pretyping/glob_ops.ml @@ -15,9 +15,7 @@ open Glob_term (* Untyped intermediate terms, after ASTs and before constr. *) -let cases_pattern_loc = function - PatVar(loc,_) -> loc - | PatCstr(loc,_,_,_) -> loc +let cases_pattern_loc (loc, _) = loc let cases_predicate_names tml = List.flatten (List.map (function @@ -47,9 +45,9 @@ let case_style_eq s1 s2 = match s1, s2 with | RegularStyle, RegularStyle -> true | _ -> false -let rec cases_pattern_eq p1 p2 = match p1, p2 with -| PatVar (_, na1), PatVar (_, na2) -> Name.equal na1 na2 -| PatCstr (_, c1, pl1, na1), PatCstr (_, c2, pl2, na2) -> +let rec cases_pattern_eq (_loc1, p1) (_loc2, p2) = match p1, p2 with +| PatVar na1, PatVar na2 -> Name.equal na1 na2 +| PatCstr (c1, pl1, na1), PatCstr (c2, pl2, na2) -> eq_constructor c1 c2 && List.equal cases_pattern_eq pl1 pl2 && Name.equal na1 na2 | _ -> false @@ -423,18 +421,19 @@ let map_tomatch_binders f ((c,(na,inp)) as x) : tomatch_tuple = if r == inp then x else c,(f na, r) -let rec map_case_pattern_binders f = function - | PatVar (loc,na) as x -> +let rec map_case_pattern_binders f = Loc.map (function + | PatVar na as x -> let r = f na in if r == na then x - else PatVar (loc,r) - | PatCstr (loc,c,ps,na) as x -> + else PatVar r + | PatCstr (c,ps,na) as x -> let rna = f na in let rps = CList.smartmap (fun p -> map_case_pattern_binders f p) ps in if rna == na && rps == ps then x - else PatCstr(loc,c,rps,rna) + else PatCstr(c,rps,rna) + ) let map_cases_branch_binders f ((loc,il,cll,rhs) as x) : cases_clause = (* spiwack: not sure if I must do something with the list of idents. @@ -558,26 +557,27 @@ let rec cases_pattern_of_glob_constr na = function | Name _ -> (* Unable to manage the presence of both an alias and a variable *) raise Not_found - | Anonymous -> PatVar (loc,Name id) + | Anonymous -> Loc.tag ~loc @@ PatVar (Name id) end - | GHole (loc,_,_,_) -> PatVar (loc,na) + | GHole (loc,_,_,_) -> Loc.tag ~loc @@ PatVar na | GRef (loc,ConstructRef cstr,_) -> - PatCstr (loc,cstr,[],na) + Loc.tag ~loc @@ PatCstr (cstr,[],na) | GApp (loc,GRef (_,ConstructRef cstr,_),l) -> - PatCstr (loc,cstr,List.map (cases_pattern_of_glob_constr Anonymous) l,na) + Loc.tag ~loc @@ PatCstr (cstr,List.map (cases_pattern_of_glob_constr Anonymous) l,na) | _ -> raise Not_found (* Turn a closed cases pattern into a glob_constr *) -let rec glob_constr_of_closed_cases_pattern_aux = function - | PatCstr (loc,cstr,[],Anonymous) -> +let rec glob_constr_of_closed_cases_pattern_aux x = Loc.with_loc (fun ~loc -> function + | PatCstr (cstr,[],Anonymous) -> GRef (loc,ConstructRef cstr,None) - | PatCstr (loc,cstr,l,Anonymous) -> + | PatCstr (cstr,l,Anonymous) -> let ref = GRef (loc,ConstructRef cstr,None) in GApp (loc,ref, List.map glob_constr_of_closed_cases_pattern_aux l) | _ -> raise Not_found + ) x let glob_constr_of_closed_cases_pattern = function - | PatCstr (loc,cstr,l,na) -> - na,glob_constr_of_closed_cases_pattern_aux (PatCstr (loc,cstr,l,Anonymous)) + | loc, PatCstr (cstr,l,na) -> + na,glob_constr_of_closed_cases_pattern_aux (loc, PatCstr (cstr,l,Anonymous)) | _ -> raise Not_found diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index b16d04495..8c570dffe 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -376,7 +376,7 @@ let rec pat_of_raw metas vars = function [0,tags,pat_of_raw metas vars c]) | GCases (loc,sty,p,[c,(na,indnames)],brs) -> let get_ind = function - | (_,_,[PatCstr(_,(ind,_),_,_)],_)::_ -> Some ind + | (_,_,[_, PatCstr((ind,_),_,_)],_)::_ -> Some ind | _ -> None in let ind_tags,ind = match indnames with @@ -408,15 +408,15 @@ let rec pat_of_raw metas vars = function and pats_of_glob_branches loc metas vars ind brs = let get_arg = function - | PatVar(_,na) -> + | _, PatVar na -> name_iter (fun n -> metas := n::!metas) na; na - | PatCstr(loc,_,_,_) -> err ~loc (Pp.str "Non supported pattern.") + | loc, PatCstr(_,_,_) -> err ~loc (Pp.str "Non supported pattern.") in let rec get_pat indexes = function | [] -> false, [] - | [(_,_,[PatVar(_,Anonymous)],GHole _)] -> true, [] (* ends with _ => _ *) - | (_,_,[PatCstr(_,(indsp,j),lv,_)],br) :: brs -> + | [(_,_,[_, PatVar(Anonymous)],GHole _)] -> true, [] (* ends with _ => _ *) + | (_,_,[_, PatCstr((indsp,j),lv,_)],br) :: brs -> let () = match ind with | Some sp when eq_ind sp indsp -> () | _ -> -- cgit v1.2.3 From be83b52cf50ed4c596e40cfd52da03258a7a4a18 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Mon, 16 Jan 2017 13:22:42 +0100 Subject: [location] Move Glob_term.predicate_pattern to located. We continue the uniformization pass. No big news here, trying to be minimally invasive. --- interp/constrextern.ml | 2 +- interp/constrintern.ml | 2 +- interp/notation_ops.ml | 8 ++++---- intf/glob_term.mli | 2 +- pretyping/cases.ml | 8 ++++---- pretyping/detyping.ml | 8 ++++---- pretyping/glob_ops.ml | 10 +++++----- pretyping/patternops.ml | 4 ++-- 8 files changed, 22 insertions(+), 22 deletions(-) (limited to 'intf') diff --git a/interp/constrextern.ml b/interp/constrextern.ml index b3059f5d0..d45f3a9f1 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -749,7 +749,7 @@ let rec extern inctx scopes vars r = | Name _, _ -> Some (Loc.ghost,na) in (sub_extern false scopes vars tm, na', - Option.map (fun (loc,ind,nal) -> + Option.map (fun (loc,(ind,nal)) -> let args = List.map (fun x -> Loc.tag @@ PatVar x) nal in let fullargs = add_cpatt_for_params ind args in extern_ind_pattern_in_scope scopes vars ind fullargs diff --git a/interp/constrintern.ml b/interp/constrintern.ml index 4960d7332..f814205dc 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -1836,7 +1836,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let _,args_rel = List.chop nparams (List.rev mip.Declarations.mind_arity_ctxt) in canonize_args args_rel l (Id.Set.elements forbidden_names_for_gen) [] [] in - match_to_do, Some (cases_pattern_expr_loc t,ind,List.rev_map snd nal) + match_to_do, Some (cases_pattern_expr_loc t,(ind,List.rev_map snd nal)) | None -> [], None in (tm',(snd na,typ)), extra_id, match_td diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml index 29f42d0e9..a25fd81f3 100644 --- a/interp/notation_ops.ml +++ b/interp/notation_ops.ml @@ -179,7 +179,7 @@ let glob_constr_of_notation_constr_with_binders loc g f e = function | Some (ind,nal) -> let e',nal' = List.fold_right (fun na (e',nal) -> let e',na' = g e' na in e',na'::nal) nal (e',[]) in - e',Some (loc,ind,nal') in + e',Some (loc,(ind,nal')) in let e',na' = g e' na in (e',(f e tm,(na',t'))::tml')) tml (e,[]) in let fold (idl,e) na = let (e,na) = g e na in ((name_cons na idl,e),na) in @@ -356,8 +356,8 @@ let notation_constr_and_vars_of_glob_constr a = List.map (fun (tm,(na,x)) -> add_name found na; Option.iter - (fun (_,_,nl) -> List.iter (add_name found) nl) x; - (aux tm,(na,Option.map (fun (_,ind,nal) -> (ind,nal)) x))) tml, + (fun (_,(_,nl)) -> List.iter (add_name found) nl) x; + (aux tm,(na,Option.map (fun (_,(ind,nal)) -> (ind,nal)) x))) tml, List.map f eqnl) | GLetTuple (loc,nal,(na,po),b,c) -> add_name found na; @@ -589,7 +589,7 @@ let abstract_return_type_context pi mklam tml rtno = rtno let abstract_return_type_context_glob_constr = - abstract_return_type_context (fun (_,_,nal) -> nal) + abstract_return_type_context (fun (_,(_,nal)) -> nal) (fun na c -> GLambda(Loc.ghost,na,Explicit,GHole(Loc.ghost,Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None),c)) diff --git a/intf/glob_term.mli b/intf/glob_term.mli index 5e771245c..7a43c44f9 100644 --- a/intf/glob_term.mli +++ b/intf/glob_term.mli @@ -67,7 +67,7 @@ and fix_kind = | GCoFix of int and predicate_pattern = - Name.t * (Loc.t * inductive * Name.t list) option + Name.t * (inductive * Name.t list) Loc.located option (** [(na,id)] = "as 'na' in 'id'" where if [id] is [Some(l,I,k,args)]. *) and tomatch_tuple = (glob_constr * predicate_pattern) diff --git a/pretyping/cases.ml b/pretyping/cases.ml index ce4610e3b..531485935 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -342,7 +342,7 @@ let unify_tomatch_with_patterns evdref env loc typ pats realnames = let find_tomatch_tycon evdref env loc = function (* Try if some 'in I ...' is present and can be used as a constraint *) - | Some (_,ind,realnal) -> + | Some (_,(ind,realnal)) -> mk_tycon (inductive_template evdref env loc ind),Some (List.rev realnal) | None -> empty_tycon,None @@ -360,7 +360,7 @@ let coerce_row typing_fun evdref env pats (tomatch,(_,indopt)) = unify_tomatch_with_patterns evdref env loc typ pats realnames in (j.uj_val,t) -let coerce_to_indtype typing_fun evdref env matx tomatchl = +let coerce_to_indtype typing_fun evdref env matx tomatchl = let pats = List.map (fun r -> r.patterns) matx in let matx' = match matrix_transpose pats with | [] -> List.map (fun _ -> []) tomatchl (* no patterns at all *) @@ -1852,7 +1852,7 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign = | None -> (match bo with | None -> [LocalAssum (na, lift n typ)] | Some b -> [LocalDef (na, lift n b, lift n typ)]) - | Some (loc,_,_) -> + | Some (loc,_) -> user_err ~loc (str"Unexpected type annotation for a term of non inductive type.")) | IsInd (term,IndType(indf,realargs),_) -> @@ -1863,7 +1863,7 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign = let arsign = List.map (fun d -> map_rel_decl EConstr.of_constr d) arsign in let realnal = match t with - | Some (loc,ind',realnal) -> + | Some (loc,(ind',realnal)) -> if not (eq_ind ind ind') then user_err ~loc (str "Wrong inductive type."); if not (Int.equal nrealargs_ctxt (List.length realnal)) then diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index eef6d34ac..f3018ac64 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -380,7 +380,7 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl = | _ -> Anonymous, typ in let aliastyp = if List.for_all (Name.equal Anonymous) nl then None - else Some (dl,indsp,nl) in + else Some (dl,(indsp,nl)) in n, aliastyp, Some typ in let constructs = Array.init (Array.length bl) (fun i -> (indsp,i+1)) in @@ -844,9 +844,9 @@ let rec subst_glob_constr subst raw = let a' = subst_glob_constr subst a in let (n,topt) = x in let topt' = Option.smartmap - (fun (loc,(sp,i),y as t) -> + (fun ((loc,((sp,i),y) as t)) -> let sp' = subst_mind subst sp in - if sp == sp' then t else (loc,(sp',i),y)) topt in + if sp == sp' then t else (loc,((sp',i),y))) topt in if a == a' && topt == topt' then y else (a',(n,topt'))) rl and branches' = List.smartmap (fun (loc,idl,cpl,r as branch) -> @@ -919,4 +919,4 @@ let simple_cases_matrix_of_branches ind brs = let return_type_of_predicate ind nrealargs_tags pred = let nal,p = it_destRLambda_or_LetIn_names (nrealargs_tags@[false]) pred in - (List.hd nal, Some (Loc.ghost, ind, List.tl nal)), Some p + (List.hd nal, Some (Loc.tag (ind, List.tl nal))), Some p diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml index 27ceabf4e..4cccaaf8f 100644 --- a/pretyping/glob_ops.ml +++ b/pretyping/glob_ops.ml @@ -20,7 +20,7 @@ let cases_pattern_loc (loc, _) = loc let cases_predicate_names tml = List.flatten (List.map (function | (tm,(na,None)) -> [na] - | (tm,(na,Some (_,_,nal))) -> na::nal) tml) + | (tm,(na,Some (_,(_,nal)))) -> na::nal) tml) let mkGApp loc p t = match p with @@ -103,7 +103,7 @@ let rec glob_constr_eq c1 c2 = match c1, c2 with | _ -> false and tomatch_tuple_eq (c1, p1) (c2, p2) = - let eqp (_, i1, na1) (_, i2, na2) = + let eqp (_, (i1, na1)) (_, (i2, na2)) = eq_ind i1 i2 && List.equal Name.equal na1 na2 in let eq_pred (n1, o1) (n2, o2) = Name.equal n1 n2 && Option.equal eqp o1 o2 in @@ -411,10 +411,10 @@ let bound_glob_vars = probably be no significant penalty in doing reallocation as pattern-matching expressions are usually rather small. *) -let map_inpattern_binders f ((loc,id,nal) as x) = +let map_inpattern_binders f ((loc,(id,nal)) as x) = let r = CList.smartmap f nal in if r == nal then x - else loc,id,r + else loc,(id,r) let map_tomatch_binders f ((c,(na,inp)) as x) : tomatch_tuple = let r = Option.smartmap (fun p -> map_inpattern_binders f p) inp in @@ -525,7 +525,7 @@ let rec rename_glob_vars l = function (* Lazy strategy: we fail if a collision with renaming occurs, rather than renaming further *) | GCases (loc,ci,po,tomatchl,cls) -> let test_pred_pat (na,ino) = - test_na l na; Option.iter (fun (_,_,nal) -> List.iter (test_na l) nal) ino in + test_na l na; Option.iter (fun (_,(_,nal)) -> List.iter (test_na l) nal) ino in let test_clause idl = List.iter (test_id l) idl in let po = Option.map (rename_glob_vars l) po in let tomatchl = Util.List.map_left (fun (tm,x) -> test_pred_pat x; (rename_glob_vars l tm,x)) tomatchl in diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index 8c570dffe..48ae93f3e 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -380,13 +380,13 @@ let rec pat_of_raw metas vars = function | _ -> None in let ind_tags,ind = match indnames with - | Some (_,ind,nal) -> Some (List.length nal), Some ind + | Some (_,(ind,nal)) -> Some (List.length nal), Some ind | None -> None, get_ind brs in let ext,brs = pats_of_glob_branches loc metas vars ind brs in let pred = match p,indnames with - | Some p, Some (_,_,nal) -> + | Some p, Some (_,(_,nal)) -> let nvars = na :: List.rev nal @ vars in rev_it_mkPLambda nal (mkPLambda na (pat_of_raw metas nvars p)) | (None | Some (GHole _)), _ -> PMeta None -- cgit v1.2.3 From 158f40db9482ead89befbf9bc9ad45ff8a60b75f Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Tue, 17 Jan 2017 14:23:53 +0100 Subject: [location] Switch glob_constr to Loc.located --- dev/top_printers.ml | 2 +- interp/constrextern.ml | 165 +++++++------- interp/constrintern.ml | 147 ++++++------ interp/implicit_quantifiers.ml | 33 +-- interp/notation.ml | 23 +- interp/notation_ops.ml | 289 ++++++++++++------------ intf/glob_term.mli | 40 ++-- lib/loc.ml | 1 + lib/loc.mli | 1 + plugins/decl_mode/g_decl_mode.ml4 | 387 -------------------------------- plugins/funind/glob_term_to_relation.ml | 102 ++++----- plugins/funind/glob_termops.ml | 328 +++++++++++++-------------- plugins/funind/glob_termops.mli | 7 +- plugins/funind/indfun.ml | 20 +- plugins/funind/indfun_common.ml | 10 +- plugins/funind/merge.ml | 52 ++--- plugins/funind/recdef.ml | 18 +- plugins/ltac/extratactics.ml4 | 12 +- plugins/ltac/g_rewrite.ml4 | 2 +- plugins/ltac/pptactic.ml | 4 +- plugins/ltac/tacintern.ml | 22 +- plugins/ltac/tacinterp.ml | 4 +- plugins/setoid_ring/newring.ml | 2 +- plugins/ssrmatching/ssrmatching.ml4 | 39 ++-- plugins/syntax/ascii_syntax.ml | 14 +- plugins/syntax/nat_syntax.ml | 17 +- plugins/syntax/numbers_syntax.ml | 84 +++---- plugins/syntax/r_syntax.ml | 32 +-- plugins/syntax/string_syntax.ml | 16 +- plugins/syntax/z_syntax.ml | 56 ++--- pretyping/cases.ml | 26 +-- pretyping/detyping.ml | 208 ++++++++--------- pretyping/detyping.mli | 6 +- pretyping/glob_ops.ml | 268 +++++++++++----------- pretyping/patternops.ml | 45 ++-- pretyping/pretyping.ml | 36 +-- tactics/hipattern.ml | 20 +- vernac/command.ml | 14 +- 38 files changed, 1082 insertions(+), 1470 deletions(-) delete mode 100644 plugins/decl_mode/g_decl_mode.ml4 (limited to 'intf') diff --git a/dev/top_printers.ml b/dev/top_printers.ml index 474cc85c1..52cf8cf97 100644 --- a/dev/top_printers.ml +++ b/dev/top_printers.ml @@ -26,7 +26,7 @@ open Clenv let _ = Detyping.print_evar_arguments := true let _ = Detyping.print_universes := true let _ = set_bool_option_value ["Printing";"Matching"] false -let _ = Detyping.set_detype_anonymous (fun _ _ -> raise Not_found) +let _ = Detyping.set_detype_anonymous (fun ?loc _ -> raise Not_found) (* std_ppcmds *) let pp x = Pp.pp_with !Topfmt.std_ft x diff --git a/interp/constrextern.ml b/interp/constrextern.ml index d45f3a9f1..bbc98dd28 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -157,7 +157,7 @@ let insert_pat_alias loc p = function (**********************************************************************) (* conversion of references *) -let extern_evar loc n l = Loc.tag @@ CEvar (n,l) +let extern_evar n l = CEvar (n,l) (** We allow customization of the global_reference printer. For instance, in the debugger the tables of global references @@ -475,7 +475,7 @@ exception Expl (* Implicit args indexes are in ascending order *) (* inctx is useful only if there is a last argument to be deduced from ctxt *) -let explicitize loc inctx impl (cf,f) args = +let explicitize inctx impl (cf,f) args = let impl = if !Constrintern.parsing_explicit then [] else impl in let n = List.length args in let rec exprec q = function @@ -512,41 +512,41 @@ let explicitize loc inctx impl (cf,f) args = let args1 = exprec 1 (args1,impl1) in let args2 = exprec (i+1) (args2,impl2) in let ip = Some (List.length args1) in - Loc.tag ~loc @@ CApp ((ip,f),args1@args2) + CApp ((ip,f),args1@args2) | None -> let args = exprec 1 (args,impl) in - if List.is_empty args then f else Loc.tag ~loc @@ CApp ((None, f), args) + if List.is_empty args then snd f else CApp ((None, f), args) in try expl () with Expl -> - let f',us = match f with _loc, CRef (f,us) -> f,us | _ -> assert false in + let f',us = match f with (_loc, CRef (f,us)) -> f,us | _ -> assert false in let ip = if !print_projections then ip else None in - Loc.tag ~loc @@ CAppExpl ((ip, f', us), List.map Lazy.force args) + CAppExpl ((ip, f', us), List.map Lazy.force args) let is_start_implicit = function | imp :: _ -> is_status_implicit imp && maximal_insertion_of imp | [] -> false -let extern_global loc impl f us = +let extern_global impl f us = if not !Constrintern.parsing_explicit && is_start_implicit impl then - Loc.tag ~loc @@ CAppExpl ((None, f, us), []) + CAppExpl ((None, f, us), []) else - Loc.tag ~loc @@ CRef (f,us) + CRef (f,us) -let extern_app loc inctx impl (cf,f) us args = +let extern_app inctx impl (cf,f) us args = if List.is_empty args then (* If coming from a notation "Notation a := @b" *) - Loc.tag ~loc @@ CAppExpl ((None, f, us), []) + CAppExpl ((None, f, us), []) else if not !Constrintern.parsing_explicit && ((!Flags.raw_print || (!print_implicits && not !print_implicits_explicit_args)) && List.exists is_status_implicit impl) then let args = List.map Lazy.force args in - Loc.tag ~loc @@ CAppExpl ((is_projection (List.length args) cf,f,us), args) + CAppExpl ((is_projection (List.length args) cf,f,us), args) else - explicitize loc inctx impl (cf, Loc.tag ~loc @@ CRef (f,us)) args + explicitize inctx impl (cf, Loc.tag @@ CRef (f,us)) args let rec fill_arg_scopes args subscopes scopes = match args, subscopes with | [], _ -> [] @@ -560,7 +560,7 @@ let extern_args extern env args = List.map map args let match_coercion_app = function - | GApp (loc,GRef (_,r,_),args) -> Some (loc, r, 0, args) + | (loc, GApp ((_, GRef (r,_)),args)) -> Some (loc, r, 0, args) | _ -> None let rec remove_coercions inctx c = @@ -582,13 +582,13 @@ let rec remove_coercions inctx c = been confused with ordinary application or would have need a surrounding context and the coercion to funclass would have been made explicit to match *) - if List.is_empty l then a' else GApp (loc,a',l) + if List.is_empty l then a' else Loc.tag ~loc @@ GApp (a',l) | _ -> c with Not_found -> c) | _ -> c let rec flatten_application = function - | GApp (loc,GApp(_,a,l'),l) -> flatten_application (GApp (loc,a,l'@l)) + | (loc, GApp ((_, GApp(a,l')),l)) -> flatten_application (Loc.tag ~loc @@ GApp (a,l'@l)) | a -> a (**********************************************************************) @@ -616,7 +616,7 @@ let extern_optimal_prim_token scopes r r' = let extended_glob_local_binder_of_decl loc = function | (p,bk,None,t) -> GLocalAssum (loc,p,bk,t) - | (p,bk,Some x,GHole (_, _, Misctypes.IntroAnonymous, None)) -> GLocalDef (loc,p,bk,x,None) + | (p,bk,Some x,(_,GHole ( _, Misctypes.IntroAnonymous, None))) -> GLocalDef (loc,p,bk,x,None) | (p,bk,Some x,t) -> GLocalDef (loc,p,bk,x,Some t) (**********************************************************************) @@ -642,25 +642,25 @@ let rec extern inctx scopes vars r = let r'' = flatten_application r' in if !Flags.raw_print || !print_no_symbol then raise No_match; extern_notation scopes vars r'' (uninterp_notations r'') - with No_match -> match r' with - | GRef (loc,ref,us) -> - extern_global loc (select_stronger_impargs (implicits_of_global ref)) + with No_match -> Loc.map_with_loc (fun ~loc -> function + | GRef (ref,us) -> + extern_global (select_stronger_impargs (implicits_of_global ref)) (extern_reference loc vars ref) (extern_universes us) - | GVar (loc,id) -> Loc.tag ~loc @@ CRef (Ident (loc,id),None) + | GVar id -> CRef (Ident (loc,id),None) - | GEvar (loc,n,[]) when !print_meta_as_hole -> Loc.tag ~loc @@ CHole (None, Misctypes.IntroAnonymous, None) + | GEvar (n,[]) when !print_meta_as_hole -> CHole (None, Misctypes.IntroAnonymous, None) - | GEvar (loc,n,l) -> - extern_evar loc n (List.map (on_snd (extern false scopes vars)) l) + | GEvar (n,l) -> + extern_evar n (List.map (on_snd (extern false scopes vars)) l) - | GPatVar (loc,(b,n)) -> Loc.tag ~loc @@ + | GPatVar (b,n) -> if !print_meta_as_hole then CHole (None, Misctypes.IntroAnonymous, None) else if b then CPatVar n else CEvar (n,[]) - | GApp (loc,f,args) -> + | GApp (f,args) -> (match f with - | GRef (rloc,ref,us) -> + | (rloc, GRef (ref,us)) -> let subscopes = find_arguments_scope ref in let args = fill_arg_scopes args subscopes (snd scopes) in begin @@ -701,42 +701,42 @@ let rec extern inctx scopes vars r = let head = extern true scopes vars arg in ip q locs' tail ((extern_reference loc Id.Set.empty (ConstRef c), head) :: acc) in - Loc.tag ~loc @@ CRecord (List.rev (ip projs locals args [])) + CRecord (List.rev (ip projs locals args [])) with | Not_found | No_match | Exit -> let args = extern_args (extern true) vars args in - extern_app loc inctx + extern_app inctx (select_stronger_impargs (implicits_of_global ref)) (Some ref,extern_reference rloc vars ref) (extern_universes us) args end - + | _ -> - explicitize loc inctx [] (None,sub_extern false scopes vars f) + explicitize inctx [] (None,sub_extern false scopes vars f) (List.map (fun c -> lazy (sub_extern true scopes vars c)) args)) - | GLetIn (loc,na,b,t,c) -> - Loc.tag ~loc @@ CLetIn ((loc,na),sub_extern false scopes vars b, + | GLetIn (na,b,t,c) -> + CLetIn ((loc,na),sub_extern false scopes vars b, Option.map (extern_typ scopes vars) t, extern inctx scopes (add_vname vars na) c) - | GProd (loc,na,bk,t,c) -> + | GProd (na,bk,t,c) -> let t = extern_typ scopes vars t in let (idl,c) = factorize_prod scopes (add_vname vars na) na bk t c in - Loc.tag ~loc @@ CProdN ([(Loc.ghost,na)::idl,Default bk,t],c) + CProdN ([(Loc.ghost,na)::idl,Default bk,t],c) - | GLambda (loc,na,bk,t,c) -> + | GLambda (na,bk,t,c) -> let t = extern_typ scopes vars t in let (idl,c) = factorize_lambda inctx scopes (add_vname vars na) na bk t c in - Loc.tag ~loc @@ CLambdaN ([(Loc.ghost,na)::idl,Default bk,t],c) + CLambdaN ([(Loc.ghost,na)::idl,Default bk,t],c) - | GCases (loc,sty,rtntypopt,tml,eqns) -> + | GCases (sty,rtntypopt,tml,eqns) -> let vars' = List.fold_right (name_fold Id.Set.add) (cases_predicate_names tml) vars in let rtntypopt' = Option.map (extern_typ scopes vars') rtntypopt in let tml = List.map (fun (tm,(na,x)) -> let na' = match na,tm with - | Anonymous, GVar (_, id) -> + | Anonymous, (_, GVar id) -> begin match rtntypopt with | None -> None | Some ntn -> @@ -745,7 +745,7 @@ let rec extern inctx scopes vars r = else None end | Anonymous, _ -> None - | Name id, GVar (_,id') when Id.equal id id' -> None + | Name id, (_, GVar id') when Id.equal id id' -> None | Name _, _ -> Some (Loc.ghost,na) in (sub_extern false scopes vars tm, na', @@ -757,22 +757,22 @@ let rec extern inctx scopes vars r = tml in let eqns = List.map (extern_eqn inctx scopes vars) eqns in - Loc.tag ~loc @@ CCases (sty,rtntypopt',tml,eqns) + CCases (sty,rtntypopt',tml,eqns) - | GLetTuple (loc,nal,(na,typopt),tm,b) -> - Loc.tag ~loc @@ CLetTuple (List.map (fun na -> (Loc.ghost,na)) nal, + | GLetTuple (nal,(na,typopt),tm,b) -> + CLetTuple (List.map (fun na -> (Loc.ghost,na)) nal, (Option.map (fun _ -> (Loc.ghost,na)) typopt, Option.map (extern_typ scopes (add_vname vars na)) typopt), sub_extern false scopes vars tm, extern inctx scopes (List.fold_left add_vname vars nal) b) - | GIf (loc,c,(na,typopt),b1,b2) -> - Loc.tag ~loc @@ CIf (sub_extern false scopes vars c, + | GIf (c,(na,typopt),b1,b2) -> + CIf (sub_extern false scopes vars c, (Option.map (fun _ -> (Loc.ghost,na)) typopt, Option.map (extern_typ scopes (add_vname vars na)) typopt), sub_extern inctx scopes vars b1, sub_extern inctx scopes vars b2) - | GRec (loc,fk,idv,blv,tyv,bv) -> + | GRec (fk,idv,blv,tyv,bv) -> let vars' = Array.fold_right Id.Set.add idv vars in (match fk with | GFix (nv,n) -> @@ -792,7 +792,7 @@ let rec extern inctx scopes vars r = ((Loc.ghost, fi), (n, ro), bl, extern_typ scopes vars0 ty, extern false scopes vars1 def)) idv in - Loc.tag ~loc @@ CFix ((loc,idv.(n)),Array.to_list listdecl) + CFix ((loc,idv.(n)),Array.to_list listdecl) | GCoFix n -> let listdecl = Array.mapi (fun i fi -> @@ -803,15 +803,16 @@ let rec extern inctx scopes vars r = ((Loc.ghost, fi),bl,extern_typ scopes vars0 tyv.(i), sub_extern false scopes vars1 bv.(i))) idv in - Loc.tag ~loc @@ CCoFix ((loc,idv.(n)),Array.to_list listdecl)) + CCoFix ((loc,idv.(n)),Array.to_list listdecl)) - | GSort (loc,s) -> Loc.tag ~loc @@ CSort (extern_glob_sort s) + | GSort s -> CSort (extern_glob_sort s) - | GHole (loc,e,naming,_) -> Loc.tag ~loc @@ CHole (Some e, naming, None) (** TODO: extern tactics. *) + | GHole (e,naming,_) -> CHole (Some e, naming, None) (** TODO: extern tactics. *) - | GCast (loc,c, c') -> - Loc.tag ~loc @@ CCast (sub_extern true scopes vars c, + | GCast (c, c') -> + CCast (sub_extern true scopes vars c, Miscops.map_cast_type (extern_typ scopes vars) c') + ) r' and extern_typ (_,scopes) = extern true (Notation.current_type_scope_name (),scopes) @@ -867,7 +868,7 @@ and extern_local_binder scopes vars = function let (assums,ids,l) = extern_local_binder scopes vars l in (assums,ids, CLocalPattern(Loc.ghost,p,ty) :: l) -and extern_eqn inctx scopes vars (loc,ids,pl,c) = +and extern_eqn inctx scopes vars (loc,(ids,pl,c)) = Loc.tag ~loc ([loc,List.map (extern_cases_pattern_in_scope scopes vars) pl], extern inctx scopes vars c) @@ -878,13 +879,13 @@ and extern_notation (tmp_scope,scopes as allscopes) vars t = function try if List.mem keyrule !print_non_active_notations then raise No_match; (* Adjusts to the number of arguments expected by the notation *) - let (t,args,argsscopes,argsimpls) = match t,n with - | GApp (_,f,args), Some n + let (t,args,argsscopes,argsimpls) = match snd t,n with + | GApp (f,args), Some n when List.length args >= n -> let args1, args2 = List.chop n args in let subscopes, impls = - match f with - | GRef (_,ref,us) -> + match snd f with + | GRef (ref,us) -> let subscopes = try List.skipn n (find_arguments_scope ref) with Failure _ -> [] in @@ -896,15 +897,15 @@ and extern_notation (tmp_scope,scopes as allscopes) vars t = function subscopes,impls | _ -> [], [] in - (if Int.equal n 0 then f else GApp (Loc.ghost,f,args1)), + (if Int.equal n 0 then f else Loc.tag @@ GApp (f,args1)), args2, subscopes, impls - | GApp (_,(GRef (_,ref,us) as f),args), None -> + | GApp ((_, GRef (ref,us) as f),args), None -> let subscopes = find_arguments_scope ref in let impls = select_impargs_size (List.length args) (implicits_of_global ref) in f, args, subscopes, impls - | GRef (_,ref,us), Some 0 -> GApp (Loc.ghost,t,[]), [], [], [] + | GRef (ref,us), Some 0 -> Loc.tag @@ GApp (t,[]), [], [], [] | _, None -> t, [], [], [] | _ -> raise No_match in (* Try matching ... *) @@ -945,7 +946,7 @@ and extern_notation (tmp_scope,scopes as allscopes) vars t = function else let args = fill_arg_scopes args argsscopes scopes in let args = extern_args (extern true) vars args in - explicitize loc false argsimpls (None,e) args + Loc.tag ~loc @@ explicitize false argsimpls (None,e) args with No_match -> extern_notation allscopes vars t rules @@ -965,8 +966,6 @@ let extern_glob_type vars c = (******************************************************************) (* Main translation function from constr -> constr_expr *) -let loc = Loc.ghost (* for constr and pattern, locations are lost *) - let extern_constr_gen lax goal_concl_style scopt env sigma t = (* "goal_concl_style" means do alpha-conversion using the "goal" convention *) (* i.e.: avoid using the names of goal/section/rel variables and the short *) @@ -1008,11 +1007,11 @@ let extern_closed_glob ?lax goal_concl_style env sigma t = let any_any_branch = (* | _ => _ *) - (loc,[],[Loc.tag ~loc @@ PatVar Anonymous],GHole (loc,Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None)) + Loc.tag ([],[Loc.tag @@ PatVar Anonymous], Loc.tag @@ GHole (Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None)) -let rec glob_of_pat env sigma = function - | PRef ref -> GRef (loc,ref,None) - | PVar id -> GVar (loc,id) +let rec glob_of_pat env sigma pat = Loc.tag @@ match pat with + | PRef ref -> GRef (ref,None) + | PVar id -> GVar id | PEvar (evk,l) -> let test decl = function PVar id' -> Id.equal (NamedDecl.get_id decl) id' | _ -> false in let l = Evd.evar_instance_array test (Evd.find sigma evk) l in @@ -1020,36 +1019,36 @@ let rec glob_of_pat env sigma = function | None -> Id.of_string "__" | Some id -> id in - GEvar (loc,id,List.map (on_snd (glob_of_pat env sigma)) l) + GEvar (id,List.map (on_snd (glob_of_pat env sigma)) l) | PRel n -> let id = try match lookup_name_of_rel n env with | Name id -> id | Anonymous -> anomaly ~label:"glob_constr_of_pattern" (Pp.str "index to an anonymous variable") with Not_found -> Id.of_string ("_UNBOUND_REL_"^(string_of_int n)) in - GVar (loc,id) - | PMeta None -> GHole (loc,Evar_kinds.InternalHole, Misctypes.IntroAnonymous,None) - | PMeta (Some n) -> GPatVar (loc,(false,n)) - | PProj (p,c) -> GApp (loc,GRef (loc, ConstRef (Projection.constant p),None), + GVar id + | PMeta None -> GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous,None) + | PMeta (Some n) -> GPatVar (false,n) + | PProj (p,c) -> GApp (Loc.tag @@ GRef (ConstRef (Projection.constant p),None), [glob_of_pat env sigma c]) | PApp (f,args) -> - GApp (loc,glob_of_pat env sigma f,Array.map_to_list (glob_of_pat env sigma) args) + GApp (glob_of_pat env sigma f,Array.map_to_list (glob_of_pat env sigma) args) | PSoApp (n,args) -> - GApp (loc,GPatVar (loc,(true,n)), + GApp (Loc.tag @@ GPatVar (true,n), List.map (glob_of_pat env sigma) args) | PProd (na,t,c) -> - GProd (loc,na,Explicit,glob_of_pat env sigma t,glob_of_pat (na::env) sigma c) + GProd (na,Explicit,glob_of_pat env sigma t,glob_of_pat (na::env) sigma c) | PLetIn (na,b,t,c) -> - GLetIn (loc,na,glob_of_pat env sigma b, Option.map (glob_of_pat env sigma) t, + GLetIn (na,glob_of_pat env sigma b, Option.map (glob_of_pat env sigma) t, glob_of_pat (na::env) sigma c) | PLambda (na,t,c) -> - GLambda (loc,na,Explicit,glob_of_pat env sigma t, glob_of_pat (na::env) sigma c) + GLambda (na,Explicit,glob_of_pat env sigma t, glob_of_pat (na::env) sigma c) | PIf (c,b1,b2) -> - GIf (loc, glob_of_pat env sigma c, (Anonymous,None), + GIf (glob_of_pat env sigma c, (Anonymous,None), glob_of_pat env sigma b1, glob_of_pat env sigma b2) | PCase ({cip_style=LetStyle; cip_ind_tags=None},PMeta None,tm,[(0,n,b)]) -> let nal,b = it_destRLambda_or_LetIn_names n (glob_of_pat env sigma b) in - GLetTuple (loc,nal,(Anonymous,None),glob_of_pat env sigma tm,b) + GLetTuple (nal,(Anonymous,None),glob_of_pat env sigma tm,b) | PCase (info,p,tm,bl) -> let mat = match bl, info.cip_ind with | [], _ -> [] @@ -1066,10 +1065,10 @@ let rec glob_of_pat env sigma = function return_type_of_predicate ind nargs (glob_of_pat env sigma p) | _ -> anomaly (Pp.str "PCase with non-trivial predicate but unknown inductive") in - GCases (loc,RegularStyle,rtn,[glob_of_pat env sigma tm,indnames],mat) - | PFix f -> Detyping.detype_names false [] env (Global.env()) sigma (EConstr.of_constr (mkFix f)) (** FIXME bad env *) - | PCoFix c -> Detyping.detype_names false [] env (Global.env()) sigma (EConstr.of_constr (mkCoFix c)) - | PSort s -> GSort (loc,s) + GCases (RegularStyle,rtn,[glob_of_pat env sigma tm,indnames],mat) + | PFix f -> Loc.obj @@ Detyping.detype_names false [] env (Global.env()) sigma (EConstr.of_constr (mkFix f)) (** FIXME bad env *) + | PCoFix c -> Loc.obj @@ Detyping.detype_names false [] env (Global.env()) sigma (EConstr.of_constr (mkCoFix c)) + | PSort s -> GSort s let extern_constr_pattern env sigma pat = extern true (None,[]) Id.Set.empty (glob_of_pat env sigma pat) diff --git a/interp/constrintern.ml b/interp/constrintern.ml index f814205dc..cc7203ac0 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -304,12 +304,12 @@ let reset_tmp_scope env = {env with tmp_scope = None} let rec it_mkGProd loc2 env body = match env with - (loc1, (na, bk, t)) :: tl -> it_mkGProd loc2 tl (GProd (Loc.merge loc1 loc2, na, bk, t, body)) + (loc1, (na, bk, t)) :: tl -> it_mkGProd loc2 tl (Loc.tag ~loc:(Loc.merge loc1 loc2) @@ GProd (na, bk, t, body)) | [] -> body let rec it_mkGLambda loc2 env body = match env with - (loc1, (na, bk, t)) :: tl -> it_mkGLambda loc2 tl (GLambda (Loc.merge loc1 loc2, na, bk, t, body)) + (loc1, (na, bk, t)) :: tl -> it_mkGLambda loc2 tl (Loc.tag ~loc:(Loc.merge loc1 loc2) @@ GLambda (na, bk, t, body)) | [] -> body (**********************************************************************) @@ -322,14 +322,14 @@ let build_impls = function let impls_type_list ?(args = []) = let rec aux acc = function - |GProd (_,na,bk,_,c) -> aux ((build_impls bk na)::acc) c + |_, GProd (na,bk,_,c) -> aux ((build_impls bk na)::acc) c |_ -> (Variable,[],List.append args (List.rev acc),[]) in aux [] let impls_term_list ?(args = []) = let rec aux acc = function - |GLambda (_,na,bk,_,c) -> aux ((build_impls bk na)::acc) c - |GRec (_, fix_kind, nas, args, tys, bds) -> + |_, GLambda (na,bk,_,c) -> aux ((build_impls bk na)::acc) c + |_, GRec (fix_kind, nas, args, tys, bds) -> let nb = match fix_kind with |GFix (_, n) -> n | GCoFix n -> n in let acc' = List.fold_left (fun a (na, bk, _, _) -> (build_impls bk na)::a) acc args.(nb) in aux acc' bds.(nb) @@ -346,12 +346,12 @@ let rec check_capture ty = function () let locate_if_hole loc na = function - | GHole (_,_,naming,arg) -> + | _, GHole (_,naming,arg) -> (try match na with | Name id -> glob_constr_of_notation_constr loc (Reserve.find_reserved_type id) | Anonymous -> raise Not_found - with Not_found -> GHole (loc, Evar_kinds.BinderType na, naming, arg)) + with Not_found -> Loc.tag ~loc @@ GHole (Evar_kinds.BinderType na, naming, arg)) | x -> x let reset_hidden_inductive_implicit_test env = @@ -397,7 +397,7 @@ let intern_generalized_binder ?(global_level=false) intern_type lvar env fvs in let bl = List.map (fun (id, loc) -> - (loc, (Name id, b, GHole (loc, Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None)))) + (loc, (Name id, b, Loc.tag ~loc @@ GHole (Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None)))) fvs in let na = match na with @@ -458,7 +458,7 @@ let glob_local_binder_of_extended = function | GLocalAssum (loc,na,bk,t) -> (na,bk,None,t) | GLocalDef (loc,na,bk,c,Some t) -> (na,bk,Some c,t) | GLocalDef (loc,na,bk,c,None) -> - let t = GHole(loc,Evar_kinds.BinderType na,Misctypes.IntroAnonymous,None) in + let t = Loc.tag ~loc @@ GHole(Evar_kinds.BinderType na,Misctypes.IntroAnonymous,None) in (na,bk,Some c,t) | GLocalPattern (loc,_,_,_,_) -> Loc.raise ~loc (Stream.Error "pattern with quote not allowed here.") @@ -517,10 +517,12 @@ let intern_generalization intern env lvar loc bk ak c = in if pi then (fun (id, loc') acc -> - GProd (Loc.merge loc' loc, Name id, bk, GHole (loc', Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None), acc)) + Loc.tag ~loc:(Loc.merge loc' loc) @@ + GProd (Name id, bk, Loc.tag ~loc:loc' @@ GHole (Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None), acc)) else (fun (id, loc') acc -> - GLambda (Loc.merge loc' loc, Name id, bk, GHole (loc', Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None), acc)) + Loc.tag ~loc:(Loc.merge loc' loc) @@ + GLambda (Name id, bk, Loc.tag ~loc:loc' @@ GHole (Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None), acc)) in List.fold_right (fun (id, loc as lid) (env, acc) -> let env' = push_name_env lvar (Variable,[],[],[]) env (loc, Name id) in @@ -572,10 +574,10 @@ let make_letins = (fun a c -> match a with | LPLetIn (loc,(na,b,t)) -> - GLetIn(loc,na,b,t,c) + Loc.tag ~loc @@ GLetIn(na,b,t,c) | LPCases (loc,(cp,il),id) -> - let tt = (GVar(loc,id),(Name id,None)) in - GCases(loc,Misctypes.LetPatternStyle,None,[tt],[(loc,il,[cp],c)])) + let tt = (Loc.tag ~loc @@ GVar id, (Name id,None)) in + Loc.tag ~loc @@ GCases(Misctypes.LetPatternStyle,None,[tt],[(loc,(il,[cp],c))])) let rec subordinate_letins letins = function (* binders come in reverse order; the non-let are returned in reverse order together *) @@ -660,7 +662,7 @@ let instantiate_notation_constr loc intern ntnvars subst infos c = let bindings = Id.Map.map mk_env terms in Some (Genintern.generic_substitute_notation bindings arg) in - GHole (loc, knd, naming, arg) + Loc.tag ~loc @@ GHole (knd, naming, arg) | NBinderList (x,y,iter,terminator) -> (try (* All elements of the list are in scopes (scopt,subscopes) *) @@ -678,22 +680,22 @@ let instantiate_notation_constr loc intern ntnvars subst infos c = let a,letins = snd (Option.get binderopt) in let e = make_letins letins (aux subst' infos c') in let (loc,(na,bk,t)) = a in - GProd (loc,na,bk,t,e) + Loc.tag ~loc @@ GProd (na,bk,t,e) | NLambda (Name id,NHole _,c') when option_mem_assoc id binderopt -> let a,letins = snd (Option.get binderopt) in let (loc,(na,bk,t)) = a in - GLambda (loc,na,bk,t,make_letins letins (aux subst' infos c')) + Loc.tag ~loc @@ GLambda (na,bk,t,make_letins letins (aux subst' infos c')) (* Two special cases to keep binder name synchronous with BinderType *) | NProd (na,NHole(Evar_kinds.BinderType na',naming,arg),c') when Name.equal na na' -> let subinfos,na = traverse_binder subst avoid subinfos na in - let ty = GHole (loc,Evar_kinds.BinderType na,naming,arg) in - GProd (loc,na,Explicit,ty,aux subst' subinfos c') + let ty = Loc.tag ~loc @@ GHole (Evar_kinds.BinderType na,naming,arg) in + Loc.tag ~loc @@ GProd (na,Explicit,ty,aux subst' subinfos c') | NLambda (na,NHole(Evar_kinds.BinderType na',naming,arg),c') when Name.equal na na' -> let subinfos,na = traverse_binder subst avoid subinfos na in - let ty = GHole (loc,Evar_kinds.BinderType na,naming,arg) in - GLambda (loc,na,Explicit,ty,aux subst' subinfos c') + let ty = Loc.tag ~loc @@ GHole (Evar_kinds.BinderType na,naming,arg) in + Loc.tag ~loc @@ GLambda (na,Explicit,ty,aux subst' subinfos c') | t -> glob_constr_of_notation_constr_with_binders loc (traverse_binder subst avoid) (aux subst') subinfos t @@ -705,11 +707,12 @@ let instantiate_notation_constr loc intern ntnvars subst infos c = intern {env with tmp_scope = scopt; scopes = subscopes @ env.scopes} a with Not_found -> + Loc.tag ~loc ( try - GVar (loc, Id.Map.find id renaming) + GVar (Id.Map.find id renaming) with Not_found -> (* Happens for local notation joint with inductive/fixpoint defs *) - GVar (loc,id) + GVar id) in aux (terms,None,None) infos c let split_by_type ids = @@ -744,7 +747,7 @@ let string_of_ty = function | Variable -> "var" let gvar (loc, id) us = match us with -| None -> GVar (loc, id) +| None -> Loc.tag ~loc @@ GVar id | Some _ -> user_err ~loc (str "Variable " ++ pr_id id ++ str " cannot have a universe instance") @@ -786,25 +789,25 @@ let intern_var genv (ltacvars,ntnvars) namedctx loc id us = let impls = implicits_of_global ref in let scopes = find_arguments_scope ref in Dumpglob.dump_reference loc "<>" (string_of_qualid (Decls.variable_secpath id)) "var"; - GRef (loc, ref, us), impls, scopes, [] + Loc.tag ~loc @@ GRef (ref, us), impls, scopes, [] with e when CErrors.noncritical e -> (* [id] a goal variable *) gvar (loc,id) us, [], [], [] let find_appl_head_data c = - match c with - | GRef (loc,ref,_) as x -> + match Loc.obj c with + | GRef (ref,_) -> let impls = implicits_of_global ref in let scopes = find_arguments_scope ref in - x, impls, scopes, [] - | GApp (_,GRef (_,ref,_),l) as x + c, impls, scopes, [] + | GApp ((_, GRef (ref,_)),l) when l != [] && Flags.version_strictly_greater Flags.V8_2 -> let n = List.length l in let impls = implicits_of_global ref in let scopes = find_arguments_scope ref in - x, List.map (drop_first_implicits n) impls, + c, List.map (drop_first_implicits n) impls, List.skipn_at_least n scopes,[] - | x -> x,[],[],[] + | _ -> c,[],[],[] let error_not_enough_arguments loc = user_err ~loc (str "Abbreviation is not applied enough.") @@ -836,7 +839,7 @@ let intern_reference ref = (* Is it a global reference or a syntactic definition? *) let intern_qualid loc qid intern env lvar us args = match intern_extended_global_of_qualid (loc,qid) with - | TrueGlobal ref -> GRef (loc, ref, us), true, args + | TrueGlobal ref -> (Loc.tag ~loc @@ GRef (ref, us)), true, args | SynDef sp -> let (ids,c) = Syntax_def.search_syntactic_definition sp in let nids = List.length ids in @@ -850,9 +853,9 @@ let intern_qualid loc qid intern env lvar us args = let c = instantiate_notation_constr loc intern lvar subst infos c in let c = match us, c with | None, _ -> c - | Some _, GRef (loc, ref, None) -> GRef (loc, ref, us) - | Some _, GApp (loc, GRef (loc', ref, None), arg) -> - GApp (loc, GRef (loc', ref, us), arg) + | Some _, (loc, GRef (ref, None)) -> Loc.tag ~loc @@ GRef (ref, us) + | Some _, (loc, GApp ((loc', GRef (ref, None)), arg)) -> + Loc.tag ~loc @@ GApp (Loc.tag ~loc:loc' @@ GRef (ref, us), arg) | Some _, _ -> user_err ~loc (str "Notation " ++ pr_qualid qid ++ str " cannot have a universe instance, its expanded head @@ -863,7 +866,7 @@ let intern_qualid loc qid intern env lvar us args = (* Rule out section vars since these should have been found by intern_var *) let intern_non_secvar_qualid loc qid intern env lvar us args = match intern_qualid loc qid intern env lvar us args with - | GRef (_, VarRef _, _),_,_ -> raise Not_found + | (_, GRef (VarRef _, _)),_,_ -> raise Not_found | r -> r let intern_applied_reference intern env namedctx (_, ntnvars as lvar) us args = function @@ -1470,8 +1473,8 @@ let get_implicit_name n imps = Some (Impargs.name_of_implicit (List.nth imps (n-1))) let set_hole_implicit i b = function - | GRef (loc,r,_) | GApp (_,GRef (loc,r,_),_) -> (loc,Evar_kinds.ImplicitArg (r,i,b),Misctypes.IntroAnonymous,None) - | GVar (loc,id) -> (loc,Evar_kinds.ImplicitArg (VarRef id,i,b),Misctypes.IntroAnonymous,None) + | (loc, GRef (r,_)) | (_, GApp ((loc, (GRef (r,_))),_)) -> Loc.tag ~loc (Evar_kinds.ImplicitArg (r,i,b),Misctypes.IntroAnonymous,None) + | (loc, GVar id) -> Loc.tag ~loc (Evar_kinds.ImplicitArg (VarRef id,i,b),Misctypes.IntroAnonymous,None) | _ -> anomaly (Pp.str "Only refs have implicits") let exists_implicit_name id = @@ -1558,7 +1561,8 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = push_name_env ntnvars (impls_type_list ~args:fix_args tyi) en (Loc.ghost, Name name)) 0 env' lf in (a,b,c,intern {env'' with tmp_scope = None} bd)) dl idl_temp in - GRec (loc,GFix + Loc.tag ~loc @@ + GRec (GFix (Array.map (fun (ro,_,_,_) -> ro) idl,n), Array.of_list lf, Array.map (fun (_,bl,_,_) -> bl) idl, @@ -1584,7 +1588,8 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = push_name_env ntnvars (impls_type_list ~args:cofix_args tyi) en (Loc.ghost, Name name)) 0 env' lf in (b,c,intern {env'' with tmp_scope = None} bd)) dl idl_tmp in - GRec (loc,GCoFix n, + Loc.tag ~loc @@ + GRec (GCoFix n, Array.of_list lf, Array.map (fun (bl,_,_) -> bl) idl, Array.map (fun (_,ty,_) -> ty) idl, @@ -1600,7 +1605,8 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = | CLetIn (na,c1,t,c2) -> let inc1 = intern (reset_tmp_scope env) c1 in let int = Option.map (intern_type env) t in - GLetIn (loc, snd na, inc1, int, + Loc.tag ~loc @@ + GLetIn (snd na, inc1, int, intern (push_name_env ntnvars (impls_term_list inc1) env na) c2) | CNotation ("- _",([_, CPrim (Numeral p)],[],[])) when Bigint.is_strictly_pos p -> @@ -1622,7 +1628,8 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = lvar us args ref in (* Rem: GApp(_,f,[]) stands for @f *) - GApp (loc, f, intern_args env args_scopes (List.map fst args)) + Loc.tag ~loc @@ + GApp (f, intern_args env args_scopes (List.map fst args)) | CApp ((isproj,f), args) -> let f,args = match f with @@ -1687,20 +1694,21 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = (* Build a return predicate by expansion of the patterns of the "in" clause *) let thevars, thepats = List.split l in let sub_rtn = (* Some (GSort (Loc.ghost,GType None)) *) None in - let sub_tms = List.map (fun id -> GVar (Loc.ghost,id),(Name id,None)) thevars (* "match v1,..,vn" *) in - let main_sub_eqn = - (Loc.ghost,[],thepats, (* "|p1,..,pn" *) + let sub_tms = List.map (fun id -> (Loc.tag @@ GVar id),(Name id,None)) thevars (* "match v1,..,vn" *) in + let main_sub_eqn = Loc.tag @@ + ([],thepats, (* "|p1,..,pn" *) Option.cata (intern_type env') - (GHole(Loc.ghost,Evar_kinds.CasesType false,Misctypes.IntroAnonymous,None)) + (Loc.tag ~loc @@ GHole(Evar_kinds.CasesType false,Misctypes.IntroAnonymous,None)) rtnpo) (* "=> P" if there were a return predicate P, and "=> _" otherwise *) in let catch_all_sub_eqn = if List.for_all (irrefutable globalenv) thepats then [] else - [Loc.ghost,[],List.make (List.length thepats) (Loc.tag @@ PatVar Anonymous), (* "|_,..,_" *) - GHole(Loc.ghost,Evar_kinds.ImpossibleCase,Misctypes.IntroAnonymous,None)] (* "=> _" *) in - Some (GCases(Loc.ghost,Term.RegularStyle,sub_rtn,sub_tms,main_sub_eqn::catch_all_sub_eqn)) + [Loc.tag @@ ([],List.make (List.length thepats) (Loc.tag @@ PatVar Anonymous), (* "|_,..,_" *) + Loc.tag @@ GHole(Evar_kinds.ImpossibleCase,Misctypes.IntroAnonymous,None))] (* "=> _" *) in + Some (Loc.tag @@ GCases(Term.RegularStyle,sub_rtn,sub_tms,main_sub_eqn::catch_all_sub_eqn)) in let eqns' = List.map (intern_eqn (List.length tms) env) eqns in - GCases (loc, sty, rtnpo, tms, List.flatten eqns') + Loc.tag ~loc @@ + GCases (sty, rtnpo, tms, List.flatten eqns') | CLetTuple (nal, (na,po), b, c) -> let env' = reset_tmp_scope env in (* "in" is None so no match to add *) @@ -1709,7 +1717,8 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let env'' = push_name_env ntnvars (Variable,[],[],[]) (reset_hidden_inductive_implicit_test env') (Loc.ghost,na') in intern_type env'' u) po in - GLetTuple (loc, List.map snd nal, (na', p'), b', + Loc.tag ~loc @@ + GLetTuple (List.map snd nal, (na', p'), b', intern (List.fold_left (push_name_env ntnvars (Variable,[],[],[])) (reset_hidden_inductive_implicit_test env) nal) c) | CIf (c, (na,po), b1, b2) -> let env' = reset_tmp_scope env in @@ -1718,7 +1727,8 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let env'' = push_name_env ntnvars (Variable,[],[],[]) (reset_hidden_inductive_implicit_test env) (Loc.ghost,na') in intern_type env'' p) po in - GIf (loc, c', (na', p'), intern env b1, intern env b2) + Loc.tag ~loc @@ + GIf (c', (na', p'), intern env b1, intern env b2) | CHole (k, naming, solve) -> let k = match k with | None -> @@ -1743,23 +1753,29 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let (_, glb) = Genintern.generic_intern ist gen in Some glb in - GHole (loc, k, naming, solve) + Loc.tag ~loc @@ + GHole (k, naming, solve) (* Parsing pattern variables *) | CPatVar n when allow_patvar -> - GPatVar (loc, (true,n)) + Loc.tag ~loc @@ + GPatVar (true,n) | CEvar (n, []) when allow_patvar -> - GPatVar (loc, (false,n)) + Loc.tag ~loc @@ + GPatVar (false,n) (* end *) (* Parsing existential variables *) | CEvar (n, l) -> - GEvar (loc, n, List.map (on_snd (intern env)) l) + Loc.tag ~loc @@ + GEvar (n, List.map (on_snd (intern env)) l) | CPatVar _ -> raise (InternalizationError (loc,IllegalMetavariable)) (* end *) | CSort s -> - GSort(loc,s) + Loc.tag ~loc @@ + GSort s | CCast (c1, c2) -> - GCast (loc,intern env c1, Miscops.map_cast_type (intern_type env) c2) + Loc.tag ~loc @@ + GCast (intern env c1, Miscops.map_cast_type (intern_type env) c2) ) and intern_type env = intern (set_type_scope env) @@ -1790,15 +1806,15 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = List.map (fun (asubst,pl) -> let rhs = replace_vars_constr_expr asubst rhs in let rhs' = intern {env with ids = env_ids} rhs in - (loc,eqn_ids,pl,rhs')) pll + (loc,(eqn_ids,pl,rhs'))) pll and intern_case_item env forbidden_names_for_gen (tm,na,t) = (* the "match" part *) let tm' = intern env tm in (* the "as" part *) let extra_id,na = match tm', na with - | GVar (loc,id), None when not (Id.Map.mem id (snd lvar)) -> Some id,(loc,Name id) - | GRef (loc, VarRef id, _), None -> Some id,(loc,Name id) + | (loc , GVar id), None when not (Id.Map.mem id (snd lvar)) -> Some id,(loc,Name id) + | (loc, GRef (VarRef id, _)), None -> Some id,(loc,Name id) | _, None -> None,(Loc.ghost,Anonymous) | _, Some (loc,na) -> None,(loc,na) in (* the "in" part *) @@ -1870,8 +1886,9 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = (* with implicit arguments if maximal insertion is set *) [] else - GHole (set_hole_implicit (n,get_implicit_name n l) (force_inference_of imp) c) :: - aux (n+1) impl' subscopes' eargs rargs + (Loc.map (fun (a,b,c) -> GHole(a,b,c)) + (set_hole_implicit (n,get_implicit_name n l) (force_inference_of imp) c) + ) :: aux (n+1) impl' subscopes' eargs rargs end | (imp::impl', a::rargs') -> intern enva a :: aux (n+1) impl' subscopes' eargs rargs' @@ -1895,8 +1912,8 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = and smart_gapp f loc = function | [] -> f | l -> match f with - | GApp (loc', g, args) -> GApp (Loc.merge loc' loc, g, args@l) - | _ -> GApp (Loc.merge (loc_of_glob_constr f) loc, f, l) + | (loc', GApp (g, args)) -> Loc.tag ~loc:(Loc.merge loc' loc) @@ GApp (g, args@l) + | _ -> Loc.tag ~loc:(Loc.merge (loc_of_glob_constr f) loc) @@ GApp (f, l) and intern_args env subscopes = function | [] -> [] diff --git a/interp/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml index d2bebfb54..51152bb24 100644 --- a/interp/implicit_quantifiers.ml +++ b/interp/implicit_quantifiers.ml @@ -125,37 +125,38 @@ let add_name_to_ids set na = | Name id -> Id.Set.add id set let generalizable_vars_of_glob_constr ?(bound=Id.Set.empty) ?(allowed=Id.Set.empty) = - let rec vars bound vs = function - | GVar (loc,id) -> + let rec vars bound vs (loc, t) = match t with + | GVar id -> if is_freevar bound (Global.env ()) id then if Id.List.mem_assoc id vs then vs else (id, loc) :: vs else vs - | GApp (loc,f,args) -> List.fold_left (vars bound) vs (f::args) - | GLambda (loc,na,_,ty,c) | GProd (loc,na,_,ty,c) -> + + | GApp (f,args) -> List.fold_left (vars bound) vs (f::args) + | GLambda (na,_,ty,c) | GProd (na,_,ty,c) -> let vs' = vars bound vs ty in let bound' = add_name_to_ids bound na in vars bound' vs' c - | GLetIn (loc,na,b,ty,c) -> + | GLetIn (na,b,ty,c) -> let vs' = vars bound vs b in let vs'' = Option.fold_left (vars bound) vs' ty in let bound' = add_name_to_ids bound na in vars bound' vs'' c - | GCases (loc,sty,rtntypopt,tml,pl) -> + | GCases (sty,rtntypopt,tml,pl) -> let vs1 = vars_option bound vs rtntypopt in let vs2 = List.fold_left (fun vs (tm,_) -> vars bound vs tm) vs1 tml in List.fold_left (vars_pattern bound) vs2 pl - | GLetTuple (loc,nal,rtntyp,b,c) -> + | GLetTuple (nal,rtntyp,b,c) -> let vs1 = vars_return_type bound vs rtntyp in let vs2 = vars bound vs1 b in let bound' = List.fold_left add_name_to_ids bound nal in vars bound' vs2 c - | GIf (loc,c,rtntyp,b1,b2) -> + | GIf (c,rtntyp,b1,b2) -> let vs1 = vars_return_type bound vs rtntyp in let vs2 = vars bound vs1 c in let vs3 = vars bound vs2 b1 in vars bound vs3 b2 - | GRec (loc,fk,idl,bl,tyl,bv) -> + | GRec (fk,idl,bl,tyl,bv) -> let bound' = Array.fold_right Id.Set.add idl bound in let vars_fix i vs fid = let vs1,bound1 = @@ -173,11 +174,11 @@ let generalizable_vars_of_glob_constr ?(bound=Id.Set.empty) ?(allowed=Id.Set.emp vars bound1 vs2 bv.(i) in Array.fold_left_i vars_fix vs idl - | GCast (loc,c,k) -> let v = vars bound vs c in + | GCast (c,k) -> let v = vars bound vs c in (match k with CastConv t | CastVM t -> vars bound v t | _ -> v) | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> vs - and vars_pattern bound vs (loc,idl,p,c) = + and vars_pattern bound vs (loc,(idl,p,c)) = let bound' = List.fold_right Id.Set.add idl bound in vars bound' vs c @@ -309,12 +310,12 @@ let implicits_of_glob_constr ?(with_products=true) l = (ExplByPos (i, name), (true, true, true)) :: l | _ -> l in - let rec aux i c = + let rec aux i (loc, c) = let abs na bk b = add_impl i na bk (aux (succ i) b) in match c with - | GProd (loc, na, bk, t, b) -> + | GProd (na, bk, t, b) -> if with_products then abs na bk b else let () = match bk with @@ -323,9 +324,9 @@ let implicits_of_glob_constr ?(with_products=true) l = pr_name na ++ strbrk " and following binders") | _ -> () in [] - | GLambda (loc, na, bk, t, b) -> abs na bk b - | GLetIn (loc, na, b, t, c) -> aux i c - | GRec (_, fix_kind, nas, args, tys, bds) -> + | GLambda (na, bk, t, b) -> abs na bk b + | GLetIn (na, b, t, c) -> aux i b + | GRec (fix_kind, nas, args, tys, bds) -> let nb = match fix_kind with |GFix (_, n) -> n | GCoFix n -> n in List.fold_left_i (fun i l (na,bk,_,_) -> add_impl i na bk l) i (aux (List.length args.(nb) + i) bds.(nb)) args.(nb) | _ -> [] diff --git a/interp/notation.ml b/interp/notation.ml index aef089299..3bcec3001 100644 --- a/interp/notation.ml +++ b/interp/notation.ml @@ -264,12 +264,12 @@ let notations_key_table = ref (KeyMap.empty : notation_rule list KeyMap.t) let prim_token_key_table = ref KeyMap.empty let glob_prim_constr_key = function - | GApp (_,GRef (_,ref,_),_) | GRef (_,ref,_) -> RefKey (canonical_gr ref) + | _, GApp ((_, GRef (ref,_)),_) | _, GRef (ref,_) -> RefKey (canonical_gr ref) | _ -> Oth let glob_constr_keys = function - | GApp (_,GRef (_,ref,_),_) -> [RefKey (canonical_gr ref); Oth] - | GRef (_,ref,_) -> [RefKey (canonical_gr ref)] + | _, GApp ((_, GRef (ref,_)),_) -> [RefKey (canonical_gr ref); Oth] + | _, GRef (ref,_) -> [RefKey (canonical_gr ref)] | _ -> [Oth] let cases_pattern_key = function @@ -471,13 +471,14 @@ let interp_prim_token = (** [rcp_of_glob] : from [glob_constr] to [raw_cases_pattern_expr] *) -let rec rcp_of_glob looked_for = function - | GVar (loc,id) -> Loc.tag ~loc @@ RCPatAtom (Some id) - | GHole (loc,_,_,_) -> Loc.tag ~loc @@ RCPatAtom (None) - | GRef (loc,g,_) -> looked_for g; Loc.tag ~loc @@ RCPatCstr (g,[],[]) - | GApp (loc,GRef (_,g,_),l) -> - looked_for g; Loc.tag ~loc @@ RCPatCstr (g, List.map (rcp_of_glob looked_for) l,[]) +let rec rcp_of_glob looked_for gt = Loc.map (function + | GVar id -> RCPatAtom (Some id) + | GHole (_,_,_) -> RCPatAtom None + | GRef (g,_) -> looked_for g; RCPatCstr (g,[],[]) + | GApp ((_, GRef (g,_)),l) -> + looked_for g; RCPatCstr (g, List.map (rcp_of_glob looked_for) l,[]) | _ -> raise Not_found + ) gt let interp_prim_token_cases_pattern_expr loc looked_for p = interp_prim_token_gen (rcp_of_glob looked_for) loc p @@ -521,8 +522,8 @@ let uninterp_prim_token_ind_pattern ind args = if not b then raise Notation_ops.No_match; let args' = List.map (fun x -> snd (glob_constr_of_closed_cases_pattern x)) args in - let ref = GRef (Loc.ghost,ref,None) in - match numpr (GApp (Loc.ghost,ref,args')) with + let ref = Loc.tag @@ GRef (ref,None) in + match numpr (Loc.tag @@ GApp (ref,args')) with | None -> raise Notation_ops.No_match | Some n -> (sc,n) with Not_found -> raise Notation_ops.No_match diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml index a25fd81f3..32c900504 100644 --- a/interp/notation_ops.ml +++ b/interp/notation_ops.ml @@ -24,19 +24,19 @@ open Notation_term let on_true_do b f c = if b then (f c; b) else b -let compare_glob_constr f add t1 t2 = match t1,t2 with - | GRef (_,r1,_), GRef (_,r2,_) -> eq_gr r1 r2 - | GVar (_,v1), GVar (_,v2) -> on_true_do (Id.equal v1 v2) add (Name v1) - | GApp (_,f1,l1), GApp (_,f2,l2) -> f f1 f2 && List.for_all2eq f l1 l2 - | GLambda (_,na1,bk1,ty1,c1), GLambda (_,na2,bk2,ty2,c2) +let compare_glob_constr f add (_l1, t1) (_l2, t2) = match t1,t2 with + | GRef (r1,_), GRef (r2,_) -> eq_gr r1 r2 + | GVar v1, GVar v2 -> on_true_do (Id.equal v1 v2) add (Name v1) + | GApp (f1,l1), GApp (f2,l2) -> f f1 f2 && List.for_all2eq f l1 l2 + | GLambda (na1,bk1,ty1,c1), GLambda (na2,bk2,ty2,c2) when Name.equal na1 na2 && Constrexpr_ops.binding_kind_eq bk1 bk2 -> on_true_do (f ty1 ty2 && f c1 c2) add na1 - | GProd (_,na1,bk1,ty1,c1), GProd (_,na2,bk2,ty2,c2) + | GProd (na1,bk1,ty1,c1), GProd (na2,bk2,ty2,c2) when Name.equal na1 na2 && Constrexpr_ops.binding_kind_eq bk1 bk2 -> on_true_do (f ty1 ty2 && f c1 c2) add na1 | GHole _, GHole _ -> true - | GSort (_,s1), GSort (_,s2) -> Miscops.glob_sort_eq s1 s2 - | GLetIn (_,na1,b1,t1,c1), GLetIn (_,na2,b2,t2,c2) when Name.equal na1 na2 -> + | GSort s1, GSort s2 -> Miscops.glob_sort_eq s1 s2 + | GLetIn (na1,b1,t1,c1), GLetIn (na2,b2,t2,c2) when Name.equal na1 na2 -> on_true_do (f b1 b2 && f c1 c2) add na1 | (GCases _ | GRec _ | GPatVar _ | GEvar _ | GLetTuple _ | GIf _ | GCast _),_ @@ -129,49 +129,51 @@ let rec cases_pattern_fold_map loc g e = Loc.with_unloc (function let subst_binder_type_vars l = function | Evar_kinds.BinderType (Name id) -> let id = - try match Id.List.assoc id l with GVar(_,id') -> id' | _ -> id + try match snd @@ Id.List.assoc id l with GVar id' -> id' | _ -> id with Not_found -> id in Evar_kinds.BinderType (Name id) | e -> e -let rec subst_glob_vars l = function - | GVar (_,id) as r -> (try Id.List.assoc id l with Not_found -> r) - | GProd (loc,Name id,bk,t,c) -> +let rec subst_glob_vars l gc = Loc.map (function + | GVar id as r -> (try snd @@ Id.List.assoc id l with Not_found -> r) + | GProd (Name id,bk,t,c) -> let id = - try match Id.List.assoc id l with GVar(_,id') -> id' | _ -> id + try match snd @@ Id.List.assoc id l with GVar id' -> id' | _ -> id with Not_found -> id in - GProd (loc,Name id,bk,subst_glob_vars l t,subst_glob_vars l c) - | GLambda (loc,Name id,bk,t,c) -> + GProd (Name id,bk,subst_glob_vars l t,subst_glob_vars l c) + | GLambda (Name id,bk,t,c) -> let id = - try match Id.List.assoc id l with GVar(_,id') -> id' | _ -> id + try match snd @@ Id.List.assoc id l with GVar id' -> id' | _ -> id with Not_found -> id in - GLambda (loc,Name id,bk,subst_glob_vars l t,subst_glob_vars l c) - | GHole (loc,x,naming,arg) -> GHole (loc,subst_binder_type_vars l x,naming,arg) - | r -> map_glob_constr (subst_glob_vars l) r (* assume: id is not binding *) + GLambda (Name id,bk,subst_glob_vars l t,subst_glob_vars l c) + | GHole (x,naming,arg) -> GHole (subst_binder_type_vars l x,naming,arg) + | _ -> snd @@ map_glob_constr (subst_glob_vars l) gc (* assume: id is not binding *) + ) gc let ldots_var = Id.of_string ".." -let glob_constr_of_notation_constr_with_binders loc g f e = function - | NVar id -> GVar (loc,id) - | NApp (a,args) -> GApp (loc,f e a, List.map (f e) args) +let glob_constr_of_notation_constr_with_binders loc g f e nc = + let lt x = Loc.tag ~loc x in lt @@ match nc with + | NVar id -> GVar id + | NApp (a,args) -> GApp (f e a, List.map (f e) args) | NList (x,y,iter,tail,swap) -> let t = f e tail in let it = f e iter in - let innerl = (ldots_var,t)::(if swap then [] else [x,GVar(loc,y)]) in - let inner = GApp (loc,GVar (loc,ldots_var),[subst_glob_vars innerl it]) in - let outerl = (ldots_var,inner)::(if swap then [x,GVar(loc,y)] else []) in - subst_glob_vars outerl it + let innerl = (ldots_var,t)::(if swap then [] else [x, lt @@ GVar y]) in + let inner = lt @@ GApp (lt @@ GVar (ldots_var),[subst_glob_vars innerl it]) in + let outerl = (ldots_var,inner)::(if swap then [x, lt @@ GVar y] else []) in + Loc.obj @@ subst_glob_vars outerl it | NBinderList (x,y,iter,tail) -> let t = f e tail in let it = f e iter in - let innerl = [(ldots_var,t);(x,GVar(loc,y))] in - let inner = GApp (loc,GVar (loc,ldots_var),[subst_glob_vars innerl it]) in + let innerl = [(ldots_var,t);(x, lt @@ GVar y)] in + let inner = lt @@ GApp (lt @@ GVar ldots_var,[subst_glob_vars innerl it]) in let outerl = [(ldots_var,inner)] in - subst_glob_vars outerl it + Loc.obj @@ subst_glob_vars outerl it | NLambda (na,ty,c) -> - let e',na = g e na in GLambda (loc,na,Explicit,f e ty,f e' c) + let e',na = g e na in GLambda (na,Explicit,f e ty,f e' c) | NProd (na,ty,c) -> - let e',na = g e na in GProd (loc,na,Explicit,f e ty,f e' c) + let e',na = g e na in GProd (na,Explicit,f e ty,f e' c) | NLetIn (na,b,t,c) -> - let e',na = g e na in GLetIn (loc,na,f e b,Option.map (f e) t,f e' c) + let e',na = g e na in GLetIn (na,f e b,Option.map (f e) t,f e' c) | NCases (sty,rtntypopt,tml,eqnl) -> let e',tml' = List.fold_right (fun (tm,(na,t)) (e',tml') -> let e',t' = match t with @@ -186,25 +188,25 @@ let glob_constr_of_notation_constr_with_binders loc g f e = function let eqnl' = List.map (fun (patl,rhs) -> let ((idl,e),patl) = List.fold_map (cases_pattern_fold_map loc fold) ([],e) patl in - (loc,idl,patl,f e rhs)) eqnl in - GCases (loc,sty,Option.map (f e') rtntypopt,tml',eqnl') + lt (idl,patl,f e rhs)) eqnl in + GCases (sty,Option.map (f e') rtntypopt,tml',eqnl') | NLetTuple (nal,(na,po),b,c) -> let e',nal = List.fold_map g e nal in let e'',na = g e na in - GLetTuple (loc,nal,(na,Option.map (f e'') po),f e b,f e' c) + GLetTuple (nal,(na,Option.map (f e'') po),f e b,f e' c) | NIf (c,(na,po),b1,b2) -> let e',na = g e na in - GIf (loc,f e c,(na,Option.map (f e') po),f e b1,f e b2) + GIf (f e c,(na,Option.map (f e') po),f e b1,f e b2) | NRec (fk,idl,dll,tl,bl) -> let e,dll = Array.fold_map (List.fold_map (fun e (na,oc,b) -> let e,na = g e na in (e,(na,Explicit,Option.map (f e) oc,f e b)))) e dll in let e',idl = Array.fold_map (to_id g) e idl in - GRec (loc,fk,idl,dll,Array.map (f e) tl,Array.map (f e') bl) - | NCast (c,k) -> GCast (loc,f e c,Miscops.map_cast_type (f e) k) - | NSort x -> GSort (loc,x) - | NHole (x, naming, arg) -> GHole (loc, x, naming, arg) - | NRef x -> GRef (loc,x,None) + GRec (fk,idl,dll,Array.map (f e) tl,Array.map (f e') bl) + | NCast (c,k) -> GCast (f e c,Miscops.map_cast_type (f e) k) + | NSort x -> GSort x + | NHole (x, naming, arg) -> GHole (x, naming, arg) + | NRef x -> GRef (x,None) let glob_constr_of_notation_constr loc x = let rec aux () x = @@ -220,13 +222,13 @@ let add_name r = function Anonymous -> () | Name id -> add_id r id let split_at_recursive_part c = let sub = ref None in let rec aux = function - | GApp (loc0,GVar(loc,v),c::l) when Id.equal v ldots_var -> + | loc0, GApp ((loc,GVar v),c::l) when Id.equal v ldots_var -> begin match !sub with | None -> let () = sub := Some c in begin match l with - | [] -> GVar (loc, ldots_var) - | _ :: _ -> GApp (loc0, GVar (loc, ldots_var), l) + | [] -> Loc.tag ~loc @@ GVar ldots_var + | _ :: _ -> Loc.tag ~loc:loc0 @@ GApp (Loc.tag ~loc @@ GVar ldots_var, l) end | Some _ -> (* Not narrowed enough to find only one recursive part *) @@ -237,13 +239,13 @@ let split_at_recursive_part c = match !sub with | None -> (* No recursive pattern found *) raise Not_found | Some c -> - match outer_iterator with - | GVar (_,v) when Id.equal v ldots_var -> (* Not enough context *) raise Not_found + match Loc.obj outer_iterator with + | GVar v when Id.equal v ldots_var -> (* Not enough context *) raise Not_found | _ -> outer_iterator, c let subtract_loc loc1 loc2 = Loc.make_loc (fst (Loc.unloc loc1),fst (Loc.unloc loc2)-1) -let check_is_hole id = function GHole _ -> () | t -> +let check_is_hole id = function _, GHole _ -> () | t -> user_err ~loc:(loc_of_glob_constr t) (strbrk "In recursive notation with binders, " ++ pr_id id ++ strbrk " is expected to come without type.") @@ -257,19 +259,19 @@ type recursive_pattern_kind = let compare_recursive_parts found f f' (iterator,subc) = let diff = ref None in let terminator = ref None in - let rec aux c1 c2 = match c1,c2 with - | GVar(_,v), term when Id.equal v ldots_var -> + let rec aux (l1, c1) (l2, c2) = match c1, c2 with + | GVar v, term when Id.equal v ldots_var -> (* We found the pattern *) assert (match !terminator with None -> true | Some _ -> false); - terminator := Some term; + terminator := Some (l2, term); true - | GApp (_,GVar(_,v),l1), GApp (_,term,l2) when Id.equal v ldots_var -> + | GApp ((_, GVar v),l1), GApp (term, l2) when Id.equal v ldots_var -> (* We found the pattern, but there are extra arguments *) (* (this allows e.g. alternative (recursive) notation of application) *) assert (match !terminator with None -> true | Some _ -> false); terminator := Some term; List.for_all2eq aux l1 l2 - | GVar (_,x), GVar (_,y) when not (Id.equal x y) -> + | GVar x, GVar y when not (Id.equal x y) -> (* We found the position where it differs *) let lassoc = match !terminator with None -> false | Some _ -> true in let x,y = if lassoc then y,x else x,y in @@ -279,8 +281,8 @@ let compare_recursive_parts found f f' (iterator,subc) = true | Some _ -> false end - | GLambda (_,Name x,_,t_x,c), GLambda (_,Name y,_,t_y,term) - | GProd (_,Name x,_,t_x,c), GProd (_,Name y,_,t_y,term) -> + | GLambda (Name x,_,t_x,c), GLambda (Name y,_,t_y,term) + | GProd (Name x,_,t_x,c), GProd (Name y,_,t_y,term) -> (* We found a binding position where it differs *) begin match !diff with | None -> @@ -289,7 +291,7 @@ let compare_recursive_parts found f f' (iterator,subc) = | Some _ -> false end | _ -> - compare_glob_constr aux (add_name found) c1 c2 in + compare_glob_constr aux (add_name found) (l1, c1) (l2, c2) in if aux iterator subc then match !diff with | None -> @@ -312,13 +314,13 @@ let compare_recursive_parts found f f' (iterator,subc) = (pi1 !found, (x,y) :: pi2 !found, pi3 !found),x,y,lassoc in let iterator = f' (if lassoc then iterator - else subst_glob_vars [x,GVar(Loc.ghost,y)] iterator) in + else subst_glob_vars [x, Loc.tag @@ GVar y] iterator) in (* found have been collected by compare_constr *) found := newfound; NList (x,y,iterator,f (Option.get !terminator),lassoc) | Some (x,y,RecursiveBinders (t_x,t_y)) -> let newfound = (pi1 !found, pi2 !found, (x,y) :: pi3 !found) in - let iterator = f' (subst_glob_vars [x,GVar(Loc.ghost,y)] iterator) in + let iterator = f' (subst_glob_vars [x, Loc.tag @@ GVar y] iterator) in (* found have been collected by compare_constr *) found := newfound; check_is_hole x t_x; @@ -336,22 +338,22 @@ let notation_constr_and_vars_of_glob_constr a = try compare_recursive_parts found aux aux' (split_at_recursive_part c) with Not_found -> found := keepfound; - match c with - | GApp (_,GVar (loc,f),[c]) when Id.equal f ldots_var -> + match snd c with + | GApp ((loc, GVar f),[c]) when Id.equal f ldots_var -> (* Fall on the second part of the recursive pattern w/o having found the first part *) user_err ~loc (str "Cannot find where the recursive pattern starts.") - | c -> + | _c -> aux' c - and aux' = function - | GVar (_,id) -> add_id found id; NVar id - | GApp (_,g,args) -> NApp (aux g, List.map aux args) - | GLambda (_,na,bk,ty,c) -> add_name found na; NLambda (na,aux ty,aux c) - | GProd (_,na,bk,ty,c) -> add_name found na; NProd (na,aux ty,aux c) - | GLetIn (_,na,b,t,c) -> add_name found na; NLetIn (na,aux b,Option.map aux t,aux c) - | GCases (_,sty,rtntypopt,tml,eqnl) -> - let f (_,idl,pat,rhs) = List.iter (add_id found) idl; (pat,aux rhs) in + and aux' x = Loc.with_unloc (function + | GVar id -> add_id found id; NVar id + | GApp (g,args) -> NApp (aux g, List.map aux args) + | GLambda (na,bk,ty,c) -> add_name found na; NLambda (na,aux ty,aux c) + | GProd (na,bk,ty,c) -> add_name found na; NProd (na,aux ty,aux c) + | GLetIn (na,b,t,c) -> add_name found na; NLetIn (na,aux b,Option.map aux t, aux c) + | GCases (sty,rtntypopt,tml,eqnl) -> + let f (_,(idl,pat,rhs)) = List.iter (add_id found) idl; (pat,aux rhs) in NCases (sty,Option.map aux rtntypopt, List.map (fun (tm,(na,x)) -> add_name found na; @@ -359,29 +361,29 @@ let notation_constr_and_vars_of_glob_constr a = (fun (_,(_,nl)) -> List.iter (add_name found) nl) x; (aux tm,(na,Option.map (fun (_,(ind,nal)) -> (ind,nal)) x))) tml, List.map f eqnl) - | GLetTuple (loc,nal,(na,po),b,c) -> + | GLetTuple (nal,(na,po),b,c) -> add_name found na; List.iter (add_name found) nal; NLetTuple (nal,(na,Option.map aux po),aux b,aux c) - | GIf (loc,c,(na,po),b1,b2) -> + | GIf (c,(na,po),b1,b2) -> add_name found na; NIf (aux c,(na,Option.map aux po),aux b1,aux b2) - | GRec (_,fk,idl,dll,tl,bl) -> + | GRec (fk,idl,dll,tl,bl) -> Array.iter (add_id found) idl; let dll = Array.map (List.map (fun (na,bk,oc,b) -> if bk != Explicit then error "Binders marked as implicit not allowed in notations."; add_name found na; (na,Option.map aux oc,aux b))) dll in NRec (fk,idl,dll,Array.map aux tl,Array.map aux bl) - | GCast (_,c,k) -> NCast (aux c,Miscops.map_cast_type aux k) - | GSort (_,s) -> NSort s - | GHole (_,w,naming,arg) -> + | GCast (c,k) -> NCast (aux c,Miscops.map_cast_type aux k) + | GSort s -> NSort s + | GHole (w,naming,arg) -> if arg != None then has_ltac := true; NHole (w, naming, arg) - | GRef (_,r,_) -> NRef r + | GRef (r,_) -> NRef r | GEvar _ | GPatVar _ -> error "Existential variables not allowed in notations." - + ) x in let t = aux a in (* Side effect *) @@ -590,8 +592,8 @@ let abstract_return_type_context pi mklam tml rtno = let abstract_return_type_context_glob_constr = abstract_return_type_context (fun (_,(_,nal)) -> nal) - (fun na c -> - GLambda(Loc.ghost,na,Explicit,GHole(Loc.ghost,Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None),c)) + (fun na c -> Loc.tag @@ + GLambda(na,Explicit,Loc.tag @@ GHole(Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None),c)) let abstract_return_type_context_notation_constr = abstract_return_type_context snd @@ -663,18 +665,19 @@ let add_binding_env alp (terms,onlybinders,termlists,binderlists) var v = let add_bindinglist_env (terms,onlybinders,termlists,binderlists) x bl = (terms,onlybinders,termlists,(x,bl)::binderlists) -let rec pat_binder_of_term = function - | GVar (loc, id) -> Loc.tag ~loc @@ PatVar (Name id) - | GApp (loc, GRef (_,ConstructRef cstr,_), l) -> +let rec pat_binder_of_term t = Loc.map (function + | GVar id -> PatVar (Name id) + | GApp ((_, GRef (ConstructRef cstr,_)), l) -> let nparams = Inductiveops.inductive_nparams (fst cstr) in let _,l = List.chop nparams l in - Loc.tag ~loc @@ PatCstr (cstr, List.map pat_binder_of_term l, Anonymous) + PatCstr (cstr, List.map pat_binder_of_term l, Anonymous) | _ -> raise No_match + ) t let bind_term_env alp (terms,onlybinders,termlists,binderlists as sigma) var v = try let v' = Id.List.assoc var terms in - match v, v' with + match Loc.obj v, Loc.obj v' with | GHole _, _ -> sigma | _, GHole _ -> let sigma = Id.List.remove_assoc var terms,onlybinders,termlists,binderlists in @@ -688,7 +691,7 @@ let bind_termlist_env alp (terms,onlybinders,termlists,binderlists as sigma) var try let vl' = Id.List.assoc var termlists in let unify_term v v' = - match v, v' with + match Loc.obj v, Loc.obj v' with | GHole _, _ -> v' | _, GHole _ -> v | _, _ -> if glob_constr_eq (alpha_rename (snd alp) v) v' then v' else raise No_match in @@ -704,8 +707,8 @@ let bind_termlist_env alp (terms,onlybinders,termlists,binderlists as sigma) var let bind_term_as_binding_env alp (terms,onlybinders,termlists,binderlists as sigma) var id = try - match Id.List.assoc var terms with - | GVar (_,id') -> + match Loc.obj @@ Id.List.assoc var terms with + | GVar id' -> (if not (Id.equal id id') then (fst alp,(id,id')::snd alp) else alp), sigma | _ -> anomaly (str "A term which can be a binder has to be a variable") @@ -713,7 +716,7 @@ let bind_term_as_binding_env alp (terms,onlybinders,termlists,binderlists as sig (* The matching against a term allowing to find the instance has not been found yet *) (* If it will be a different name, we shall unfortunately fail *) (* TODO: look at the consequences for alp *) - alp, add_env alp sigma var (GVar (Loc.ghost,id)) + alp, add_env alp sigma var (Loc.tag @@ GVar id) let bind_binding_as_term_env alp (terms,onlybinders,termlists,binderlists as sigma) var id = try @@ -782,7 +785,7 @@ let bind_bindinglist_env alp (terms,onlybinders,termlists,binderlists as sigma) let unify_pat alp p p' = try fold_cases_pattern_eq unify_name alp p p' with Failure _ -> raise No_match in let unify_term alp v v' = - match v, v' with + match Loc.obj v, Loc.obj v' with | GHole _, _ -> v' | _, GHole _ -> v | _, _ -> if glob_constr_eq (alpha_rename (snd alp) v) v' then v else raise No_match in @@ -831,7 +834,7 @@ let bind_bindinglist_as_term_env alp (terms,onlybinders,termlists,binderlists) v else raise No_match in let unify_term_binder c b' = match c, b' with - | GVar (loc, id), GLocalAssum (_, na', bk', t') -> + | (_, GVar id), GLocalAssum (loc, na', bk', t') -> GLocalAssum (loc, unify_id id na', bk', t') | c, GLocalPattern (loc, (p',ids), id, bk', t') -> let p = pat_binder_of_term c in @@ -892,21 +895,22 @@ let rec match_cases_pattern_binders metas acc (_, pat1) (_, pat2) = let glue_letin_with_decls = true -let rec match_iterated_binders islambda decls = function - | GLambda (loc,Name p,bk,t,GCases (_,LetPatternStyle,None,[(GVar(_,e),_)],[(_,ids,[cp],b)])) +let rec match_iterated_binders islambda decls bi = Loc.with_loc (fun ~loc -> function + | GLambda (Name p,bk,t,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b))]))) when islambda && Id.equal p e -> - match_iterated_binders islambda (GLocalPattern (loc,(cp,ids),p,bk,t)::decls) b - | GLambda (loc,na,bk,t,b) when islambda -> - match_iterated_binders islambda (GLocalAssum (loc,na,bk,t)::decls) b - | GProd (loc,Name p,bk,t,GCases (_,LetPatternStyle,None,[(GVar(_,e),_)],[(_,ids,[cp],b)])) + match_iterated_binders islambda (GLocalPattern(loc,(cp,ids),p,bk,t)::decls) b + | GLambda (na,bk,t,b) when islambda -> + match_iterated_binders islambda (GLocalAssum(loc,na,bk,t)::decls) b + | GProd (Name p,bk,t,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b))]))) when not islambda && Id.equal p e -> - match_iterated_binders islambda (GLocalPattern (loc,(cp,ids),p,bk,t)::decls) b - | GProd (loc,(Name _ as na),bk,t,b) when not islambda -> - match_iterated_binders islambda (GLocalAssum (loc,na,bk,t)::decls) b - | GLetIn (loc,na,c,t,b) when glue_letin_with_decls -> + match_iterated_binders islambda (GLocalPattern(loc,(cp,ids),p,bk,t)::decls) b + | GProd ((Name _ as na),bk,t,b) when not islambda -> + match_iterated_binders islambda (GLocalAssum(loc,na,bk,t)::decls) b + | GLetIn (na,c,t,b) when glue_letin_with_decls -> match_iterated_binders islambda (GLocalDef (loc,na,Explicit (*?*), c,t)::decls) b - | b -> (decls,b) + | b -> (decls, Loc.tag ~loc b) + ) bi let remove_sigma x (terms,onlybinders,termlists,binderlists) = (Id.List.remove_assoc x terms,onlybinders,termlists,binderlists) @@ -967,91 +971,92 @@ let does_not_come_from_already_eta_expanded_var = (* The following test is then an approximation of what can be done *) (* optimally (whether other looping situations can occur remains to be *) (* checked). *) - function GVar _ -> false | _ -> true + function _loc, GVar _ -> false | _ -> true let rec match_ inner u alp metas sigma a1 a2 = - match (a1,a2) with + let loc, a1_val = Loc.to_pair a1 in + match a1_val, a2 with (* Matching notation variable *) - | r1, NVar id2 when is_term_meta id2 metas -> bind_term_env alp sigma id2 r1 - | GVar (_,id1), NVar id2 when is_onlybinding_meta id2 metas -> bind_binding_as_term_env alp sigma id2 id1 - | r1, NVar id2 when is_bindinglist_meta id2 metas -> bind_term_env alp sigma id2 r1 + | r1, NVar id2 when is_term_meta id2 metas -> bind_term_env alp sigma id2 a1 + | GVar id1, NVar id2 when is_onlybinding_meta id2 metas -> bind_binding_as_term_env alp sigma id2 id1 + | r1, NVar id2 when is_bindinglist_meta id2 metas -> bind_term_env alp sigma id2 a1 (* Matching recursive notations for terms *) | r1, NList (x,y,iter,termin,lassoc) -> - match_termlist (match_hd u alp) alp metas sigma r1 x y iter termin lassoc + match_termlist (match_hd u alp) alp metas sigma a1 x y iter termin lassoc (* "λ p, let 'cp = p in t" -> "λ 'cp, t" *) - | GLambda (loc,Name p,bk,t1,GCases (_,LetPatternStyle,None,[(GVar(_,e),_)],[(_,ids,[cp],b1)])), + | GLambda (Name p,bk,t1,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b1))]))), NBinderList (x,_,NLambda (Name _id2,_,b2),termin) when Id.equal p e -> let (decls,b) = match_iterated_binders true [GLocalPattern(loc,(cp,ids),p,bk,t1)] b1 in let alp,sigma = bind_bindinglist_env alp sigma x decls in match_in u alp metas sigma b termin (* Matching recursive notations for binders: ad hoc cases supporting let-in *) - | GLambda (loc,na1,bk,t1,b1), NBinderList (x,_,NLambda (Name _id2,_,b2),termin)-> + | GLambda (na1,bk,t1,b1), NBinderList (x,_,NLambda (Name _id2,_,b2),termin)-> let (decls,b) = match_iterated_binders true [GLocalAssum (loc,na1,bk,t1)] b1 in (* TODO: address the possibility that termin is a Lambda itself *) let alp,sigma = bind_bindinglist_env alp sigma x decls in match_in u alp metas sigma b termin (* "∀ p, let 'cp = p in t" -> "∀ 'cp, t" *) - | GProd (loc,Name p,bk,t1,GCases (_,LetPatternStyle,None,[(GVar(_,e),_)],[(_,ids,[cp],b1)])), + | GProd (Name p,bk,t1,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b1))]))), NBinderList (x,_,NProd (Name _id2,_,b2),(NVar v as termin)) when Id.equal p e -> let (decls,b) = match_iterated_binders true [GLocalPattern (loc,(cp,ids),p,bk,t1)] b1 in let alp,sigma = bind_bindinglist_env alp sigma x decls in match_in u alp metas sigma b termin - | GProd (loc,na1,bk,t1,b1), NBinderList (x,_,NProd (Name _id2,_,b2),termin) + | GProd (na1,bk,t1,b1), NBinderList (x,_,NProd (Name _id2,_,b2),termin) when na1 != Anonymous -> let (decls,b) = match_iterated_binders false [GLocalAssum (loc,na1,bk,t1)] b1 in (* TODO: address the possibility that termin is a Prod itself *) let alp,sigma = bind_bindinglist_env alp sigma x decls in match_in u alp metas sigma b termin (* Matching recursive notations for binders: general case *) - | r, NBinderList (x,y,iter,termin) -> - match_binderlist_with_app (match_hd u) alp metas sigma r x y iter termin + | _r, NBinderList (x,y,iter,termin) -> + match_binderlist_with_app (match_hd u) alp metas sigma a1 x y iter termin (* Matching individual binders as part of a recursive pattern *) - | GLambda (loc,Name p,bk,t,GCases (_,LetPatternStyle,None,[(GVar(_,e),_)],[(_,ids,[cp],b1)])), + | GLambda (Name p,bk,t,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b1))]))), NLambda (Name id,_,b2) when is_bindinglist_meta id metas -> let alp,sigma = bind_bindinglist_env alp sigma id [GLocalPattern (loc,(cp,ids),p,bk,t)] in match_in u alp metas sigma b1 b2 - | GLambda (loc,na,bk,t,b1), NLambda (Name id,_,b2) + | GLambda (na,bk,t,b1), NLambda (Name id,_,b2) when is_bindinglist_meta id metas -> let alp,sigma = bind_bindinglist_env alp sigma id [GLocalAssum (loc,na,bk,t)] in match_in u alp metas sigma b1 b2 - | GProd (loc,na,bk,t,b1), NProd (Name id,_,b2) + | GProd (na,bk,t,b1), NProd (Name id,_,b2) when is_bindinglist_meta id metas && na != Anonymous -> let alp,sigma = bind_bindinglist_env alp sigma id [GLocalAssum (loc,na,bk,t)] in match_in u alp metas sigma b1 b2 (* Matching compositionally *) - | GVar (_,id1), NVar id2 when alpha_var id1 id2 (fst alp) -> sigma - | GRef (_,r1,_), NRef r2 when (eq_gr r1 r2) -> sigma - | GApp (loc,f1,l1), NApp (f2,l2) -> + | GVar id1, NVar id2 when alpha_var id1 id2 (fst alp) -> sigma + | GRef (r1,_), NRef r2 when (eq_gr r1 r2) -> sigma + | GApp (f1,l1), NApp (f2,l2) -> let n1 = List.length l1 and n2 = List.length l2 in let f1,l1,f2,l2 = if n1 < n2 then let l21,l22 = List.chop (n2-n1) l2 in f1,l1, NApp (f2,l21), l22 else if n1 > n2 then - let l11,l12 = List.chop (n1-n2) l1 in GApp (loc,f1,l11),l12, f2,l2 + let l11,l12 = List.chop (n1-n2) l1 in Loc.tag ~loc @@ GApp (f1,l11),l12, f2,l2 else f1,l1, f2, l2 in let may_use_eta = does_not_come_from_already_eta_expanded_var f1 in List.fold_left2 (match_ may_use_eta u alp metas) (match_in u alp metas sigma f1 f2) l1 l2 - | GLambda (_,na1,_,t1,b1), NLambda (na2,t2,b2) -> + | GLambda (na1,_,t1,b1), NLambda (na2,t2,b2) -> match_binders u alp metas na1 na2 (match_in u alp metas sigma t1 t2) b1 b2 - | GProd (_,na1,_,t1,b1), NProd (na2,t2,b2) -> + | GProd (na1,_,t1,b1), NProd (na2,t2,b2) -> match_binders u alp metas na1 na2 (match_in u alp metas sigma t1 t2) b1 b2 - | GLetIn (_,na1,b1,_,c1), NLetIn (na2,b2,None,c2) - | GLetIn (_,na1,b1,None,c1), NLetIn (na2,b2,_,c2) -> + | GLetIn (na1,b1,_,c1), NLetIn (na2,b2,None,c2) + | GLetIn (na1,b1,None,c1), NLetIn (na2,b2,_,c2) -> match_binders u alp metas na1 na2 (match_in u alp metas sigma b1 b2) c1 c2 - | GLetIn (_,na1,b1,Some t1,c1), NLetIn (na2,b2,Some t2,c2) -> + | GLetIn (na1,b1,Some t1,c1), NLetIn (na2,b2,Some t2,c2) -> match_binders u alp metas na1 na2 (match_in u alp metas (match_in u alp metas sigma b1 b2) t1 t2) c1 c2 - | GCases (_,sty1,rtno1,tml1,eqnl1), NCases (sty2,rtno2,tml2,eqnl2) + | GCases (sty1,rtno1,tml1,eqnl1), NCases (sty2,rtno2,tml2,eqnl2) when sty1 == sty2 && Int.equal (List.length tml1) (List.length tml2) && Int.equal (List.length eqnl1) (List.length eqnl2) -> @@ -1065,17 +1070,17 @@ let rec match_ inner u alp metas sigma a1 a2 = (fun s (tm1,_) (tm2,_) -> match_in u alp metas s tm1 tm2) sigma tml1 tml2 in List.fold_left2 (match_equations u alp metas) sigma eqnl1 eqnl2 - | GLetTuple (_,nal1,(na1,to1),b1,c1), NLetTuple (nal2,(na2,to2),b2,c2) + | GLetTuple (nal1,(na1,to1),b1,c1), NLetTuple (nal2,(na2,to2),b2,c2) when Int.equal (List.length nal1) (List.length nal2) -> let sigma = match_opt (match_binders u alp metas na1 na2) sigma to1 to2 in let sigma = match_in u alp metas sigma b1 b2 in let (alp,sigma) = List.fold_left2 (match_names metas) (alp,sigma) nal1 nal2 in match_in u alp metas sigma c1 c2 - | GIf (_,a1,(na1,to1),b1,c1), NIf (a2,(na2,to2),b2,c2) -> + | GIf (a1,(na1,to1),b1,c1), NIf (a2,(na2,to2),b2,c2) -> let sigma = match_opt (match_binders u alp metas na1 na2) sigma to1 to2 in List.fold_left2 (match_in u alp metas) sigma [a1;b1;c1] [a2;b2;c2] - | GRec (_,fk1,idl1,dll1,tl1,bl1), NRec (fk2,idl2,dll2,tl2,bl2) + | GRec (fk1,idl1,dll1,tl1,bl1), NRec (fk2,idl2,dll2,tl2,bl2) when match_fix_kind fk1 fk2 && Int.equal (Array.length idl1) (Array.length idl2) && Array.for_all2 (fun l1 l2 -> Int.equal (List.length l1) (List.length l2)) dll1 dll2 -> @@ -1089,13 +1094,13 @@ let rec match_ inner u alp metas sigma a1 a2 = let alp,sigma = Array.fold_right2 (fun id1 id2 alsig -> match_names metas alsig (Name id1) (Name id2)) idl1 idl2 (alp,sigma) in Array.fold_left2 (match_in u alp metas) sigma bl1 bl2 - | GCast(_,c1,CastConv t1), NCast (c2,CastConv t2) - | GCast(_,c1,CastVM t1), NCast (c2,CastVM t2) -> + | GCast(c1,CastConv t1), NCast (c2,CastConv t2) + | GCast(c1,CastVM t1), NCast (c2,CastVM t2) -> match_in u alp metas (match_in u alp metas sigma c1 c2) t1 t2 - | GCast(_,c1, CastCoerce), NCast(c2, CastCoerce) -> + | GCast(c1, CastCoerce), NCast(c2, CastCoerce) -> match_in u alp metas sigma c1 c2 - | GSort (_,GType _), NSort (GType _) when not u -> sigma - | GSort (_,s1), NSort s2 when Miscops.glob_sort_eq s1 s2 -> sigma + | GSort (GType _), NSort (GType _) when not u -> sigma + | GSort s1, NSort s2 when Miscops.glob_sort_eq s1 s2 -> sigma | GPatVar _, NHole _ -> (*Don't hide Metas, they bind in ltac*) raise No_match | a, NHole _ -> sigma @@ -1105,11 +1110,11 @@ let rec match_ inner u alp metas sigma a1 a2 = otherwise how to ensure it corresponds to a well-typed eta-expansion; we make an exception for types which are metavariables: this is useful e.g. to print "{x:_ & P x}" knowing that notation "{x & P x}" is not defined. *) - | b1, NLambda (Name id as na,(NHole _ | NVar _ as t2),b2) when inner -> + | _b1, NLambda (Name id as na,(NHole _ | NVar _ as t2),b2) when inner -> let avoid = - free_glob_vars b1 @ (* as in Namegen: *) glob_visible_short_qualid b1 in + free_glob_vars a1 @ (* as in Namegen: *) glob_visible_short_qualid a1 in let id' = Namegen.next_ident_away id avoid in - let t1 = GHole(Loc.ghost,Evar_kinds.BinderType (Name id'),Misctypes.IntroAnonymous,None) in + let t1 = Loc.tag @@ GHole(Evar_kinds.BinderType (Name id'),Misctypes.IntroAnonymous,None) in let sigma = match t2 with | NHole _ -> sigma | NVar id2 -> bind_term_env alp sigma id2 t1 @@ -1119,7 +1124,7 @@ let rec match_ inner u alp metas sigma a1 a2 = bind_bindinglist_env alp sigma id [GLocalAssum (Loc.ghost,Name id',Explicit,t1)] else match_names metas (alp,sigma) (Name id') na in - match_in u alp metas sigma (mkGApp Loc.ghost b1 (GVar (Loc.ghost,id'))) b2 + match_in u alp metas sigma (mkGApp Loc.ghost a1 (Loc.tag @@ GVar id')) b2 | (GRec _ | GEvar _), _ | _,_ -> raise No_match @@ -1132,7 +1137,7 @@ and match_binders u alp metas na1 na2 sigma b1 b2 = let (alp,sigma) = match_names metas (alp,sigma) na1 na2 in match_in u alp metas sigma b1 b2 -and match_equations u alp metas sigma (_,_,patl1,rhs1) (patl2,rhs2) = +and match_equations u alp metas sigma (_,(_,patl1,rhs1)) (patl2,rhs2) = (* patl1 and patl2 have the same length because they respectively correspond to some tml1 and tml2 that have the same length *) let (alp,sigma) = @@ -1140,9 +1145,9 @@ and match_equations u alp metas sigma (_,_,patl1,rhs1) (patl2,rhs2) = (alp,sigma) patl1 patl2 in match_in u alp metas sigma rhs1 rhs2 -let term_of_binder = function - | Name id -> GVar (Loc.ghost,id) - | Anonymous -> GHole (Loc.ghost,Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None) +let term_of_binder bi = Loc.tag @@ match bi with + | Name id -> GVar id + | Anonymous -> GHole (Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None) type glob_decl2 = (name, cases_pattern) Util.union * Decl_kinds.binding_kind * @@ -1157,7 +1162,7 @@ let match_notation_constr u c (metas,pat) = with Not_found -> (* Happens for binders bound to Anonymous *) (* Find a better way to propagate Anonymous... *) - GVar (Loc.ghost,x) in + Loc.tag @@GVar x in List.fold_right (fun (x,(scl,typ)) (terms',termlists',binders') -> match typ with | NtnTypeConstr -> diff --git a/intf/glob_term.mli b/intf/glob_term.mli index 7a43c44f9..a8311e093 100644 --- a/intf/glob_term.mli +++ b/intf/glob_term.mli @@ -31,29 +31,29 @@ type cases_pattern_r = and cases_pattern = cases_pattern_r Loc.located (** Representation of an internalized (or in other words globalized) term. *) -type glob_constr = - | GRef of (Loc.t * global_reference * glob_level list option) +type glob_constr_r = + | GRef of global_reference * glob_level list option (** An identifier that represents a reference to an object defined either in the (global) environment or in the (local) context. *) - | GVar of (Loc.t * Id.t) + | GVar of Id.t (** An identifier that cannot be regarded as "GRef". Bound variables are typically represented this way. *) - | GEvar of Loc.t * existential_name * (Id.t * glob_constr) list - | GPatVar of Loc.t * (bool * patvar) (** Used for patterns only *) - | GApp of Loc.t * glob_constr * glob_constr list - | GLambda of Loc.t * Name.t * binding_kind * glob_constr * glob_constr - | GProd of Loc.t * Name.t * binding_kind * glob_constr * glob_constr - | GLetIn of Loc.t * Name.t * glob_constr * glob_constr option * glob_constr - | GCases of Loc.t * case_style * glob_constr option * tomatch_tuples * cases_clauses - (** [GCases(l,style,r,tur,cc)] = "match 'tur' return 'r' with 'cc'" (in [MatchStyle]) *) - | GLetTuple of Loc.t * Name.t list * (Name.t * glob_constr option) * - glob_constr * glob_constr - | GIf of Loc.t * glob_constr * (Name.t * glob_constr option) * glob_constr * glob_constr - | GRec of Loc.t * fix_kind * Id.t array * glob_decl list array * - glob_constr array * glob_constr array - | GSort of Loc.t * glob_sort - | GHole of (Loc.t * Evar_kinds.t * intro_pattern_naming_expr * Genarg.glob_generic_argument option) - | GCast of Loc.t * glob_constr * glob_constr cast_type + | GEvar of existential_name * (Id.t * glob_constr) list + | GPatVar of bool * patvar (** Used for patterns only *) + | GApp of glob_constr * glob_constr list + | GLambda of Name.t * binding_kind * glob_constr * glob_constr + | GProd of Name.t * binding_kind * glob_constr * glob_constr + | GLetIn of Name.t * glob_constr * glob_constr option * glob_constr + | GCases of case_style * glob_constr option * tomatch_tuples * cases_clauses + (** [GCases(style,r,tur,cc)] = "match 'tur' return 'r' with 'cc'" (in [MatchStyle]) *) + | GLetTuple of Name.t list * (Name.t * glob_constr option) * glob_constr * glob_constr + | GIf of glob_constr * (Name.t * glob_constr option) * glob_constr * glob_constr + | GRec of fix_kind * Id.t array * glob_decl list array * + glob_constr array * glob_constr array + | GSort of glob_sort + | GHole of Evar_kinds.t * intro_pattern_naming_expr * Genarg.glob_generic_argument option + | GCast of glob_constr * glob_constr cast_type +and glob_constr = glob_constr_r Loc.located and glob_decl = Name.t * binding_kind * glob_constr option * glob_constr @@ -74,7 +74,7 @@ and tomatch_tuple = (glob_constr * predicate_pattern) and tomatch_tuples = tomatch_tuple list -and cases_clause = (Loc.t * Id.t list * cases_pattern list * glob_constr) +and cases_clause = (Id.t list * cases_pattern list * glob_constr) Loc.located (** [(p,il,cl,t)] = "|'cl' => 't'". Precondition: the free variables of [t] are members of [il]. *) and cases_clauses = cases_clause list diff --git a/lib/loc.ml b/lib/loc.ml index 8d7432ff4..2a785fac4 100644 --- a/lib/loc.ml +++ b/lib/loc.ml @@ -62,6 +62,7 @@ type 'a located = t * 'a let to_pair x = x let tag ?loc x = Option.default ghost loc, x +let obj (_,x) = x let with_loc f (loc, x) = f ~loc x let with_unloc f (_,x) = f x diff --git a/lib/loc.mli b/lib/loc.mli index 3f484bc4c..10f63a8dd 100644 --- a/lib/loc.mli +++ b/lib/loc.mli @@ -60,6 +60,7 @@ type 'a located = t * 'a val to_pair : 'a located -> t * 'a val tag : ?loc:t -> 'a -> 'a located +val obj : 'a located -> 'a val with_loc : (loc:t -> 'a -> 'b) -> 'a located -> 'b val with_unloc : ('a -> 'b) -> 'a located -> 'b diff --git a/plugins/decl_mode/g_decl_mode.ml4 b/plugins/decl_mode/g_decl_mode.ml4 deleted file mode 100644 index 2415080f3..000000000 --- a/plugins/decl_mode/g_decl_mode.ml4 +++ /dev/null @@ -1,387 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* - pr_goal { Evd.it = goal ; sigma = sigma } - | _ -> - (* spiwack: it's not very nice to have to call proof global - here, a more robust solution would be to add a hook for - [Printer.pr_open_subgoals] in proof modes, in order to - compute the end command. Yet a more robust solution would be - to have focuses give explanations of their unfocusing - behaviour. *) - let p = Proof_global.give_me_the_proof () in - let close_cmd = Decl_mode.get_end_command p in - str "Subproof completed, now type " ++ str close_cmd ++ str "." - -let interp_proof_instr _ { Evd.it = gl ; sigma = sigma }= - Decl_interp.interp_proof_instr - (Decl_mode.get_info sigma gl) - (Goal.V82.env sigma gl) - (sigma) - -let vernac_decl_proof () = - let pf = Proof_global.give_me_the_proof () in - if Proof.is_done pf then - CErrors.error "Nothing left to prove here." - else - begin - Decl_proof_instr.go_to_proof_mode () ; - Proof_global.set_proof_mode "Declarative" - end - -(* spiwack: some bureaucracy is not performed here *) -let vernac_return () = - begin - Decl_proof_instr.return_from_tactic_mode () ; - Proof_global.set_proof_mode "Declarative" - end - -let vernac_proof_instr instr = - Decl_proof_instr.proof_instr instr - -(* Before we can write an new toplevel command (see below) - which takes a [proof_instr] as argument, we need to declare - how to parse it, print it, globalise it and interprete it. - Normally we could do that easily through ARGUMENT EXTEND, - but as the parsing is fairly complicated we will do it manually to - indirect through the [proof_instr] grammar entry. *) -(* spiwack: proposal: doing that directly from argextend.ml4, maybe ? *) - -(* Only declared at raw level, because only used in vernac commands. *) -let wit_proof_instr : (raw_proof_instr, glob_proof_instr, proof_instr) Genarg.genarg_type = - Genarg.make0 "proof_instr" - -(* We create a new parser entry [proof_mode]. The Declarative proof mode - will replace the normal parser entry for tactics with this one. *) -let proof_mode : vernac_expr Gram.entry = - Gram.entry_create "vernac:proof_command" -(* Auxiliary grammar entry. *) -let proof_instr : raw_proof_instr Gram.entry = - Pcoq.create_generic_entry Pcoq.utactic "proof_instr" (Genarg.rawwit wit_proof_instr) - -let _ = Pptactic.declare_extra_genarg_pprule wit_proof_instr - pr_raw_proof_instr pr_glob_proof_instr pr_proof_instr - -let classify_proof_instr = function - | { instr = Pescape |Pend B_proof } -> VtProofMode "Classic", VtNow - | _ -> Vernac_classifier.classify_as_proofstep - -(* We use the VERNAC EXTEND facility with a custom non-terminal - to populate [proof_mode] with a new toplevel interpreter. - The "-" indicates that the rule does not start with a distinguished - string. *) -VERNAC proof_mode EXTEND ProofInstr - [ - proof_instr(instr) ] => [classify_proof_instr instr] -> [ vernac_proof_instr instr ] -END - -(* It is useful to use GEXTEND directly to call grammar entries that have been - defined previously VERNAC EXTEND. In this case we allow, in proof mode, - the use of commands like Check or Print. VERNAC EXTEND does quite a bit of - bureaucracy for us, but it is not needed in this sort of case, and it would require - to have an ARGUMENT EXTEND version of the "proof_mode" grammar entry. *) -GEXTEND Gram - GLOBAL: proof_mode ; - - proof_mode: LAST - [ [ c=G_vernac.subgoal_command -> c (Some (Vernacexpr.SelectNth 1)) ] ] - ; -END - -(* We register a new proof mode here *) - -let _ = - Proof_global.register_proof_mode { Proof_global. - name = "Declarative" ; (* name for identifying and printing *) - (* function [set] goes from No Proof Mode to - Declarative Proof Mode performing side effects *) - set = begin fun () -> - (* We set the command non terminal to - [proof_mode] (which we just defined). *) - Pcoq.set_command_entry proof_mode ; - (* We substitute the goal printer, by the one we built - for the proof mode. *) - Printer.set_printer_pr { Printer.default_printer_pr with - Printer.pr_goal = pr_goal; - pr_subgoals = pr_subgoals; } - end ; - (* function [reset] goes back to No Proof Mode from - Declarative Proof Mode *) - reset = begin fun () -> - (* We restore the command non terminal to - [noedit_mode]. *) - Pcoq.set_command_entry Pcoq.Vernac_.noedit_mode ; - (* We restore the goal printer to default *) - Printer.set_printer_pr Printer.default_printer_pr - end - } - -VERNAC COMMAND EXTEND DeclProof -[ "proof" ] => [ VtProofMode "Declarative", VtNow ] -> [ vernac_decl_proof () ] -END -VERNAC COMMAND EXTEND DeclReturn -[ "return" ] => [ VtProofMode "Declarative", VtNow ] -> [ vernac_return () ] -END - -let none_is_empty = function - None -> [] - | Some l -> l - -GEXTEND Gram -GLOBAL: proof_instr; - thesis : - [[ "thesis" -> Plain - | "thesis"; "for"; i=ident -> (For i) - ]]; - statement : - [[ i=ident ; ":" ; c=constr -> {st_label=Name i;st_it=c} - | i=ident -> {st_label=Anonymous; - st_it= Loc.tag ~loc:!@loc @@ Constrexpr.CRef (Libnames.Ident (!@loc, i), None)} - | c=constr -> {st_label=Anonymous;st_it=c} - ]]; - constr_or_thesis : - [[ t=thesis -> Thesis t ] | - [ c=constr -> This c - ]]; - statement_or_thesis : - [ - [ t=thesis -> {st_label=Anonymous;st_it=Thesis t} ] - | - [ i=ident ; ":" ; cot=constr_or_thesis -> {st_label=Name i;st_it=cot} - | i=ident -> {st_label=Anonymous; - st_it=This (Loc.tag ~loc:!@loc @@ Constrexpr.CRef (Libnames.Ident (!@loc, i), None))} - | c=constr -> {st_label=Anonymous;st_it=This c} - ] - ]; - justification_items : - [[ -> Some [] - | "by"; l=LIST1 constr SEP "," -> Some l - | "by"; "*" -> None ]] - ; - justification_method : - [[ -> None - | "using"; tac = tactic -> Some tac ]] - ; - simple_cut_or_thesis : - [[ ls = statement_or_thesis; - j = justification_items; - taco = justification_method - -> {cut_stat=ls;cut_by=j;cut_using=taco} ]] - ; - simple_cut : - [[ ls = statement; - j = justification_items; - taco = justification_method - -> {cut_stat=ls;cut_by=j;cut_using=taco} ]] - ; - elim_type: - [[ IDENT "induction" -> ET_Induction - | IDENT "cases" -> ET_Case_analysis ]] - ; - block_type : - [[ IDENT "claim" -> B_claim - | IDENT "focus" -> B_focus - | IDENT "proof" -> B_proof - | et=elim_type -> B_elim et ]] - ; - elim_obj: - [[ IDENT "on"; c=constr -> Real c - | IDENT "of"; c=simple_cut -> Virtual c ]] - ; - elim_step: - [[ IDENT "consider" ; - h=consider_vars ; IDENT "from" ; c=constr -> Pconsider (c,h) - | IDENT "per"; et=elim_type; obj=elim_obj -> Pper (et,obj) - | IDENT "suffices"; ls=suff_clause; - j = justification_items; - taco = justification_method - -> Psuffices {cut_stat=ls;cut_by=j;cut_using=taco} ]] - ; - rew_step : - [[ "~=" ; c=simple_cut -> (Rhs,c) - | "=~" ; c=simple_cut -> (Lhs,c)]] - ; - cut_step: - [[ "then"; tt=elim_step -> Pthen tt - | "then"; c=simple_cut_or_thesis -> Pthen (Pcut c) - | IDENT "thus"; tt=rew_step -> Pthus (let s,c=tt in Prew (s,c)) - | IDENT "thus"; c=simple_cut_or_thesis -> Pthus (Pcut c) - | IDENT "hence"; c=simple_cut_or_thesis -> Phence (Pcut c) - | tt=elim_step -> tt - | tt=rew_step -> let s,c=tt in Prew (s,c); - | IDENT "have"; c=simple_cut_or_thesis -> Pcut c; - | IDENT "claim"; c=statement -> Pclaim c; - | IDENT "focus"; IDENT "on"; c=statement -> Pfocus c; - | "end"; bt = block_type -> Pend bt; - | IDENT "escape" -> Pescape ]] - ; - (* examiner s'il est possible de faire R _ et _ R pour R une relation qcq*) - loc_id: - [[ id=ident -> fun x -> (!@loc,(id,x)) ]]; - hyp: - [[ id=loc_id -> id None ; - | id=loc_id ; ":" ; c=constr -> id (Some c)]] - ; - consider_vars: - [[ name=hyp -> [Hvar name] - | name=hyp; ","; v=consider_vars -> (Hvar name) :: v - | name=hyp; - IDENT "such"; IDENT "that"; h=consider_hyps -> (Hvar name)::h - ]] - ; - consider_hyps: - [[ st=statement; IDENT "and"; h=consider_hyps -> Hprop st::h - | st=statement; IDENT "and"; - IDENT "consider" ; v=consider_vars -> Hprop st::v - | st=statement -> [Hprop st] - ]] - ; - assume_vars: - [[ name=hyp -> [Hvar name] - | name=hyp; ","; v=assume_vars -> (Hvar name) :: v - | name=hyp; - IDENT "such"; IDENT "that"; h=assume_hyps -> (Hvar name)::h - ]] - ; - assume_hyps: - [[ st=statement; IDENT "and"; h=assume_hyps -> Hprop st::h - | st=statement; IDENT "and"; - IDENT "we"; IDENT "have" ; v=assume_vars -> Hprop st::v - | st=statement -> [Hprop st] - ]] - ; - assume_clause: - [[ IDENT "we" ; IDENT "have" ; v=assume_vars -> v - | h=assume_hyps -> h ]] - ; - suff_vars: - [[ name=hyp; IDENT "to"; IDENT "show" ; c = constr_or_thesis -> - [Hvar name],c - | name=hyp; ","; v=suff_vars -> - let (q,c) = v in ((Hvar name) :: q),c - | name=hyp; - IDENT "such"; IDENT "that"; h=suff_hyps -> - let (q,c) = h in ((Hvar name) :: q),c - ]]; - suff_hyps: - [[ st=statement; IDENT "and"; h=suff_hyps -> - let (q,c) = h in (Hprop st::q),c - | st=statement; IDENT "and"; - IDENT "to" ; IDENT "have" ; v=suff_vars -> - let (q,c) = v in (Hprop st::q),c - | st=statement; IDENT "to"; IDENT "show" ; c = constr_or_thesis -> - [Hprop st],c - ]] - ; - suff_clause: - [[ IDENT "to" ; IDENT "have" ; v=suff_vars -> v - | h=suff_hyps -> h ]] - ; - let_vars: - [[ name=hyp -> [Hvar name] - | name=hyp; ","; v=let_vars -> (Hvar name) :: v - | name=hyp; IDENT "be"; - IDENT "such"; IDENT "that"; h=let_hyps -> (Hvar name)::h - ]] - ; - let_hyps: - [[ st=statement; IDENT "and"; h=let_hyps -> Hprop st::h - | st=statement; IDENT "and"; "let"; v=let_vars -> Hprop st::v - | st=statement -> [Hprop st] - ]]; - given_vars: - [[ name=hyp -> [Hvar name] - | name=hyp; ","; v=given_vars -> (Hvar name) :: v - | name=hyp; IDENT "such"; IDENT "that"; h=given_hyps -> (Hvar name)::h - ]] - ; - given_hyps: - [[ st=statement; IDENT "and"; h=given_hyps -> Hprop st::h - | st=statement; IDENT "and"; IDENT "given"; v=given_vars -> Hprop st::v - | st=statement -> [Hprop st] - ]]; - suppose_vars: - [[name=hyp -> [Hvar name] - |name=hyp; ","; v=suppose_vars -> (Hvar name) :: v - |name=hyp; OPT[IDENT "be"]; - IDENT "such"; IDENT "that"; h=suppose_hyps -> (Hvar name)::h - ]] - ; - suppose_hyps: - [[ st=statement_or_thesis; IDENT "and"; h=suppose_hyps -> Hprop st::h - | st=statement_or_thesis; IDENT "and"; IDENT "we"; IDENT "have"; - v=suppose_vars -> Hprop st::v - | st=statement_or_thesis -> [Hprop st] - ]] - ; - suppose_clause: - [[ IDENT "we"; IDENT "have"; v=suppose_vars -> v; - | h=suppose_hyps -> h ]] - ; - intro_step: - [[ IDENT "suppose" ; h=assume_clause -> Psuppose h - | IDENT "suppose" ; IDENT "it"; IDENT "is" ; c=pattern LEVEL "0" ; - po=OPT[ "with"; p=LIST1 hyp SEP ","-> p ] ; - ho=OPT[ IDENT "and" ; h=suppose_clause -> h ] -> - Pcase (none_is_empty po,c,none_is_empty ho) - | "let" ; v=let_vars -> Plet v - | IDENT "take"; witnesses = LIST1 constr SEP "," -> Ptake witnesses - | IDENT "assume"; h=assume_clause -> Passume h - | IDENT "given"; h=given_vars -> Pgiven h - | IDENT "define"; id=ident; args=LIST0 hyp; - "as"; body=constr -> Pdefine(id,args,body) - | IDENT "reconsider"; id=ident; "as" ; typ=constr -> Pcast (This id,typ) - | IDENT "reconsider"; t=thesis; "as" ; typ=constr -> Pcast (Thesis t ,typ) - ]] - ; - emphasis : - [[ -> 0 - | "*" -> 1 - | "**" -> 2 - | "***" -> 3 - ]] - ; - bare_proof_instr: - [[ c = cut_step -> c ; - | i = intro_step -> i ]] - ; - proof_instr : - [[ e=emphasis;i=bare_proof_instr;"." -> {emph=e;instr=i}]] - ; -END;; diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml index 85d465a4b..07a0d5a50 100644 --- a/plugins/funind/glob_term_to_relation.ml +++ b/plugins/funind/glob_term_to_relation.ml @@ -274,10 +274,10 @@ let make_discr_match_el = *) let make_discr_match_brl i = List.map_i - (fun j (_,idl,patl,_) -> + (fun j (_,(idl,patl,_)) -> Loc.tag @@ if Int.equal j i - then (Loc.ghost,idl,patl, mkGRef (Lazy.force coq_True_ref)) - else (Loc.ghost,idl,patl, mkGRef (Lazy.force coq_False_ref)) + then (idl,patl, mkGRef (Lazy.force coq_True_ref)) + else (idl,patl, mkGRef (Lazy.force coq_False_ref)) ) 0 (* @@ -464,13 +464,13 @@ let rec pattern_to_term_and_type env typ = Loc.with_unloc (function *) -let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = +let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = observe (str " Entering : " ++ Printer.pr_glob_constr rt); match rt with - | GRef _ | GVar _ | GEvar _ | GPatVar _ | GSort _ | GHole _ -> + | _, GRef _ | _, GVar _ | _, GEvar _ | _, GPatVar _ | _, GSort _ | _, GHole _ -> (* do nothing (except changing type of course) *) mk_result [] rt avoid - | GApp(_,_,_) -> + | _, GApp(_,_) -> let f,args = glob_decompose_app rt in let args_res : (glob_constr list) build_entry_return = List.fold_right (* create the arguments lists of constructors and combine them *) @@ -482,20 +482,20 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = (mk_result [] [] avoid) in begin - match f with + match Loc.obj f with | GLambda _ -> let rec aux t l = match l with | [] -> t - | u::l -> + | u::l -> Loc.tag @@ match t with - | GLambda(loc,na,_,nat,b) -> - GLetIn(Loc.ghost,na,u,None,aux b l) + | loc, GLambda(na,_,nat,b) -> + GLetIn(na,u,None,aux b l) | _ -> - GApp(Loc.ghost,t,l) + GApp(t,l) in build_entry_lc env funnames avoid (aux f args) - | GVar(_,id) when Id.Set.mem id funnames -> + | GVar id when Id.Set.mem id funnames -> (* if we have [f t1 ... tn] with [f]$\in$[fnames] then we create a fresh variable [res], add [res] and its "value" (i.e. [res v1 ... vn]) to each @@ -536,7 +536,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = args_res.result } | GApp _ -> assert false (* we have collected all the app in [glob_decompose_app] *) - | GLetIn(_,n,v,t,b) -> + | GLetIn(n,v,t,b) -> (* if we have [(let x := v in b) t1 ... tn] , we discard our work and compute the list of constructor for [let x = v in (b t1 ... tn)] up to alpha conversion @@ -550,7 +550,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = let new_b = replace_var_by_term id - (GVar(Loc.ghost,id)) + (Loc.tag @@ GVar id) b in (Name new_id,new_b,new_avoid) @@ -568,7 +568,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = *) let f_res = build_entry_lc env funnames args_res.to_avoid f in combine_results combine_app f_res args_res - | GCast(_,b,_) -> + | GCast(b,_) -> (* for an applied cast we just trash the cast part and restart the work. @@ -579,7 +579,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = | GProd _ -> error "Cannot apply a type" end (* end of the application treatement *) - | GLambda(_,n,_,t,b) -> + | _, GLambda(n,_,t,b) -> (* we first compute the list of constructor corresponding to the body of the function, then the one corresponding to the type @@ -594,7 +594,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = let new_env = raw_push_named (new_n,None,t) env in let b_res = build_entry_lc new_env funnames avoid b in combine_results (combine_lam new_n) t_res b_res - | GProd(_,n,_,t,b) -> + | _, GProd(n,_,t,b) -> (* we first compute the list of constructor corresponding to the body of the function, then the one corresponding to the type @@ -604,13 +604,13 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = let new_env = raw_push_named (n,None,t) env in let b_res = build_entry_lc new_env funnames avoid b in combine_results (combine_prod n) t_res b_res - | GLetIn(loc,n,v,typ,b) -> + | loc, GLetIn(n,v,typ,b) -> (* we first compute the list of constructor corresponding to the body of the function, then the one corresponding to the value [t] and combine the two result *) - let v = match typ with None -> v | Some t -> GCast (loc,v,CastConv t) in + let v = match typ with None -> v | Some t -> Loc.tag ~loc @@ GCast (v,CastConv t) in let v_res = build_entry_lc env funnames avoid v in let v_as_constr,ctx = Pretyping.understand env (Evd.from_env env) v in let v_type = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr v_as_constr) in @@ -622,13 +622,13 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = in let b_res = build_entry_lc new_env funnames avoid b in combine_results (combine_letin n) v_res b_res - | GCases(_,_,_,el,brl) -> + | _, GCases(_,_,el,brl) -> (* we create the discrimination function and treat the case itself *) let make_discr = make_discr_match brl in build_entry_lc_from_case env funnames make_discr el brl avoid - | GIf(_,b,(na,e_option),lhs,rhs) -> + | _, GIf(b,(na,e_option),lhs,rhs) -> let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in let b_typ = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr b_as_constr) in let (ind,_) = @@ -642,7 +642,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = assert (Int.equal (Array.length case_pats) 2); let brl = List.map_i - (fun i x -> (Loc.ghost,[],[case_pats.(i)],x)) + (fun i x -> Loc.tag ([],[case_pats.(i)],x)) 0 [lhs;rhs] in @@ -651,7 +651,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = in (* Pp.msgnl (str "new case := " ++ Printer.pr_glob_constr match_expr); *) build_entry_lc env funnames avoid match_expr - | GLetTuple(_,nal,_,b,e) -> + | _, GLetTuple(nal,_,b,e) -> begin let nal_as_glob_constr = List.map @@ -673,14 +673,14 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = let case_pats = build_constructors_of_type (fst ind) nal_as_glob_constr in assert (Int.equal (Array.length case_pats) 1); let br = - (Loc.ghost,[],[case_pats.(0)],e) + (Loc.ghost,([],[case_pats.(0)],e)) in let match_expr = mkGCases(None,[b,(Anonymous,None)],[br]) in build_entry_lc env funnames avoid match_expr end - | GRec _ -> error "Not handled GRec" - | GCast(_,b,_) -> + | _, GRec _ -> error "Not handled GRec" + | _, GCast(b,_) -> build_entry_lc env funnames avoid b and build_entry_lc_from_case env funname make_discr (el:tomatch_tuples) @@ -740,7 +740,7 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve | [] -> (* computed_branches *) {result = [];to_avoid = avoid} | br::brl' -> (* alpha conversion to prevent name clashes *) - let _,idl,patl,return = alpha_br avoid br in + let _,(idl,patl,return) = alpha_br avoid br in let new_avoid = idl@avoid in (* for now we can no more use idl as an identifier *) (* building a list of precondition stating that we are not in this branch (will be used in the following recursive calls) @@ -862,9 +862,9 @@ let is_res id = -let same_raw_term rt1 rt2 = +let same_raw_term (_,rt1) (_,rt2) = match rt1,rt2 with - | GRef(_,r1,_), GRef (_,r2,_) -> Globnames.eq_gr r1 r2 + | GRef(r1,_), GRef (r2,_) -> Globnames.eq_gr r1 r2 | GHole _, GHole _ -> true | _ -> false let decompose_raw_eq lhs rhs = @@ -897,15 +897,15 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = observe (str "rebuilding : " ++ pr_glob_constr rt); let open Context.Rel.Declaration in match rt with - | GProd(_,n,k,t,b) -> + | _, GProd(n,k,t,b) -> let not_free_in_t id = not (is_free_in id t) in let new_crossed_types = t::crossed_types in begin match t with - | GApp(_,(GVar(_,res_id) as res_rt),args') when is_res res_id -> + | _, GApp(((_, GVar(res_id)) as res_rt),args') when is_res res_id -> begin match args' with - | (GVar(_,this_relname))::args' -> + | (_, (GVar this_relname))::args' -> (*i The next call to mk_rel_id is valid since we are constructing the graph Ensures by: obvious @@ -927,7 +927,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = | _ -> (* the first args is the name of the function! *) assert false end - | GApp(loc1,GRef(loc2,eq_as_ref,_),[ty;GVar(loc3,id);rt]) + | loc1, GApp((loc2, GRef(eq_as_ref,_)),[ty;(loc3, GVar id);rt]) when Globnames.eq_gr eq_as_ref (Lazy.force Coqlib.coq_eq_ref) && n == Anonymous -> begin @@ -964,9 +964,8 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = let params,arg' = ((Util.List.chop nparam args')) in - let rt_typ = - GApp(Loc.ghost, - GRef (Loc.ghost,Globnames.IndRef (fst ind),None), + let rt_typ = Loc.tag @@ + GApp(Loc.tag @@ GRef (Globnames.IndRef (fst ind),None), (List.map (fun p -> Detyping.detype false [] env (Evd.from_env env) @@ -976,7 +975,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = (mkGHole ())))) in let eq' = - GApp(loc1,GRef(loc2,jmeq,None),[ty;GVar(loc3,id);rt_typ;rt]) + Loc.tag ~loc:loc1 @@ GApp(Loc.tag ~loc:loc2 @@GRef(jmeq,None),[ty;Loc.tag ~loc:loc3 @@ GVar id;rt_typ;rt]) in observe (str "computing new type for jmeq : " ++ pr_glob_constr eq'); let eq'_as_constr,ctx = Pretyping.understand env (Evd.from_env env) eq' in @@ -1045,7 +1044,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = mkGProd(n,t,new_b),id_to_exclude else new_b, Id.Set.add id id_to_exclude *) - | GApp(loc1,GRef(loc2,eq_as_ref,_),[ty;rt1;rt2]) + | loc1, GApp((loc2, GRef(eq_as_ref,_)),[ty;rt1;rt2]) when Globnames.eq_gr eq_as_ref (Lazy.force Coqlib.coq_eq_ref) && n == Anonymous -> begin @@ -1096,7 +1095,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = (Id.Set.filter not_free_in_t id_to_exclude) | _ -> mkGProd(n,t,new_b),Id.Set.filter not_free_in_t id_to_exclude end - | GLambda(_,n,k,t,b) -> + | _, GLambda(n,k,t,b) -> begin let not_free_in_t id = not (is_free_in id t) in let new_crossed_types = t :: crossed_types in @@ -1115,14 +1114,14 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = then new_b, Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude) else - GProd(Loc.ghost,n,k,t,new_b),Id.Set.filter not_free_in_t id_to_exclude + Loc.tag @@ GProd(n,k,t,new_b),Id.Set.filter not_free_in_t id_to_exclude | _ -> anomaly (Pp.str "Should not have an anonymous function here") (* We have renamed all the anonymous functions during alpha_renaming phase *) end - | GLetIn(loc,n,v,t,b) -> + | loc, GLetIn(n,v,t,b) -> begin - let t = match t with None -> v | Some t -> GCast (loc,v,CastConv t) in + let t = match t with None -> v | Some t -> Loc.tag ~loc @@ GCast (v,CastConv t) in let not_free_in_t id = not (is_free_in id t) in let evd = (Evd.from_env env) in let t',ctx = Pretyping.understand env evd t in @@ -1138,10 +1137,10 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = match n with | Name id when Id.Set.mem id id_to_exclude && depth >= nb_args -> new_b,Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude) - | _ -> GLetIn(Loc.ghost,n,t,None,new_b), (* HOPING IT WOULD WORK *) + | _ -> Loc.tag @@ GLetIn(n,t,None,new_b), (* HOPING IT WOULD WORK *) Id.Set.filter not_free_in_t id_to_exclude end - | GLetTuple(_,nal,(na,rto),t,b) -> + | _, GLetTuple(nal,(na,rto),t,b) -> assert (Option.is_empty rto); begin let not_free_in_t id = not (is_free_in id t) in @@ -1164,7 +1163,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = (* | Name id when Id.Set.mem id id_to_exclude -> *) (* new_b,Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude) *) (* | _ -> *) - GLetTuple(Loc.ghost,nal,(na,None),t,new_b), + Loc.tag @@ GLetTuple(nal,(na,None),t,new_b), Id.Set.filter not_free_in_t (Id.Set.union id_to_exclude id_to_exclude') end @@ -1190,16 +1189,16 @@ let rebuild_cons env nb_args relname args crossed_types rt = TODO: Find a valid way to deal with implicit arguments here! *) -let rec compute_cst_params relnames params = function +let rec compute_cst_params relnames params gt = Loc.with_unloc (function | GRef _ | GVar _ | GEvar _ | GPatVar _ -> params - | GApp(_,GVar(_,relname'),rtl) when Id.Set.mem relname' relnames -> + | GApp((_, GVar relname'),rtl) when Id.Set.mem relname' relnames -> compute_cst_params_from_app [] (params,rtl) - | GApp(_,f,args) -> + | GApp(f,args) -> List.fold_left (compute_cst_params relnames) params (f::args) - | GLambda(_,_,_,t,b) | GProd(_,_,_,t,b) | GLetTuple(_,_,_,t,b) -> + | GLambda(_,_,t,b) | GProd(_,_,t,b) | GLetTuple(_,_,t,b) -> let t_params = compute_cst_params relnames params t in compute_cst_params relnames t_params b - | GLetIn(_,_,v,t,b) -> + | GLetIn(_,v,t,b) -> let v_params = compute_cst_params relnames params v in let t_params = Option.fold_left (compute_cst_params relnames) v_params t in compute_cst_params relnames t_params b @@ -1210,10 +1209,11 @@ let rec compute_cst_params relnames params = function | GHole _ -> params | GIf _ | GRec _ | GCast _ -> raise (UserError(Some "compute_cst_params", str "Not handled case")) + ) gt and compute_cst_params_from_app acc (params,rtl) = match params,rtl with | _::_,[] -> assert false (* the rel has at least nargs + 1 arguments ! *) - | ((Name id,_,None) as param)::params',(GVar(_,id'))::rtl' + | ((Name id,_,None) as param)::params',(_, GVar id')::rtl' when Id.compare id id' == 0 -> compute_cst_params_from_app (param::acc) (params',rtl') | _ -> List.rev acc diff --git a/plugins/funind/glob_termops.ml b/plugins/funind/glob_termops.ml index 51ca8c471..01e607412 100644 --- a/plugins/funind/glob_termops.ml +++ b/plugins/funind/glob_termops.ml @@ -10,16 +10,16 @@ open Misctypes Some basic functions to rebuild glob_constr In each of them the location is Loc.ghost *) -let mkGRef ref = GRef(Loc.ghost,ref,None) -let mkGVar id = GVar(Loc.ghost,id) -let mkGApp(rt,rtl) = GApp(Loc.ghost,rt,rtl) -let mkGLambda(n,t,b) = GLambda(Loc.ghost,n,Explicit,t,b) -let mkGProd(n,t,b) = GProd(Loc.ghost,n,Explicit,t,b) -let mkGLetIn(n,b,t,c) = GLetIn(Loc.ghost,n,b,t,c) -let mkGCases(rto,l,brl) = GCases(Loc.ghost,Term.RegularStyle,rto,l,brl) -let mkGSort s = GSort(Loc.ghost,s) -let mkGHole () = GHole(Loc.ghost,Evar_kinds.BinderType Anonymous,Misctypes.IntroAnonymous,None) -let mkGCast(b,t) = GCast(Loc.ghost,b,CastConv t) +let mkGRef ref = Loc.tag @@ GRef(ref,None) +let mkGVar id = Loc.tag @@ GVar(id) +let mkGApp(rt,rtl) = Loc.tag @@ GApp(rt,rtl) +let mkGLambda(n,t,b) = Loc.tag @@ GLambda(n,Explicit,t,b) +let mkGProd(n,t,b) = Loc.tag @@ GProd(n,Explicit,t,b) +let mkGLetIn(n,b,t,c) = Loc.tag @@ GLetIn(n,b,t,c) +let mkGCases(rto,l,brl) = Loc.tag @@ GCases(Term.RegularStyle,rto,l,brl) +let mkGSort s = Loc.tag @@ GSort(s) +let mkGHole () = Loc.tag @@ GHole(Evar_kinds.BinderType Anonymous,Misctypes.IntroAnonymous,None) +let mkGCast(b,t) = Loc.tag @@ GCast(b,CastConv t) (* Some basic functions to decompose glob_constrs @@ -27,7 +27,7 @@ let mkGCast(b,t) = GCast(Loc.ghost,b,CastConv t) *) let glob_decompose_prod = let rec glob_decompose_prod args = function - | GProd(_,n,k,t,b) -> + | _, GProd(n,k,t,b) -> glob_decompose_prod ((n,t)::args) b | rt -> args,rt in @@ -35,9 +35,9 @@ let glob_decompose_prod = let glob_decompose_prod_or_letin = let rec glob_decompose_prod args = function - | GProd(_,n,k,t,b) -> + | _, GProd(n,k,t,b) -> glob_decompose_prod ((n,None,Some t)::args) b - | GLetIn(_,n,b,t,c) -> + | _, GLetIn(n,b,t,c) -> glob_decompose_prod ((n,Some b,t)::args) c | rt -> args,rt in @@ -59,7 +59,7 @@ let glob_decompose_prod_n n = if i<=0 then args,c else match c with - | GProd(_,n,_,t,b) -> + | _, GProd(n,_,t,b) -> glob_decompose_prod (i-1) ((n,t)::args) b | rt -> args,rt in @@ -71,9 +71,9 @@ let glob_decompose_prod_or_letin_n n = if i<=0 then args,c else match c with - | GProd(_,n,_,t,b) -> + | _, GProd(n,_,t,b) -> glob_decompose_prod (i-1) ((n,None,Some t)::args) b - | GLetIn(_,n,b,t,c) -> + | _, GLetIn(n,b,t,c) -> glob_decompose_prod (i-1) ((n,Some b,t)::args) c | rt -> args,rt in @@ -84,7 +84,7 @@ let glob_decompose_app = let rec decompose_rapp acc rt = (* msgnl (str "glob_decompose_app on : "++ Printer.pr_glob_constr rt); *) match rt with - | GApp(_,rt,rtl) -> + | _, GApp(rt,rtl) -> decompose_rapp (List.fold_left (fun y x -> x::y) acc rtl) rt | rt -> rt,List.rev acc in @@ -120,75 +120,70 @@ let remove_name_from_mapping mapping na = let change_vars = let rec change_vars mapping rt = - match rt with - | GRef _ -> rt - | GVar(loc,id) -> + Loc.map (function + | GRef _ as x -> x + | GVar id -> let new_id = try Id.Map.find id mapping with Not_found -> id in - GVar(loc,new_id) - | GEvar _ -> rt - | GPatVar _ -> rt - | GApp(loc,rt',rtl) -> - GApp(loc, - change_vars mapping rt', + GVar(new_id) + | GEvar _ as x -> x + | GPatVar _ as x -> x + | GApp(rt',rtl) -> + GApp(change_vars mapping rt', List.map (change_vars mapping) rtl ) - | GLambda(loc,name,k,t,b) -> - GLambda(loc, - name, + | GLambda(name,k,t,b) -> + GLambda(name, k, change_vars mapping t, change_vars (remove_name_from_mapping mapping name) b ) - | GProd(loc,name,k,t,b) -> - GProd(loc, - name, + | GProd(name,k,t,b) -> + GProd( name, k, change_vars mapping t, change_vars (remove_name_from_mapping mapping name) b ) - | GLetIn(loc,name,def,typ,b) -> - GLetIn(loc, - name, + | GLetIn(name,def,typ,b) -> + GLetIn(name, change_vars mapping def, Option.map (change_vars mapping) typ, change_vars (remove_name_from_mapping mapping name) b ) - | GLetTuple(loc,nal,(na,rto),b,e) -> + | GLetTuple(nal,(na,rto),b,e) -> let new_mapping = List.fold_left remove_name_from_mapping mapping nal in - GLetTuple(loc, - nal, + GLetTuple(nal, (na, Option.map (change_vars mapping) rto), change_vars mapping b, change_vars new_mapping e ) - | GCases(loc,sty,infos,el,brl) -> - GCases(loc,sty, + | GCases(sty,infos,el,brl) -> + GCases(sty, infos, List.map (fun (e,x) -> (change_vars mapping e,x)) el, List.map (change_vars_br mapping) brl ) - | GIf(loc,b,(na,e_option),lhs,rhs) -> - GIf(loc, - change_vars mapping b, + | GIf(b,(na,e_option),lhs,rhs) -> + GIf(change_vars mapping b, (na,Option.map (change_vars mapping) e_option), change_vars mapping lhs, change_vars mapping rhs ) | GRec _ -> error "Local (co)fixes are not supported" - | GSort _ -> rt - | GHole _ -> rt - | GCast(loc,b,c) -> - GCast(loc,change_vars mapping b, + | GSort _ as x -> x + | GHole _ as x -> x + | GCast(b,c) -> + GCast(change_vars mapping b, Miscops.map_cast_type (change_vars mapping) c) - and change_vars_br mapping ((loc,idl,patl,res) as br) = + ) rt + and change_vars_br mapping ((loc,(idl,patl,res)) as br) = let new_mapping = List.fold_right Id.Map.remove idl mapping in if Id.Map.is_empty new_mapping then br - else (loc,idl,patl,change_vars new_mapping res) + else (loc,(idl,patl,change_vars new_mapping res)) in change_vars @@ -259,30 +254,30 @@ let raw_get_pattern_id pat acc = let get_pattern_id pat = raw_get_pattern_id pat [] -let rec alpha_rt excluded rt = - let new_rt = +let rec alpha_rt excluded (loc, rt) = + let new_rt = Loc.tag ~loc @@ match rt with | GRef _ | GVar _ | GEvar _ | GPatVar _ -> rt - | GLambda(loc,Anonymous,k,t,b) -> + | GLambda(Anonymous,k,t,b) -> let new_id = Namegen.next_ident_away (Id.of_string "_x") excluded in let new_excluded = new_id :: excluded in let new_t = alpha_rt new_excluded t in let new_b = alpha_rt new_excluded b in - GLambda(loc,Name new_id,k,new_t,new_b) - | GProd(loc,Anonymous,k,t,b) -> + GLambda(Name new_id,k,new_t,new_b) + | GProd(Anonymous,k,t,b) -> let new_t = alpha_rt excluded t in let new_b = alpha_rt excluded b in - GProd(loc,Anonymous,k,new_t,new_b) - | GLetIn(loc,Anonymous,b,t,c) -> + GProd(Anonymous,k,new_t,new_b) + | GLetIn(Anonymous,b,t,c) -> let new_b = alpha_rt excluded b in let new_t = Option.map (alpha_rt excluded) t in let new_c = alpha_rt excluded c in - GLetIn(loc,Anonymous,new_b,new_t,new_c) - | GLambda(loc,Name id,k,t,b) -> + GLetIn(Anonymous,new_b,new_t,new_c) + | GLambda(Name id,k,t,b) -> let new_id = Namegen.next_ident_away id excluded in let t,b = if Id.equal new_id id - then t,b + then t, b else let replace = change_vars (Id.Map.add id new_id Id.Map.empty) in (t,replace b) @@ -290,8 +285,8 @@ let rec alpha_rt excluded rt = let new_excluded = new_id::excluded in let new_t = alpha_rt new_excluded t in let new_b = alpha_rt new_excluded b in - GLambda(loc,Name new_id,k,new_t,new_b) - | GProd(loc,Name id,k,t,b) -> + GLambda(Name new_id,k,new_t,new_b) + | GProd(Name id,k,t,b) -> let new_id = Namegen.next_ident_away id excluded in let new_excluded = new_id::excluded in let t,b = @@ -303,8 +298,8 @@ let rec alpha_rt excluded rt = in let new_t = alpha_rt new_excluded t in let new_b = alpha_rt new_excluded b in - GProd(loc,Name new_id,k,new_t,new_b) - | GLetIn(loc,Name id,b,t,c) -> + GProd(Name new_id,k,new_t,new_b) + | GLetIn(Name id,b,t,c) -> let new_id = Namegen.next_ident_away id excluded in let c = if Id.equal new_id id then c @@ -314,10 +309,9 @@ let rec alpha_rt excluded rt = let new_b = alpha_rt new_excluded b in let new_t = Option.map (alpha_rt new_excluded) t in let new_c = alpha_rt new_excluded c in - GLetIn(loc,Name new_id,new_b,new_t,new_c) - + GLetIn(Name new_id,new_b,new_t,new_c) - | GLetTuple(loc,nal,(na,rto),t,b) -> + | GLetTuple(nal,(na,rto),t,b) -> let rev_new_nal,new_excluded,mapping = List.fold_left (fun (nal,excluded,mapping) na -> @@ -344,14 +338,14 @@ let rec alpha_rt excluded rt = let new_t = alpha_rt new_excluded new_t in let new_b = alpha_rt new_excluded new_b in let new_rto = Option.map (alpha_rt new_excluded) new_rto in - GLetTuple(loc,new_nal,(na,new_rto),new_t,new_b) - | GCases(loc,sty,infos,el,brl) -> + GLetTuple(new_nal,(na,new_rto),new_t,new_b) + | GCases(sty,infos,el,brl) -> let new_el = List.map (function (rt,i) -> alpha_rt excluded rt, i) el in - GCases(loc,sty,infos,new_el,List.map (alpha_br excluded) brl) - | GIf(loc,b,(na,e_o),lhs,rhs) -> - GIf(loc,alpha_rt excluded b, + GCases(sty,infos,new_el,List.map (alpha_br excluded) brl) + | GIf(b,(na,e_o),lhs,rhs) -> + GIf(alpha_rt excluded b, (na,Option.map (alpha_rt excluded) e_o), alpha_rt excluded lhs, alpha_rt excluded rhs @@ -359,66 +353,65 @@ let rec alpha_rt excluded rt = | GRec _ -> error "Not handled GRec" | GSort _ -> rt | GHole _ -> rt - | GCast (loc,b,c) -> - GCast(loc,alpha_rt excluded b, + | GCast (b,c) -> + GCast(alpha_rt excluded b, Miscops.map_cast_type (alpha_rt excluded) c) - | GApp(loc,f,args) -> - GApp(loc, - alpha_rt excluded f, + | GApp(f,args) -> + GApp(alpha_rt excluded f, List.map (alpha_rt excluded) args ) in new_rt -and alpha_br excluded (loc,ids,patl,res) = +and alpha_br excluded (loc,(ids,patl,res)) = let new_patl,new_excluded,mapping = alpha_patl excluded patl in let new_ids = List.fold_right raw_get_pattern_id new_patl [] in let new_excluded = new_ids@excluded in let renamed_res = change_vars mapping res in let new_res = alpha_rt new_excluded renamed_res in - (loc,new_ids,new_patl,new_res) + (loc,(new_ids,new_patl,new_res)) (* [is_free_in id rt] checks if [id] is a free variable in [rt] *) let is_free_in id = - let rec is_free_in = function + let rec is_free_in (loc, gt) = match gt with | GRef _ -> false - | GVar(_,id') -> Id.compare id' id == 0 + | GVar id' -> Id.compare id' id == 0 | GEvar _ -> false | GPatVar _ -> false - | GApp(_,rt,rtl) -> List.exists is_free_in (rt::rtl) - | GLambda(_,n,_,t,b) | GProd(_,n,_,t,b) -> + | GApp(rt,rtl) -> List.exists is_free_in (rt::rtl) + | GLambda(n,_,t,b) | GProd(n,_,t,b) -> let check_in_b = match n with | Name id' -> not (Id.equal id' id) | _ -> true in is_free_in t || (check_in_b && is_free_in b) - | GLetIn(_,n,b,t,c) -> + | GLetIn(n,b,t,c) -> let check_in_c = match n with | Name id' -> not (Id.equal id' id) | _ -> true in is_free_in b || Option.cata is_free_in true t || (check_in_c && is_free_in c) - | GCases(_,_,_,el,brl) -> + | GCases(_,_,el,brl) -> (List.exists (fun (e,_) -> is_free_in e) el) || List.exists is_free_in_br brl - | GLetTuple(_,nal,_,b,t) -> + | GLetTuple(nal,_,b,t) -> let check_in_nal = not (List.exists (function Name id' -> Id.equal id' id | _ -> false) nal) in is_free_in t || (check_in_nal && is_free_in b) - | GIf(_,cond,_,br1,br2) -> + | GIf(cond,_,br1,br2) -> is_free_in cond || is_free_in br1 || is_free_in br2 | GRec _ -> raise (UserError(None,str "Not handled GRec")) | GSort _ -> false | GHole _ -> false - | GCast (_,b,(CastConv t|CastVM t|CastNative t)) -> is_free_in b || is_free_in t - | GCast (_,b,CastCoerce) -> is_free_in b - and is_free_in_br (_,ids,_,rt) = + | GCast (b,(CastConv t|CastVM t|CastNative t)) -> is_free_in b || is_free_in t + | GCast (b,CastCoerce) -> is_free_in b + and is_free_in_br (_,(ids,_,rt)) = (not (Id.List.mem id ids)) && is_free_in rt in is_free_in @@ -452,60 +445,55 @@ let rec pattern_to_term pt = Loc.with_unloc (function let replace_var_by_term x_id term = - let rec replace_var_by_pattern rt = + let rec replace_var_by_pattern (loc, rt) = Loc.tag ~loc @@ match rt with | GRef _ -> rt - | GVar(_,id) when Id.compare id x_id == 0 -> term + | GVar id when Id.compare id x_id == 0 -> Loc.obj term | GVar _ -> rt | GEvar _ -> rt | GPatVar _ -> rt - | GApp(loc,rt',rtl) -> - GApp(loc, - replace_var_by_pattern rt', + | GApp(rt',rtl) -> + GApp(replace_var_by_pattern rt', List.map replace_var_by_pattern rtl ) - | GLambda(_,Name id,_,_,_) when Id.compare id x_id == 0 -> rt - | GLambda(loc,name,k,t,b) -> - GLambda(loc, - name, + | GLambda(Name id,_,_,_) when Id.compare id x_id == 0 -> rt + | GLambda(name,k,t,b) -> + GLambda(name, k, replace_var_by_pattern t, replace_var_by_pattern b ) - | GProd(_,Name id,_,_,_) when Id.compare id x_id == 0 -> rt - | GProd(loc,name,k,t,b) -> - GProd(loc, - name, + | GProd(Name id,_,_,_) when Id.compare id x_id == 0 -> rt + | GProd(name,k,t,b) -> + GProd( name, k, replace_var_by_pattern t, replace_var_by_pattern b ) - | GLetIn(_,Name id,_,_,_) when Id.compare id x_id == 0 -> rt - | GLetIn(loc,name,def,typ,b) -> - GLetIn(loc, - name, + | GLetIn(Name id,_,_,_) when Id.compare id x_id == 0 -> rt + | GLetIn(name,def,typ,b) -> + GLetIn(name, replace_var_by_pattern def, Option.map (replace_var_by_pattern) typ, replace_var_by_pattern b ) - | GLetTuple(_,nal,_,_,_) + | GLetTuple(nal,_,_,_) when List.exists (function Name id -> Id.equal id x_id | _ -> false) nal -> rt - | GLetTuple(loc,nal,(na,rto),def,b) -> - GLetTuple(loc, - nal, + | GLetTuple(nal,(na,rto),def,b) -> + GLetTuple(nal, (na,Option.map replace_var_by_pattern rto), replace_var_by_pattern def, replace_var_by_pattern b ) - | GCases(loc,sty,infos,el,brl) -> - GCases(loc,sty, + | GCases(sty,infos,el,brl) -> + GCases(sty, infos, List.map (fun (e,x) -> (replace_var_by_pattern e,x)) el, List.map replace_var_by_pattern_br brl ) - | GIf(loc,b,(na,e_option),lhs,rhs) -> - GIf(loc, replace_var_by_pattern b, + | GIf(b,(na,e_option),lhs,rhs) -> + GIf(replace_var_by_pattern b, (na,Option.map replace_var_by_pattern e_option), replace_var_by_pattern lhs, replace_var_by_pattern rhs @@ -513,13 +501,13 @@ let replace_var_by_term x_id term = | GRec _ -> raise (UserError(None,str "Not handled GRec")) | GSort _ -> rt | GHole _ -> rt - | GCast(loc,b,c) -> - GCast(loc,replace_var_by_pattern b, + | GCast(b,c) -> + GCast(replace_var_by_pattern b, Miscops.map_cast_type replace_var_by_pattern c) - and replace_var_by_pattern_br ((loc,idl,patl,res) as br) = + and replace_var_by_pattern_br ((loc,(idl,patl,res)) as br) = if List.exists (fun id -> Id.compare id x_id == 0) idl then br - else (loc,idl,patl,replace_var_by_pattern res) + else (loc,(idl,patl,replace_var_by_pattern res)) in replace_var_by_pattern @@ -590,22 +578,22 @@ let id_of_name = function (* TODO: finish Rec caes *) let ids_of_glob_constr c = - let rec ids_of_glob_constr acc c = + let rec ids_of_glob_constr acc (loc, c) = let idof = id_of_name in match c with - | GVar (_,id) -> id::acc - | GApp (loc,g,args) -> + | GVar id -> id::acc + | GApp (g,args) -> ids_of_glob_constr [] g @ List.flatten (List.map (ids_of_glob_constr []) args) @ acc - | GLambda (loc,na,k,ty,c) -> idof na :: ids_of_glob_constr [] ty @ ids_of_glob_constr [] c @ acc - | GProd (loc,na,k,ty,c) -> idof na :: ids_of_glob_constr [] ty @ ids_of_glob_constr [] c @ acc - | GLetIn (loc,na,b,t,c) -> idof na :: ids_of_glob_constr [] b @ Option.cata (ids_of_glob_constr []) [] t @ ids_of_glob_constr [] c @ acc - | GCast (loc,c,(CastConv t|CastVM t|CastNative t)) -> ids_of_glob_constr [] c @ ids_of_glob_constr [] t @ acc - | GCast (loc,c,CastCoerce) -> ids_of_glob_constr [] c @ acc - | GIf (loc,c,(na,po),b1,b2) -> ids_of_glob_constr [] c @ ids_of_glob_constr [] b1 @ ids_of_glob_constr [] b2 @ acc - | GLetTuple (_,nal,(na,po),b,c) -> + | GLambda (na,k,ty,c) -> idof na :: ids_of_glob_constr [] ty @ ids_of_glob_constr [] c @ acc + | GProd (na,k,ty,c) -> idof na :: ids_of_glob_constr [] ty @ ids_of_glob_constr [] c @ acc + | GLetIn (na,b,t,c) -> idof na :: ids_of_glob_constr [] b @ Option.cata (ids_of_glob_constr []) [] t @ ids_of_glob_constr [] c @ acc + | GCast (c,(CastConv t|CastVM t|CastNative t)) -> ids_of_glob_constr [] c @ ids_of_glob_constr [] t @ acc + | GCast (c,CastCoerce) -> ids_of_glob_constr [] c @ acc + | GIf (c,(na,po),b1,b2) -> ids_of_glob_constr [] c @ ids_of_glob_constr [] b1 @ ids_of_glob_constr [] b2 @ acc + | GLetTuple (nal,(na,po),b,c) -> List.map idof nal @ ids_of_glob_constr [] b @ ids_of_glob_constr [] c @ acc - | GCases (loc,sty,rtntypopt,tml,brchl) -> - List.flatten (List.map (fun (_,idl,patl,c) -> idl @ ids_of_glob_constr [] c) brchl) + | GCases (sty,rtntypopt,tml,brchl) -> + List.flatten (List.map (fun (_,(idl,patl,c)) -> idl @ ids_of_glob_constr [] c) brchl) | GRec _ -> failwith "Fix inside a constructor branch" | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> [] in @@ -617,49 +605,46 @@ let ids_of_glob_constr c = let zeta_normalize = - let rec zeta_normalize_term rt = + let rec zeta_normalize_term (loc, rt) = Loc.tag ~loc @@ match rt with | GRef _ -> rt | GVar _ -> rt | GEvar _ -> rt | GPatVar _ -> rt - | GApp(loc,rt',rtl) -> - GApp(loc, - zeta_normalize_term rt', + | GApp(rt',rtl) -> + GApp(zeta_normalize_term rt', List.map zeta_normalize_term rtl ) - | GLambda(loc,name,k,t,b) -> - GLambda(loc, - name, + | GLambda(name,k,t,b) -> + GLambda(name, k, zeta_normalize_term t, zeta_normalize_term b ) - | GProd(loc,name,k,t,b) -> - GProd(loc, - name, + | GProd(name,k,t,b) -> + GProd(name, k, zeta_normalize_term t, zeta_normalize_term b ) - | GLetIn(_,Name id,def,typ,b) -> - zeta_normalize_term (replace_var_by_term id def b) - | GLetIn(loc,Anonymous,def,typ,b) -> zeta_normalize_term b - | GLetTuple(loc,nal,(na,rto),def,b) -> - GLetTuple(loc, - nal, + | GLetIn(Name id,def,typ,b) -> + Loc.obj @@ zeta_normalize_term (replace_var_by_term id def b) + | GLetIn(Anonymous,def,typ,b) -> + Loc.obj @@ zeta_normalize_term b + | GLetTuple(nal,(na,rto),def,b) -> + GLetTuple(nal, (na,Option.map zeta_normalize_term rto), zeta_normalize_term def, zeta_normalize_term b ) - | GCases(loc,sty,infos,el,brl) -> - GCases(loc,sty, + | GCases(sty,infos,el,brl) -> + GCases(sty, infos, List.map (fun (e,x) -> (zeta_normalize_term e,x)) el, List.map zeta_normalize_br brl ) - | GIf(loc,b,(na,e_option),lhs,rhs) -> - GIf(loc, zeta_normalize_term b, + | GIf(b,(na,e_option),lhs,rhs) -> + GIf(zeta_normalize_term b, (na,Option.map zeta_normalize_term e_option), zeta_normalize_term lhs, zeta_normalize_term rhs @@ -667,11 +652,11 @@ let zeta_normalize = | GRec _ -> raise (UserError(None,str "Not handled GRec")) | GSort _ -> rt | GHole _ -> rt - | GCast(loc,b,c) -> - GCast(loc,zeta_normalize_term b, + | GCast(b,c) -> + GCast(zeta_normalize_term b, Miscops.map_cast_type zeta_normalize_term c) - and zeta_normalize_br (loc,idl,patl,res) = - (loc,idl,patl,zeta_normalize_term res) + and zeta_normalize_br (loc,(idl,patl,res)) = + (loc,(idl,patl,zeta_normalize_term res)) in zeta_normalize_term @@ -687,33 +672,34 @@ let expand_as = Id.Map.add id (pattern_to_term (loc, pat)) (List.fold_left add_as map patl) | PatCstr(_,patl,_) -> List.fold_left add_as map patl in - let rec expand_as map rt = + let rec expand_as map (loc, rt) = + Loc.tag ~loc @@ match rt with | GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ -> rt - | GVar(_,id) -> + | GVar id -> begin try - Id.Map.find id map + Loc.obj @@ Id.Map.find id map with Not_found -> rt end - | GApp(loc,f,args) -> GApp(loc,expand_as map f,List.map (expand_as map) args) - | GLambda(loc,na,k,t,b) -> GLambda(loc,na,k,expand_as map t, expand_as map b) - | GProd(loc,na,k,t,b) -> GProd(loc,na,k,expand_as map t, expand_as map b) - | GLetIn(loc,na,v,typ,b) -> GLetIn(loc,na, expand_as map v,Option.map (expand_as map) typ,expand_as map b) - | GLetTuple(loc,nal,(na,po),v,b) -> - GLetTuple(loc,nal,(na,Option.map (expand_as map) po), + | GApp(f,args) -> GApp(expand_as map f,List.map (expand_as map) args) + | GLambda(na,k,t,b) -> GLambda(na,k,expand_as map t, expand_as map b) + | GProd(na,k,t,b) -> GProd(na,k,expand_as map t, expand_as map b) + | GLetIn(na,v,typ,b) -> GLetIn(na, expand_as map v,Option.map (expand_as map) typ,expand_as map b) + | GLetTuple(nal,(na,po),v,b) -> + GLetTuple(nal,(na,Option.map (expand_as map) po), expand_as map v, expand_as map b) - | GIf(loc,e,(na,po),br1,br2) -> - GIf(loc,expand_as map e,(na,Option.map (expand_as map) po), + | GIf(e,(na,po),br1,br2) -> + GIf(expand_as map e,(na,Option.map (expand_as map) po), expand_as map br1, expand_as map br2) | GRec _ -> error "Not handled GRec" - | GCast(loc,b,c) -> - GCast(loc,expand_as map b, + | GCast(b,c) -> + GCast(expand_as map b, Miscops.map_cast_type (expand_as map) c) - | GCases(loc,sty,po,el,brl) -> - GCases(loc, sty, Option.map (expand_as map) po, List.map (fun (rt,t) -> expand_as map rt,t) el, + | GCases(sty,po,el,brl) -> + GCases(sty, Option.map (expand_as map) po, List.map (fun (rt,t) -> expand_as map rt,t) el, List.map (expand_as_br map) brl) - and expand_as_br map (loc,idl,cpl,rt) = - (loc,idl,cpl, expand_as (List.fold_left add_as map cpl) rt) + and expand_as_br map (loc,(idl,cpl,rt)) = + (loc,(idl,cpl, expand_as (List.fold_left add_as map cpl) rt)) in expand_as Id.Map.empty diff --git a/plugins/funind/glob_termops.mli b/plugins/funind/glob_termops.mli index 84359a36b..25d79582f 100644 --- a/plugins/funind/glob_termops.mli +++ b/plugins/funind/glob_termops.mli @@ -82,11 +82,8 @@ val alpha_rt : Id.t list -> glob_constr -> glob_constr (* same as alpha_rt but for case branches *) val alpha_br : Id.t list -> - Loc.t * Id.t list * Glob_term.cases_pattern list * - Glob_term.glob_constr -> - Loc.t * Id.t list * Glob_term.cases_pattern list * - Glob_term.glob_constr - + Glob_term.cases_clause -> + Glob_term.cases_clause (* Reduction function *) val replace_var_by_term : diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index 6ee755323..cad96946c 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -190,18 +190,18 @@ let build_newrecursive l = let is_rec names = let names = List.fold_right Id.Set.add names Id.Set.empty in let check_id id names = Id.Set.mem id names in - let rec lookup names = function - | GVar(_,id) -> check_id id names + let rec lookup names (loc, gt) = match gt with + | GVar(id) -> check_id id names | GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ -> false - | GCast(_,b,_) -> lookup names b + | GCast(b,_) -> lookup names b | GRec _ -> error "GRec not handled" - | GIf(_,b,_,lhs,rhs) -> + | GIf(b,_,lhs,rhs) -> (lookup names b) || (lookup names lhs) || (lookup names rhs) - | GProd(_,na,_,t,b) | GLambda(_,na,_,t,b) -> + | GProd(na,_,t,b) | GLambda(na,_,t,b) -> lookup names t || lookup (Nameops.name_fold Id.Set.remove na names) b - | GLetIn(_,na,b,t,c) -> + | GLetIn(na,b,t,c) -> lookup names b || Option.cata (lookup names) true t || lookup (Nameops.name_fold Id.Set.remove na names) c - | GLetTuple(_,nal,_,t,b) -> lookup names t || + | GLetTuple(nal,_,t,b) -> lookup names t || lookup (List.fold_left (fun acc na -> Nameops.name_fold Id.Set.remove na acc) @@ -209,11 +209,11 @@ let is_rec names = nal ) b - | GApp(_,f,args) -> List.exists (lookup names) (f::args) - | GCases(_,_,_,el,brl) -> + | GApp(f,args) -> List.exists (lookup names) (f::args) + | GCases(_,_,el,brl) -> List.exists (fun (e,_) -> lookup names e) el || List.exists (lookup_br names) brl - and lookup_br names (_,idl,_,rt) = + and lookup_br names (_,(idl,_,rt)) = let new_names = List.fold_right Id.Set.remove idl names in lookup new_names rt in diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml index 7b0d7d27d..8f0c98624 100644 --- a/plugins/funind/indfun_common.ml +++ b/plugins/funind/indfun_common.ml @@ -69,9 +69,9 @@ let chop_rlambda_n = if n == 0 then List.rev acc,rt else - match rt with - | Glob_term.GLambda(_,name,k,t,b) -> chop_lambda_n ((name,t,None)::acc) (n-1) b - | Glob_term.GLetIn(_,name,v,t,b) -> chop_lambda_n ((name,v,t)::acc) (n-1) b + match Loc.obj rt with + | Glob_term.GLambda(name,k,t,b) -> chop_lambda_n ((name,t,None)::acc) (n-1) b + | Glob_term.GLetIn(name,v,t,b) -> chop_lambda_n ((name,v,t)::acc) (n-1) b | _ -> raise (CErrors.UserError(Some "chop_rlambda_n", str "chop_rlambda_n: Not enough Lambdas")) @@ -83,8 +83,8 @@ let chop_rprod_n = if n == 0 then List.rev acc,rt else - match rt with - | Glob_term.GProd(_,name,k,t,b) -> chop_prod_n ((name,t)::acc) (n-1) b + match Loc.obj rt with + | Glob_term.GProd(name,k,t,b) -> chop_prod_n ((name,t)::acc) (n-1) b | _ -> raise (CErrors.UserError(Some "chop_rprod_n",str "chop_rprod_n: Not enough products")) in chop_prod_n [] diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml index 29de56478..9dc1d48df 100644 --- a/plugins/funind/merge.ml +++ b/plugins/funind/merge.ml @@ -66,7 +66,7 @@ let string_of_name = id_of_name %> Id.to_string (** [isVarf f x] returns [true] if term [x] is of the form [(Var f)]. *) let isVarf f x = match x with - | GVar (_,x) -> Id.equal x f + | _, GVar x -> Id.equal x f | _ -> false (** [ident_global_exist id] returns true if identifier [id] is linked @@ -504,40 +504,40 @@ let shift_linked_params mib1 mib2 (lnk1:linked_var array) (lnk2:linked_var array exception NoMerge -let rec merge_app c1 c2 id1 id2 shift filter_shift_stable = +let rec merge_app (loc1, c1) (loc2, c2) id1 id2 shift filter_shift_stable = let lnk = Array.append shift.lnk1 shift.lnk2 in - match c1 , c2 with - | GApp(_,f1, arr1), GApp(_,f2,arr2) when isVarf id1 f1 && isVarf id2 f2 -> + match c1, c2 with + | GApp(f1, arr1), GApp(f2,arr2) when isVarf id1 f1 && isVarf id2 f2 -> let _ = prstr "\nICI1!\n" in let args = filter_shift_stable lnk (arr1 @ arr2) in - GApp (Loc.ghost,GVar (Loc.ghost,shift.ident) , args) - | GApp(_,f1, arr1), GApp(_,f2,arr2) -> raise NoMerge - | GLetIn(_,nme,bdy,typ,trm) , _ -> + Loc.tag @@ GApp ((Loc.tag @@ GVar shift.ident) , args) + | GApp(f1, arr1), GApp(f2,arr2) -> raise NoMerge + | GLetIn(nme,bdy,typ,trm) , _ -> let _ = prstr "\nICI2!\n" in - let newtrm = merge_app trm c2 id1 id2 shift filter_shift_stable in - GLetIn(Loc.ghost,nme,bdy,typ,newtrm) - | _, GLetIn(_,nme,bdy,typ,trm) -> + let newtrm = merge_app trm (loc2, c2) id1 id2 shift filter_shift_stable in + Loc.tag @@ GLetIn(nme,bdy,typ,newtrm) + | _, GLetIn(nme,bdy,typ,trm) -> let _ = prstr "\nICI3!\n" in - let newtrm = merge_app c1 trm id1 id2 shift filter_shift_stable in - GLetIn(Loc.ghost,nme,bdy,typ,newtrm) + let newtrm = merge_app (loc1, c1) trm id1 id2 shift filter_shift_stable in + Loc.tag @@ GLetIn(nme,bdy,typ,newtrm) | _ -> let _ = prstr "\nICI4!\n" in raise NoMerge -let rec merge_app_unsafe c1 c2 shift filter_shift_stable = +let rec merge_app_unsafe (l1, c1) (l2, c2) shift filter_shift_stable = let lnk = Array.append shift.lnk1 shift.lnk2 in match c1 , c2 with - | GApp(_,f1, arr1), GApp(_,f2,arr2) -> + | GApp(f1, arr1), GApp(f2,arr2) -> let args = filter_shift_stable lnk (arr1 @ arr2) in - GApp (Loc.ghost,GVar(Loc.ghost,shift.ident) , args) + Loc.tag @@ GApp (Loc.tag @@ GVar shift.ident, args) (* FIXME: what if the function appears in the body of the let? *) - | GLetIn(_,nme,bdy,typ,trm) , _ -> + | GLetIn(nme,bdy,typ,trm) , _ -> let _ = prstr "\nICI2 '!\n" in - let newtrm = merge_app_unsafe trm c2 shift filter_shift_stable in - GLetIn(Loc.ghost,nme,bdy,typ,newtrm) - | _, GLetIn(_,nme,bdy,typ,trm) -> + let newtrm = merge_app_unsafe trm (l2, c2) shift filter_shift_stable in + Loc.tag @@ GLetIn(nme,bdy,typ,newtrm) + | _, GLetIn(nme,bdy,typ,trm) -> let _ = prstr "\nICI3 '!\n" in - let newtrm = merge_app_unsafe c1 trm shift filter_shift_stable in - GLetIn(Loc.ghost,nme,bdy,typ,newtrm) + let newtrm = merge_app_unsafe (l1, c1) trm shift filter_shift_stable in + Loc.tag @@ GLetIn(nme,bdy,typ,newtrm) | _ -> let _ = prstr "\nICI4 '!\n" in raise NoMerge @@ -550,14 +550,14 @@ let rec merge_rec_hyps shift accrec filter_shift_stable : (Name.t * glob_constr option * glob_constr option) list = let mergeonehyp t reldecl = match reldecl with - | (nme,x,Some (GApp(_,i,args) as ind)) + | (nme,x,Some (_, GApp(i,args) as ind)) -> nme,x, Some (merge_app_unsafe ind t shift filter_shift_stable) | (nme,Some _,None) -> error "letins with recursive calls not treated yet" | (nme,None,Some _) -> assert false | (nme,None,None) | (nme,Some _,Some _) -> assert false in match ltyp with | [] -> [] - | (nme,None,Some (GApp(_,f, largs) as t)) :: lt when isVarf ind2name f -> + | (nme,None,Some (_, GApp(f, largs) as t)) :: lt when isVarf ind2name f -> let rechyps = List.map (mergeonehyp t) accrec in rechyps @ merge_rec_hyps shift accrec lt filter_shift_stable | e::lt -> e :: merge_rec_hyps shift accrec lt filter_shift_stable @@ -573,7 +573,7 @@ let find_app (nme:Id.t) ltyp = (List.map (fun x -> match x with - | _,None,Some (GApp(_,f,_)) when isVarf nme f -> raise (Found 0) + | _,None,Some (_, GApp(f,_)) when isVarf nme f -> raise (Found 0) | _ -> ()) ltyp); false @@ -633,7 +633,7 @@ let rec merge_types shift accrec1 rechyps , concl | (nme,None, Some t1)as e ::lt1 -> (match t1 with - | GApp(_,f,carr) when isVarf ind1name f -> + | _, GApp(f,carr) when isVarf ind1name f -> merge_types shift (e::accrec1) lt1 concl1 ltyp2 concl2 | _ -> let recres, recconcl2 = @@ -864,7 +864,7 @@ let mkProd_reldecl (rdecl:Context.Rel.Declaration.t) (t2:glob_constr) = | LocalAssum (nme,t) -> let t = EConstr.of_constr t in let traw = Detyping.detype false [] (Global.env()) Evd.empty t in - GProd (Loc.ghost,nme,Explicit,traw,t2) + Loc.tag @@ GProd (nme,Explicit,traw,t2) | LocalDef _ -> assert false diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index ee7b33227..c796fe7a2 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -172,7 +172,6 @@ let simpl_iter clause = let (value_f:Constr.constr list -> global_reference -> Constr.constr) = let open Term in fun al fterm -> - let d0 = Loc.ghost in let rev_x_id_l = ( List.fold_left @@ -189,16 +188,15 @@ let (value_f:Constr.constr list -> global_reference -> Constr.constr) = in let env = Environ.push_rel_context context (Global.env ()) in let glob_body = - GCases - (d0,RegularStyle,None, - [GApp(d0, GRef(d0,fterm,None), List.rev_map (fun x_id -> GVar(d0, x_id)) rev_x_id_l), + Loc.tag @@ + GCases + (RegularStyle,None, + [Loc.tag @@ GApp(Loc.tag @@ GRef(fterm,None), List.rev_map (fun x_id -> Loc.tag @@ GVar x_id) rev_x_id_l), (Anonymous,None)], - [d0, [v_id], [d0,PatCstr((destIndRef - (delayed_force coq_sig_ref),1), - [d0, PatVar(Name v_id); - d0, PatVar(Anonymous)], - Anonymous)], - GVar(d0,v_id)]) + [Loc.tag ([v_id], [Loc.tag @@ PatCstr ((destIndRef (delayed_force coq_sig_ref),1), + [Loc.tag @@ PatVar(Name v_id); Loc.tag @@ PatVar Anonymous], + Anonymous)], + Loc.tag @@ GVar v_id)]) in let body = fst (understand env (Evd.from_env env) glob_body)(*FIXME*) in it_mkLambda_or_LetIn body context diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4 index 38fdfb759..232bd851f 100644 --- a/plugins/ltac/extratactics.ml4 +++ b/plugins/ltac/extratactics.ml4 @@ -631,15 +631,15 @@ let subst_var_with_hole occ tid t = let occref = if occ > 0 then ref occ else Find_subterm.error_invalid_occurrence [occ] in let locref = ref 0 in let rec substrec = function - | GVar (_,id) as x -> + | (_, GVar id) as x -> if Id.equal id tid then (decr occref; if Int.equal !occref 0 then x else (incr locref; - GHole (Loc.make_loc (!locref,0), - Evar_kinds.QuestionMark(Evar_kinds.Define true), + Loc.tag ~loc:(Loc.make_loc (!locref,0)) @@ + GHole (Evar_kinds.QuestionMark(Evar_kinds.Define true), Misctypes.IntroAnonymous, None))) else x | c -> map_glob_constr_left_to_right substrec c in @@ -651,13 +651,13 @@ let subst_hole_with_term occ tc t = let locref = ref 0 in let occref = ref occ in let rec substrec = function - | GHole (_,Evar_kinds.QuestionMark(Evar_kinds.Define true),Misctypes.IntroAnonymous,s) -> + | _, GHole (Evar_kinds.QuestionMark(Evar_kinds.Define true),Misctypes.IntroAnonymous,s) -> decr occref; if Int.equal !occref 0 then tc else (incr locref; - GHole (Loc.make_loc (!locref,0), - Evar_kinds.QuestionMark(Evar_kinds.Define true),Misctypes.IntroAnonymous,s)) + Loc.tag ~loc:(Loc.make_loc (!locref,0)) @@ + GHole (Evar_kinds.QuestionMark(Evar_kinds.Define true),Misctypes.IntroAnonymous,s)) | c -> map_glob_constr_left_to_right substrec c in substrec t diff --git a/plugins/ltac/g_rewrite.ml4 b/plugins/ltac/g_rewrite.ml4 index c50100bf5..1f40c67b5 100644 --- a/plugins/ltac/g_rewrite.ml4 +++ b/plugins/ltac/g_rewrite.ml4 @@ -123,7 +123,7 @@ TACTIC EXTEND rewrite_strat END let clsubstitute o c = - let is_tac id = match fst (fst (snd c)) with GVar (_, id') when Id.equal id' id -> true | _ -> false in + let is_tac id = match fst (fst (snd c)) with (_, GVar id') when Id.equal id' id -> true | _ -> false in Tacticals.onAllHypsAndConcl (fun cl -> match cl with diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml index ad76ef9c6..aec2e37fd 100644 --- a/plugins/ltac/pptactic.ml +++ b/plugins/ltac/pptactic.ml @@ -1085,8 +1085,8 @@ type 'a extra_genarg_printer = let strip_prod_binders_glob_constr n (ty,_) = let rec strip_ty acc n ty = if Int.equal n 0 then (List.rev acc, (ty,None)) else - match ty with - Glob_term.GProd(loc,na,Explicit,a,b) -> + match Loc.obj ty with + Glob_term.GProd(na,Explicit,a,b) -> strip_ty (([Loc.ghost,na],(a,None))::acc) (n-1) b | _ -> error "Cannot translate fix tactic: not enough products" in strip_ty [] n ty diff --git a/plugins/ltac/tacintern.ml b/plugins/ltac/tacintern.ml index 75f890c96..e7d4c1be9 100644 --- a/plugins/ltac/tacintern.ml +++ b/plugins/ltac/tacintern.ml @@ -31,8 +31,6 @@ open Locus (** Globalization of tactic expressions : Conversion from [raw_tactic_expr] to [glob_tactic_expr] *) -let dloc = Loc.ghost - let error_tactic_expected ?loc = user_err ?loc (str "Tactic expected.") @@ -74,14 +72,14 @@ let intern_name l ist = function let strict_check = ref false -let adjust_loc loc = if !strict_check then dloc else loc +let adjust_loc loc = if !strict_check then Loc.ghost else loc (* Globalize a name which must be bound -- actually just check it is bound *) let intern_hyp ist (loc,id as locid) = if not !strict_check then locid else if find_ident id ist then - (dloc,id) + Loc.tag id else Pretype_errors.error_var_not_found ~loc id @@ -110,12 +108,12 @@ let intern_ltac_variable ist = function let intern_constr_reference strict ist = function | Ident (_,id) as r when not strict && find_hyp id ist -> - GVar (dloc,id), Some (Loc.tag @@ CRef (r,None)) + (Loc.tag @@ GVar id), Some (Loc.tag @@ CRef (r,None)) | Ident (_,id) as r when find_var id ist -> - GVar (dloc,id), if strict then None else Some (Loc.tag @@ CRef (r,None)) + (Loc.tag @@ GVar id), if strict then None else Some (Loc.tag @@ CRef (r,None)) | r -> let loc,_ as lqid = qualid_of_reference r in - GRef (loc,locate_global_with_alias lqid,None), + Loc.tag @@ GRef (locate_global_with_alias lqid,None), if strict then None else Some (Loc.tag @@ CRef (r,None)) let intern_move_location ist = function @@ -273,8 +271,8 @@ let intern_destruction_arg ist = function | clear,ElimOnIdent (loc,id) -> if !strict_check then (* If in a defined tactic, no intros-until *) - match intern_constr ist (Loc.tag @@ CRef (Ident (dloc,id), None)) with - | GVar (loc,id),_ -> clear,ElimOnIdent (loc,id) + match intern_constr ist (Loc.tag @@ CRef (Ident (Loc.tag id), None)) with + | (loc, GVar id), _ -> clear,ElimOnIdent (loc,id) | c -> clear,ElimOnConstr (c,NoBindings) else clear,ElimOnIdent (loc,id) @@ -352,10 +350,10 @@ let intern_typed_pattern_or_ref_with_occurrences ist (l,p) = | _ -> Qualid (loc,qualid_of_path (path_of_global (smart_global r))) in let sign = { Constrintern.ltac_vars = ist.ltacvars; Constrintern.ltac_bound = Id.Set.empty } in let c = Constrintern.interp_reference sign r in - match c with - | GRef (_,r,None) -> + match Loc.obj c with + | GRef (r,None) -> Inl (ArgArg (evaluable_of_global_reference ist.genv r,None)) - | GVar (_,id) -> + | GVar id -> let r = evaluable_of_global_reference ist.genv (VarRef id) in Inl (ArgArg (r,None)) | _ -> diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml index de6c44b2b..a8d8eda1d 100644 --- a/plugins/ltac/tacinterp.ml +++ b/plugins/ltac/tacinterp.ml @@ -699,7 +699,7 @@ let interp_typed_pattern ist env sigma (_,c,_) = let interp_constr_in_compound_list inj_fun dest_fun interp_fun ist env sigma l = let try_expand_ltac_var sigma x = try match dest_fun x with - | GVar (_,id), _ -> + | (_, GVar id), _ -> let v = Id.Map.find id ist.lfun in sigma, List.map inj_fun (coerce_to_constr_list env v) | _ -> @@ -1074,7 +1074,7 @@ let interp_destruction_arg ist gl arg = if Tactics.is_quantified_hypothesis id gl then keep,ElimOnIdent (loc,id) else - let c = (GVar (loc,id),Some (Loc.tag @@ CRef (Ident (loc,id),None))) in + let c = (Loc.tag ~loc @@ GVar id,Some (Loc.tag @@ CRef (Ident (loc,id),None))) in let f = { delayed = fun env sigma -> let sigma = Sigma.to_evar_map sigma in let (sigma,c) = interp_open_constr ist env sigma c in diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index dd68eac24..bef8139be 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -131,7 +131,7 @@ let closed_term_ast l = let l = List.map (fun gr -> ArgArg(Loc.ghost,gr)) l in TacFun([Name(Id.of_string"t")], TacML(Loc.ghost,tacname, - [TacGeneric (Genarg.in_gen (Genarg.glbwit Stdarg.wit_constr) (GVar(Loc.ghost,Id.of_string"t"),None)); + [TacGeneric (Genarg.in_gen (Genarg.glbwit Stdarg.wit_constr) (Loc.tag @@ GVar(Id.of_string"t"),None)); TacGeneric (Genarg.in_gen (Genarg.glbwit (Genarg.wit_list Stdarg.wit_ref)) l)])) (* let _ = add_tacdef false ((Loc.ghost,Id.of_string"ring_closed_term" diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index 1a5ef825d..f8cccf714 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -62,7 +62,6 @@ open Locusops DECLARE PLUGIN "ssrmatching_plugin" type loc = Loc.t -let dummy_loc = Loc.ghost let errorstrm = CErrors.user_err ~hdr:"ssrmatching" let loc_error loc msg = CErrors.user_err ~loc ~hdr:msg (str msg) let ppnl = Feedback.msg_info @@ -159,10 +158,10 @@ let mkCLetIn loc name bo t = Loc.tag ~loc @@ CLetIn ((loc, name), bo, None, t) let mkCCast loc t ty = Loc.tag ~loc @@ CCast (t, dC ty) (** Constructors for rawconstr *) -let mkRHole = GHole (dummy_loc, InternalHole, IntroAnonymous, None) -let mkRApp f args = if args = [] then f else GApp (dummy_loc, f, args) -let mkRCast rc rt = GCast (dummy_loc, rc, dC rt) -let mkRLambda n s t = GLambda (dummy_loc, n, Explicit, s, t) +let mkRHole = Loc.tag @@ GHole (InternalHole, IntroAnonymous, None) +let mkRApp f args = if args = [] then f else Loc.tag @@ GApp (f, args) +let mkRCast rc rt = Loc.tag @@ GCast (rc, dC rt) +let mkRLambda n s t = Loc.tag @@ GLambda (n, Explicit, s, t) (* ssrterm conbinators *) let combineCG t1 t2 f g = match t1, t2 with @@ -944,7 +943,7 @@ let glob_cpattern gs p = let name = Name (id_of_string ("_ssrpat_" ^ s)) in k, (mkRCast mkRHole (mkRLambda name mkRHole (mkRApp mkRHole l)), None) in let bind_in t1 t2 = - let d = dummy_loc in let n = Name (destCVar t1) in + let d = Loc.ghost in let n = Name (destCVar t1) in fst (glob (mkCCast d (mkCHole d) (mkCLambda d n (mkCHole d) t2))) in let check_var t2 = if not (isCVar t2) then loc_error (constr_loc t2) "Only identifiers are allowed here" in @@ -1023,7 +1022,7 @@ type pattern = Evd.evar_map * (constr, constr) ssrpattern let id_of_cpattern = function | _,(_,Some (_loc, CRef (Ident (_, x), _))) -> Some x | _,(_,Some (_loc, CAppExpl ((_, Ident (_, x), _), []))) -> Some x - | _,(GRef (_, VarRef x, _) ,None) -> Some x + | _,((_, GRef (VarRef x, _)) ,None) -> Some x | _ -> None let id_of_Cterm t = match id_of_cpattern t with | Some x -> x @@ -1123,8 +1122,8 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty = let mkG ?(k=' ') x = k,(x,None) in let decode ist t ?reccall f g = try match (pf_intern_term ist gl t) with - | GCast(_,GHole _,CastConv(GLambda(_,Name x,_,_,c))) -> f x (' ',(c,None)) - | GVar(_,id) + | _, GCast((_, GHole _),CastConv((_, GLambda(Name x,_,_,c)))) -> f x (' ',(c,None)) + | _, GVar id when Id.Map.mem id ist.lfun && not(Option.is_empty reccall) && not(Option.is_empty wit_ssrpatternarg) -> @@ -1166,17 +1165,17 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty = sigma new_evars in sigma in let red = let rec decode_red (ist,red) = match red with - | T(k,(GCast (_,GHole _,(CastConv(GLambda (_,Name id,_,_,t)))),None)) + | T(k,((_, GCast ((_, GHole _),CastConv((_, GLambda (Name id,_,_,t))))),None)) when let id = string_of_id id in let len = String.length id in (len > 8 && String.sub id 0 8 = "_ssrpat_") -> let id = string_of_id id in let len = String.length id in (match String.sub id 8 (len - 8), t with - | "In", GApp(_, _, [t]) -> decodeG t xInT (fun x -> T x) - | "In", GApp(_, _, [e; t]) -> decodeG t (eInXInT (mkG e)) (bad_enc id) - | "In", GApp(_, _, [e; t; e_in_t]) -> + | "In", (_, GApp( _, [t])) -> decodeG t xInT (fun x -> T x) + | "In", (_, GApp( _, [e; t])) -> decodeG t (eInXInT (mkG e)) (bad_enc id) + | "In", (_, GApp( _, [e; t; e_in_t])) -> decodeG t (eInXInT (mkG e)) (fun _ -> decodeG e_in_t xInT (fun _ -> assert false)) - | "As", GApp(_, _, [e; t]) -> decodeG t (eAsXInT (mkG e)) (bad_enc id) + | "As", (_, GApp(_, [e; t])) -> decodeG t (eAsXInT (mkG e)) (bad_enc id) | _ -> bad_enc id ()) | T t -> decode ist ~reccall:decode_red t xInT (fun x -> T x) | In_T t -> decode ist t inXInT inT @@ -1202,13 +1201,13 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty = pp(lazy(str"typed as: " ++ pr_pattern_w_ids red)); let mkXLetIn loc x (a,(g,c)) = match c with | Some b -> a,(g,Some (mkCLetIn loc x (mkCHole loc) b)) - | None -> a,(GLetIn (loc,x,(GHole (loc, BinderType x, IntroAnonymous, None)), None, g), None) in + | None -> a,(Loc.tag ~loc @@ GLetIn (x, Loc.tag ~loc @@ GHole (BinderType x, IntroAnonymous, None), None, g), None) in match red with | T t -> let sigma, t = interp_term ist gl t in sigma, T t | In_T t -> let sigma, t = interp_term ist gl t in sigma, In_T t | X_In_T (x, rp) | In_X_In_T (x, rp) -> let mk x p = match red with X_In_T _ -> X_In_T(x,p) | _ -> In_X_In_T(x,p) in - let rp = mkXLetIn dummy_loc (Name x) rp in + let rp = mkXLetIn Loc.ghost (Name x) rp in let sigma, rp = interp_term ist gl rp in let _, h, _, rp = destLetIn rp in let sigma = cleanup_XinE h x rp sigma in @@ -1217,7 +1216,7 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty = | E_In_X_In_T(e, x, rp) | E_As_X_In_T (e, x, rp) -> let mk e x p = match red with E_In_X_In_T _ ->E_In_X_In_T(e,x,p)|_->E_As_X_In_T(e,x,p) in - let rp = mkXLetIn dummy_loc (Name x) rp in + let rp = mkXLetIn Loc.ghost (Name x) rp in let sigma, rp = interp_term ist gl rp in let _, h, _, rp = destLetIn rp in let sigma = cleanup_XinE h x rp sigma in @@ -1375,10 +1374,10 @@ let pf_fill_occ_term gl occ t = let cl,(_,t) = fill_occ_term env concl occ sigma0 t in cl, t -let cpattern_of_id id = ' ', (GRef (dummy_loc, VarRef id, None), None) +let cpattern_of_id id = ' ', (Loc.tag @@ GRef (VarRef id, None), None) -let is_wildcard = function - | _,(_,Some (_, CHole _)|GHole _,None) -> true +let is_wildcard : cpattern -> bool = function + | _,(_,Some (_, CHole _)| (_, GHole _),None) -> true | _ -> false (* "ssrpattern" *) diff --git a/plugins/syntax/ascii_syntax.ml b/plugins/syntax/ascii_syntax.ml index ed8cc6ab0..dc0b87793 100644 --- a/plugins/syntax/ascii_syntax.ml +++ b/plugins/syntax/ascii_syntax.ml @@ -37,13 +37,13 @@ let glob_Ascii = lazy (make_reference "Ascii") open Lazy -let interp_ascii dloc p = +let interp_ascii loc p = let rec aux n p = if Int.equal n 0 then [] else let mp = p mod 2 in - GRef (dloc,(if Int.equal mp 0 then glob_false else glob_true),None) + (Loc.tag ~loc @@ GRef ((if Int.equal mp 0 then glob_false else glob_true),None)) :: (aux (n-1) (p/2)) in - GApp (dloc,GRef(dloc,force glob_Ascii,None), aux 8 p) + Loc.tag ~loc @@ GApp (Loc.tag ~loc @@ GRef(force glob_Ascii,None), aux 8 p) let interp_ascii_string dloc s = let p = @@ -59,12 +59,12 @@ let interp_ascii_string dloc s = let uninterp_ascii r = let rec uninterp_bool_list n = function | [] when Int.equal n 0 -> 0 - | GRef (_,k,_)::l when Globnames.eq_gr k glob_true -> 1+2*(uninterp_bool_list (n-1) l) - | GRef (_,k,_)::l when Globnames.eq_gr k glob_false -> 2*(uninterp_bool_list (n-1) l) + | (_, GRef (k,_))::l when Globnames.eq_gr k glob_true -> 1+2*(uninterp_bool_list (n-1) l) + | (_, GRef (k,_))::l when Globnames.eq_gr k glob_false -> 2*(uninterp_bool_list (n-1) l) | _ -> raise Non_closed_ascii in try let aux = function - | GApp (_,GRef (_,k,_),l) when Globnames.eq_gr k (force glob_Ascii) -> uninterp_bool_list 8 l + | _, GApp ((_, GRef (k,_)),l) when Globnames.eq_gr k (force glob_Ascii) -> uninterp_bool_list 8 l | _ -> raise Non_closed_ascii in Some (aux r) with @@ -80,4 +80,4 @@ let _ = Notation.declare_string_interpreter "char_scope" (ascii_path,ascii_module) interp_ascii_string - ([GRef (Loc.ghost,static_glob_Ascii,None)], uninterp_ascii_string, true) + ([Loc.tag @@ GRef (static_glob_Ascii,None)], uninterp_ascii_string, true) diff --git a/plugins/syntax/nat_syntax.ml b/plugins/syntax/nat_syntax.ml index ab262fea7..90d643b7f 100644 --- a/plugins/syntax/nat_syntax.ml +++ b/plugins/syntax/nat_syntax.ml @@ -33,14 +33,14 @@ let warn_large_nat = strbrk "may vary from 5000 to 70000 depending on your system " ++ strbrk "limits and on the command executed).") -let nat_of_int dloc n = +let nat_of_int loc n = if is_pos_or_zero n then begin if less_than threshold n then warn_large_nat (); - let ref_O = GRef (dloc, glob_O, None) in - let ref_S = GRef (dloc, glob_S, None) in + let ref_O = Loc.tag ~loc @@ GRef (glob_O, None) in + let ref_S = Loc.tag ~loc @@ GRef (glob_S, None) in let rec mk_nat acc n = if n <> zero then - mk_nat (GApp (dloc,ref_S, [acc])) (sub_1 n) + mk_nat (Loc.tag ~loc @@ GApp (ref_S, [acc])) (sub_1 n) else acc in @@ -55,10 +55,11 @@ let nat_of_int dloc n = exception Non_closed_number -let rec int_of_nat = function - | GApp (_,GRef (_,s,_),[a]) when Globnames.eq_gr s glob_S -> add_1 (int_of_nat a) - | GRef (_,z,_) when Globnames.eq_gr z glob_O -> zero +let rec int_of_nat x = Loc.with_unloc (function + | GApp ((_, GRef (s,_)),[a]) when Globnames.eq_gr s glob_S -> add_1 (int_of_nat a) + | GRef (z,_) when Globnames.eq_gr z glob_O -> zero | _ -> raise Non_closed_number + ) x let uninterp_nat p = try @@ -73,4 +74,4 @@ let _ = Notation.declare_numeral_interpreter "nat_scope" (nat_path,datatypes_module_name) nat_of_int - ([GRef (Loc.ghost,glob_S,None); GRef (Loc.ghost,glob_O,None)], uninterp_nat, true) + ([Loc.tag @@ GRef (glob_S,None); Loc.tag @@ GRef (glob_O,None)], uninterp_nat, true) diff --git a/plugins/syntax/numbers_syntax.ml b/plugins/syntax/numbers_syntax.ml index a25ddb062..8876d464a 100644 --- a/plugins/syntax/numbers_syntax.ml +++ b/plugins/syntax/numbers_syntax.ml @@ -86,10 +86,10 @@ exception Non_closed (* parses a *non-negative* integer (from bigint.ml) into an int31 wraps modulo 2^31 *) -let int31_of_pos_bigint dloc n = - let ref_construct = GRef (dloc, int31_construct, None) in - let ref_0 = GRef (dloc, int31_0, None) in - let ref_1 = GRef (dloc, int31_1, None) in +let int31_of_pos_bigint loc n = + let ref_construct = Loc.tag ~loc @@ GRef (int31_construct, None) in + let ref_0 = Loc.tag ~loc @@ GRef (int31_0, None) in + let ref_1 = Loc.tag ~loc @@ GRef (int31_1, None) in let rec args counter n = if counter <= 0 then [] @@ -97,7 +97,7 @@ let int31_of_pos_bigint dloc n = let (q,r) = div2_with_rest n in (if r then ref_1 else ref_0)::(args (counter-1) q) in - GApp (dloc, ref_construct, List.rev (args 31 n)) + Loc.tag ~loc @@ GApp (ref_construct, List.rev (args 31 n)) let error_negative dloc = CErrors.user_err ~loc:dloc ~hdr:"interp_int31" (Pp.str "int31 are only non-negative numbers.") @@ -114,12 +114,12 @@ let bigint_of_int31 = let rec args_parsing args cur = match args with | [] -> cur - | (GRef (_,b,_))::l when eq_gr b int31_0 -> args_parsing l (mult_2 cur) - | (GRef (_,b,_))::l when eq_gr b int31_1 -> args_parsing l (add_1 (mult_2 cur)) + | (_, GRef (b,_))::l when eq_gr b int31_0 -> args_parsing l (mult_2 cur) + | (_, GRef (b,_))::l when eq_gr b int31_1 -> args_parsing l (add_1 (mult_2 cur)) | _ -> raise Non_closed in function - | GApp (_, GRef (_, c, _), args) when eq_gr c int31_construct -> args_parsing args zero + | _, GApp ((_, GRef (c, _)), args) when eq_gr c int31_construct -> args_parsing args zero | _ -> raise Non_closed let uninterp_int31 i = @@ -132,7 +132,7 @@ let uninterp_int31 i = let _ = Notation.declare_numeral_interpreter int31_scope (int31_path, int31_module) interp_int31 - ([GRef (Loc.ghost, int31_construct, None)], + ([Loc.tag @@ GRef (int31_construct, None)], uninterp_int31, true) @@ -162,34 +162,34 @@ let height bi = in hght 0 base (* n must be a non-negative integer (from bigint.ml) *) -let word_of_pos_bigint dloc hght n = - let ref_W0 = GRef (dloc, zn2z_W0, None) in - let ref_WW = GRef (dloc, zn2z_WW, None) in +let word_of_pos_bigint loc hght n = + let ref_W0 = Loc.tag ~loc @@ GRef (zn2z_W0, None) in + let ref_WW = Loc.tag ~loc @@ GRef (zn2z_WW, None) in let rec decomp hgt n = if hgt <= 0 then - int31_of_pos_bigint dloc n + int31_of_pos_bigint loc n else if equal n zero then - GApp (dloc, ref_W0, [GHole (dloc, Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None)]) + Loc.tag ~loc @@ GApp (ref_W0, [Loc.tag ~loc @@ GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None)]) else let (h,l) = split_at hgt n in - GApp (dloc, ref_WW, [GHole (dloc, Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None); + Loc.tag ~loc @@ GApp (ref_WW, [Loc.tag ~loc @@ GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None); decomp (hgt-1) h; decomp (hgt-1) l]) in decomp hght n -let bigN_of_pos_bigint dloc n = +let bigN_of_pos_bigint loc n = let h = height n in - let ref_constructor = GRef (dloc, bigN_constructor h, None) in - let word = word_of_pos_bigint dloc h n in + let ref_constructor = Loc.tag ~loc @@ GRef (bigN_constructor h, None) in + let word = word_of_pos_bigint loc h n in let args = if h < n_inlined then [word] - else [Nat_syntax_plugin.Nat_syntax.nat_of_int dloc (of_int (h-n_inlined));word] + else [Nat_syntax_plugin.Nat_syntax.nat_of_int loc (of_int (h-n_inlined));word] in - GApp (dloc, ref_constructor, args) + Loc.tag ~loc @@ GApp (ref_constructor, args) -let bigN_error_negative dloc = - CErrors.user_err ~loc:dloc ~hdr:"interp_bigN" (Pp.str "bigN are only non-negative numbers.") +let bigN_error_negative loc = + CErrors.user_err ~loc ~hdr:"interp_bigN" (Pp.str "bigN are only non-negative numbers.") let interp_bigN dloc n = if is_pos_or_zero n then @@ -203,14 +203,14 @@ let interp_bigN dloc n = let bigint_of_word = let rec get_height rc = match rc with - | GApp (_,GRef(_,c,_), [_;lft;rght]) when eq_gr c zn2z_WW -> + | _, GApp ((_, GRef(c,_)), [_;lft;rght]) when eq_gr c zn2z_WW -> 1+max (get_height lft) (get_height rght) | _ -> 0 in let rec transform hght rc = match rc with - | GApp (_,GRef(_,c,_),_) when eq_gr c zn2z_W0-> zero - | GApp (_,GRef(_,c,_), [_;lft;rght]) when eq_gr c zn2z_WW-> + | _, GApp ((_, GRef(c,_)),_) when eq_gr c zn2z_W0-> zero + | _, GApp ((_, GRef(c,_)), [_;lft;rght]) when eq_gr c zn2z_WW-> let new_hght = hght-1 in add (mult (rank new_hght) (transform new_hght lft)) @@ -223,8 +223,8 @@ let bigint_of_word = let bigint_of_bigN rc = match rc with - | GApp (_,_,[one_arg]) -> bigint_of_word one_arg - | GApp (_,_,[_;second_arg]) -> bigint_of_word second_arg + | _, GApp (_,[one_arg]) -> bigint_of_word one_arg + | _, GApp (_,[_;second_arg]) -> bigint_of_word second_arg | _ -> raise Non_closed let uninterp_bigN rc = @@ -240,7 +240,7 @@ let uninterp_bigN rc = let bigN_list_of_constructors = let rec build i = if i < n_inlined+1 then - GRef (Loc.ghost, bigN_constructor i,None)::(build (i+1)) + (Loc.tag @@ GRef (bigN_constructor i,None))::(build (i+1)) else [] in @@ -256,18 +256,18 @@ let _ = Notation.declare_numeral_interpreter bigN_scope (*** Parsing for bigZ in digital notation ***) -let interp_bigZ dloc n = - let ref_pos = GRef (dloc, bigZ_pos, None) in - let ref_neg = GRef (dloc, bigZ_neg, None) in +let interp_bigZ loc n = + let ref_pos = Loc.tag ~loc @@ GRef (bigZ_pos, None) in + let ref_neg = Loc.tag ~loc @@ GRef (bigZ_neg, None) in if is_pos_or_zero n then - GApp (dloc, ref_pos, [bigN_of_pos_bigint dloc n]) + Loc.tag ~loc @@ GApp (ref_pos, [bigN_of_pos_bigint loc n]) else - GApp (dloc, ref_neg, [bigN_of_pos_bigint dloc (neg n)]) + Loc.tag ~loc @@ GApp (ref_neg, [bigN_of_pos_bigint loc (neg n)]) (* pretty printing functions for bigZ *) let bigint_of_bigZ = function - | GApp (_, GRef(_,c,_), [one_arg]) when eq_gr c bigZ_pos -> bigint_of_bigN one_arg - | GApp (_, GRef(_,c,_), [one_arg]) when eq_gr c bigZ_neg -> + | _, GApp ((_, GRef(c,_)), [one_arg]) when eq_gr c bigZ_pos -> bigint_of_bigN one_arg + | _, GApp ((_, GRef(c,_)), [one_arg]) when eq_gr c bigZ_neg -> let opp_val = bigint_of_bigN one_arg in if equal opp_val zero then raise Non_closed @@ -286,19 +286,19 @@ let uninterp_bigZ rc = let _ = Notation.declare_numeral_interpreter bigZ_scope (bigZ_path, bigZ_module) interp_bigZ - ([GRef (Loc.ghost, bigZ_pos, None); - GRef (Loc.ghost, bigZ_neg, None)], + ([Loc.tag @@ GRef (bigZ_pos, None); + Loc.tag @@ GRef (bigZ_neg, None)], uninterp_bigZ, true) (*** Parsing for bigQ in digital notation ***) -let interp_bigQ dloc n = - let ref_z = GRef (dloc, bigQ_z, None) in - GApp (dloc, ref_z, [interp_bigZ dloc n]) +let interp_bigQ loc n = + let ref_z = Loc.tag ~loc @@ GRef (bigQ_z, None) in + Loc.tag ~loc @@ GApp (ref_z, [interp_bigZ loc n]) let uninterp_bigQ rc = try match rc with - | GApp (_, GRef(_,c,_), [one_arg]) when eq_gr c bigQ_z -> + | _, GApp ((_, GRef(c,_)), [one_arg]) when eq_gr c bigQ_z -> Some (bigint_of_bigZ one_arg) | _ -> None (* we don't pretty-print yet fractions *) with Non_closed -> None @@ -307,5 +307,5 @@ let uninterp_bigQ rc = let _ = Notation.declare_numeral_interpreter bigQ_scope (bigQ_path, bigQ_module) interp_bigQ - ([GRef (Loc.ghost, bigQ_z, None)], uninterp_bigQ, + ([Loc.tag @@ GRef (bigQ_z, None)], uninterp_bigQ, true) diff --git a/plugins/syntax/r_syntax.ml b/plugins/syntax/r_syntax.ml index 8f065f528..1af3f6c5b 100644 --- a/plugins/syntax/r_syntax.ml +++ b/plugins/syntax/r_syntax.ml @@ -42,13 +42,13 @@ let glob_xO = ConstructRef path_of_xO let glob_xH = ConstructRef path_of_xH let pos_of_bignat dloc x = - let ref_xI = GRef (dloc, glob_xI, None) in - let ref_xH = GRef (dloc, glob_xH, None) in - let ref_xO = GRef (dloc, glob_xO, None) in + let ref_xI = Loc.tag @@ GRef (glob_xI, None) in + let ref_xH = Loc.tag @@ GRef (glob_xH, None) in + let ref_xO = Loc.tag @@ GRef (glob_xO, None) in let rec pos_of x = match div2_with_rest x with - | (q,false) -> GApp (dloc, ref_xO,[pos_of q]) - | (q,true) when not (Bigint.equal q zero) -> GApp (dloc,ref_xI,[pos_of q]) + | (q,false) -> Loc.tag @@ GApp (ref_xO,[pos_of q]) + | (q,true) when not (Bigint.equal q zero) -> Loc.tag @@ GApp (ref_xI,[pos_of q]) | (q,true) -> ref_xH in pos_of x @@ -58,9 +58,9 @@ let pos_of_bignat dloc x = (**********************************************************************) let rec bignat_of_pos = function - | GApp (_, GRef (_,b,_),[a]) when Globnames.eq_gr b glob_xO -> mult_2(bignat_of_pos a) - | GApp (_, GRef (_,b,_),[a]) when Globnames.eq_gr b glob_xI -> add_1(mult_2(bignat_of_pos a)) - | GRef (_, a, _) when Globnames.eq_gr a glob_xH -> Bigint.one + | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_xO -> mult_2(bignat_of_pos a) + | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_xI -> add_1(mult_2(bignat_of_pos a)) + | _, GRef (a, _) when Globnames.eq_gr a glob_xH -> Bigint.one | _ -> raise Non_closed_number (**********************************************************************) @@ -81,18 +81,18 @@ let z_of_int dloc n = if not (Bigint.equal n zero) then let sgn, n = if is_pos_or_zero n then glob_POS, n else glob_NEG, Bigint.neg n in - GApp(dloc, GRef (dloc,sgn,None), [pos_of_bignat dloc n]) + Loc.tag @@ GApp(Loc.tag @@ GRef (sgn,None), [pos_of_bignat dloc n]) else - GRef (dloc, glob_ZERO, None) + Loc.tag @@ GRef (glob_ZERO, None) (**********************************************************************) (* Printing Z via scopes *) (**********************************************************************) let bigint_of_z = function - | GApp (_, GRef (_,b,_),[a]) when Globnames.eq_gr b glob_POS -> bignat_of_pos a - | GApp (_, GRef (_,b,_),[a]) when Globnames.eq_gr b glob_NEG -> Bigint.neg (bignat_of_pos a) - | GRef (_, a, _) when Globnames.eq_gr a glob_ZERO -> Bigint.zero + | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_POS -> bignat_of_pos a + | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_NEG -> Bigint.neg (bignat_of_pos a) + | _, GRef (a, _) when Globnames.eq_gr a glob_ZERO -> Bigint.zero | _ -> raise Non_closed_number (**********************************************************************) @@ -108,14 +108,14 @@ let make_path dir id = Globnames.encode_con dir (Id.of_string id) let glob_IZR = ConstRef (make_path (make_dir rdefinitions) "IZR") let r_of_int dloc z = - GApp (dloc, GRef(dloc,glob_IZR,None), [z_of_int dloc z]) + Loc.tag @@ GApp (Loc.tag @@ GRef(glob_IZR,None), [z_of_int dloc z]) (**********************************************************************) (* Printing R via scopes *) (**********************************************************************) let bigint_of_r = function - | GApp (_,GRef (_,o,_), [a]) when Globnames.eq_gr o glob_IZR -> + | _, GApp ((_, GRef (o,_)), [a]) when Globnames.eq_gr o glob_IZR -> bigint_of_z a | _ -> raise Non_closed_number @@ -128,6 +128,6 @@ let uninterp_r p = let _ = Notation.declare_numeral_interpreter "R_scope" (r_path,["Coq";"Reals";"Rdefinitions"]) r_of_int - ([GRef (Loc.ghost,glob_IZR,None)], + ([Loc.tag @@ GRef (glob_IZR,None)], uninterp_r, false) diff --git a/plugins/syntax/string_syntax.ml b/plugins/syntax/string_syntax.ml index de0fa77ef..539670722 100644 --- a/plugins/syntax/string_syntax.ml +++ b/plugins/syntax/string_syntax.ml @@ -33,23 +33,23 @@ let glob_EmptyString = lazy (make_reference "EmptyString") open Lazy -let interp_string dloc s = +let interp_string loc s = let le = String.length s in let rec aux n = - if n = le then GRef (dloc, force glob_EmptyString, None) else - GApp (dloc,GRef (dloc, force glob_String, None), - [interp_ascii dloc (int_of_char s.[n]); aux (n+1)]) + if n = le then Loc.tag ~loc @@ GRef (force glob_EmptyString, None) else + Loc.tag ~loc @@ GApp (Loc.tag ~loc @@ GRef (force glob_String, None), + [interp_ascii loc (int_of_char s.[n]); aux (n+1)]) in aux 0 let uninterp_string r = try let b = Buffer.create 16 in let rec aux = function - | GApp (_,GRef (_,k,_),[a;s]) when eq_gr k (force glob_String) -> + | _, GApp ((_, GRef (k,_)),[a;s]) when eq_gr k (force glob_String) -> (match uninterp_ascii a with | Some c -> Buffer.add_char b (Char.chr c); aux s | _ -> raise Non_closed_string) - | GRef (_,z,_) when eq_gr z (force glob_EmptyString) -> + | _, GRef (z,_) when eq_gr z (force glob_EmptyString) -> Some (Buffer.contents b) | _ -> raise Non_closed_string @@ -61,6 +61,6 @@ let _ = Notation.declare_string_interpreter "string_scope" (string_path,["Coq";"Strings";"String"]) interp_string - ([GRef (Loc.ghost,static_glob_String,None); - GRef (Loc.ghost,static_glob_EmptyString,None)], + ([Loc.tag @@ GRef (static_glob_String,None); + Loc.tag @@ GRef (static_glob_EmptyString,None)], uninterp_string, true) diff --git a/plugins/syntax/z_syntax.ml b/plugins/syntax/z_syntax.ml index b7b5fb8a5..a00525f91 100644 --- a/plugins/syntax/z_syntax.ml +++ b/plugins/syntax/z_syntax.ml @@ -44,14 +44,14 @@ let glob_xI = ConstructRef path_of_xI let glob_xO = ConstructRef path_of_xO let glob_xH = ConstructRef path_of_xH -let pos_of_bignat dloc x = - let ref_xI = GRef (dloc, glob_xI, None) in - let ref_xH = GRef (dloc, glob_xH, None) in - let ref_xO = GRef (dloc, glob_xO, None) in +let pos_of_bignat loc x = + let ref_xI = Loc.tag ~loc @@ GRef (glob_xI, None) in + let ref_xH = Loc.tag ~loc @@ GRef (glob_xH, None) in + let ref_xO = Loc.tag ~loc @@ GRef (glob_xO, None) in let rec pos_of x = match div2_with_rest x with - | (q,false) -> GApp (dloc, ref_xO,[pos_of q]) - | (q,true) when not (Bigint.equal q zero) -> GApp (dloc,ref_xI,[pos_of q]) + | (q,false) -> Loc.tag ~loc @@ GApp (ref_xO,[pos_of q]) + | (q,true) when not (Bigint.equal q zero) -> Loc.tag ~loc @@ GApp (ref_xI,[pos_of q]) | (q,true) -> ref_xH in pos_of x @@ -69,9 +69,9 @@ let interp_positive dloc n = (**********************************************************************) let rec bignat_of_pos = function - | GApp (_, GRef (_,b,_),[a]) when Globnames.eq_gr b glob_xO -> mult_2(bignat_of_pos a) - | GApp (_, GRef (_,b,_),[a]) when Globnames.eq_gr b glob_xI -> add_1(mult_2(bignat_of_pos a)) - | GRef (_, a, _) when Globnames.eq_gr a glob_xH -> Bigint.one + | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_xO -> mult_2(bignat_of_pos a) + | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_xI -> add_1(mult_2(bignat_of_pos a)) + | _, GRef (a, _) when Globnames.eq_gr a glob_xH -> Bigint.one | _ -> raise Non_closed_number let uninterp_positive p = @@ -87,9 +87,9 @@ let uninterp_positive p = let _ = Notation.declare_numeral_interpreter "positive_scope" (positive_path,binnums) interp_positive - ([GRef (Loc.ghost, glob_xI, None); - GRef (Loc.ghost, glob_xO, None); - GRef (Loc.ghost, glob_xH, None)], + ([Loc.tag @@ GRef (glob_xI, None); + Loc.tag @@ GRef (glob_xO, None); + Loc.tag @@ GRef (glob_xH, None)], uninterp_positive, true) @@ -106,11 +106,11 @@ let glob_Npos = ConstructRef path_of_Npos let n_path = make_path binnums "N" -let n_of_binnat dloc pos_or_neg n = +let n_of_binnat loc pos_or_neg n = Loc.tag ~loc @@ if not (Bigint.equal n zero) then - GApp(dloc, GRef (dloc,glob_Npos,None), [pos_of_bignat dloc n]) + GApp(Loc.tag @@ GRef (glob_Npos,None), [pos_of_bignat loc n]) else - GRef (dloc, glob_N0, None) + GRef(glob_N0, None) let error_negative dloc = user_err ~loc:dloc ~hdr:"interp_N" (str "No negative numbers in type \"N\".") @@ -124,8 +124,8 @@ let n_of_int dloc n = (**********************************************************************) let bignat_of_n = function - | GApp (_, GRef (_,b,_),[a]) when Globnames.eq_gr b glob_Npos -> bignat_of_pos a - | GRef (_, a,_) when Globnames.eq_gr a glob_N0 -> Bigint.zero + | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_Npos -> bignat_of_pos a + | _, GRef (a,_) when Globnames.eq_gr a glob_N0 -> Bigint.zero | _ -> raise Non_closed_number let uninterp_n p = @@ -138,8 +138,8 @@ let uninterp_n p = let _ = Notation.declare_numeral_interpreter "N_scope" (n_path,binnums) n_of_int - ([GRef (Loc.ghost, glob_N0, None); - GRef (Loc.ghost, glob_Npos, None)], + ([Loc.tag @@ GRef (glob_N0, None); + Loc.tag @@ GRef (glob_Npos, None)], uninterp_n, true) @@ -157,22 +157,22 @@ let glob_ZERO = ConstructRef path_of_ZERO let glob_POS = ConstructRef path_of_POS let glob_NEG = ConstructRef path_of_NEG -let z_of_int dloc n = +let z_of_int loc n = if not (Bigint.equal n zero) then let sgn, n = if is_pos_or_zero n then glob_POS, n else glob_NEG, Bigint.neg n in - GApp(dloc, GRef (dloc,sgn,None), [pos_of_bignat dloc n]) + Loc.tag ~loc @@ GApp(Loc.tag ~loc @@ GRef(sgn,None), [pos_of_bignat loc n]) else - GRef (dloc, glob_ZERO, None) + Loc.tag ~loc @@ GRef(glob_ZERO, None) (**********************************************************************) (* Printing Z via scopes *) (**********************************************************************) let bigint_of_z = function - | GApp (_, GRef (_,b,_),[a]) when Globnames.eq_gr b glob_POS -> bignat_of_pos a - | GApp (_, GRef (_,b,_),[a]) when Globnames.eq_gr b glob_NEG -> Bigint.neg (bignat_of_pos a) - | GRef (_, a, _) when Globnames.eq_gr a glob_ZERO -> Bigint.zero + | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_POS -> bignat_of_pos a + | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_NEG -> Bigint.neg (bignat_of_pos a) + | _, GRef (a, _) when Globnames.eq_gr a glob_ZERO -> Bigint.zero | _ -> raise Non_closed_number let uninterp_z p = @@ -186,8 +186,8 @@ let uninterp_z p = let _ = Notation.declare_numeral_interpreter "Z_scope" (z_path,binnums) z_of_int - ([GRef (Loc.ghost, glob_ZERO, None); - GRef (Loc.ghost, glob_POS, None); - GRef (Loc.ghost, glob_NEG, None)], + ([Loc.tag @@ GRef (glob_ZERO, None); + Loc.tag @@ GRef (glob_POS, None); + Loc.tag @@ GRef (glob_NEG, None)], uninterp_z, true) diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 531485935..347c49f44 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -348,8 +348,8 @@ let find_tomatch_tycon evdref env loc = function empty_tycon,None let coerce_row typing_fun evdref env pats (tomatch,(_,indopt)) = - let loc = Some (loc_of_glob_constr tomatch) in - let tycon,realnames = find_tomatch_tycon evdref env loc indopt in + let loc = loc_of_glob_constr tomatch in + let tycon,realnames = find_tomatch_tycon evdref env (Some loc) indopt in let j = typing_fun tycon env evdref tomatch in let evd, j = Coercion.inh_coerce_to_base (loc_of_glob_constr tomatch) env !evdref j in evdref := evd; @@ -357,7 +357,7 @@ let coerce_row typing_fun evdref env pats (tomatch,(_,indopt)) = let t = try try_find_ind env !evdref typ realnames with Not_found -> - unify_tomatch_with_patterns evdref env loc typ pats realnames in + unify_tomatch_with_patterns evdref env (Some loc) typ pats realnames in (j.uj_val,t) let coerce_to_indtype typing_fun evdref env matx tomatchl = @@ -1535,7 +1535,7 @@ substituer après par les initiaux *) * and linearizing the _ patterns. * Syntactic correctness has already been done in astterm *) let matx_of_eqns env eqns = - let build_eqn (loc,ids,lpat,rhs) = + let build_eqn (loc,(ids,lpat,rhs)) = let initial_lpat,initial_rhs = lpat,rhs in let initial_rhs = rhs in let rhs = @@ -2059,8 +2059,8 @@ let mk_JMeq evdref typ x typ' y = let mk_JMeq_refl evdref typ x = papp evdref coq_JMeq_refl [| typ; x |] -let hole = - GHole (Loc.ghost, Evar_kinds.QuestionMark (Evar_kinds.Define false), +let hole = Loc.tag @@ + GHole (Evar_kinds.QuestionMark (Evar_kinds.Define false), Misctypes.IntroAnonymous, None) let constr_of_pat env evdref arsign pat avoid = @@ -2160,13 +2160,13 @@ let vars_of_ctx sigma ctx = match decl with | LocalDef (na,t',t) when is_topvar sigma t' -> prev, - (GApp (Loc.ghost, - (GRef (Loc.ghost, delayed_force coq_eq_refl_ref, None)), - [hole; GVar (Loc.ghost, prev)])) :: vars + (Loc.tag @@ GApp ( + (Loc.tag @@ GRef (delayed_force coq_eq_refl_ref, None)), + [hole; Loc.tag @@ GVar prev])) :: vars | _ -> match RelDecl.get_name decl with Anonymous -> invalid_arg "vars_of_ctx" - | Name n -> n, GVar (Loc.ghost, n) :: vars) + | Name n -> n, (Loc.tag @@ GVar n) :: vars) ctx (Id.of_string "vars_of_ctx_error", []) in List.rev y @@ -2289,13 +2289,13 @@ let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity = let branch_name = Id.of_string ("program_branch_" ^ (string_of_int !i)) in let branch_decl = LocalDef (Name branch_name, lift !i bbody, lift !i btype) in let branch = - let bref = GVar (Loc.ghost, branch_name) in + let bref = Loc.tag @@ GVar branch_name in match vars_of_ctx !evdref rhs_rels with [] -> bref - | l -> GApp (Loc.ghost, bref, l) + | l -> Loc.tag @@ GApp (bref, l) in let branch = match ineqs with - Some _ -> GApp (Loc.ghost, branch, [ hole ]) + Some _ -> Loc.tag @@ GApp (branch, [ hole ]) | None -> branch in incr i; diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index f3018ac64..05d6a1ad4 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -331,20 +331,20 @@ let is_nondep_branch sigma c l = let extract_nondep_branches test c b l = let rec strip l r = - match r,l with - | r, [] -> r - | GLambda (_,_,_,_,t), false::l -> strip l t - | GLetIn (_,_,_,_,t), true::l -> strip l t + match snd r,l with + | r', [] -> r + | GLambda (_,_,_,t), false::l -> strip l t + | GLetIn (_,_,_,t), true::l -> strip l t (* FIXME: do we need adjustment? *) | _,_ -> assert false in if test c l then Some (strip l b) else None let it_destRLambda_or_LetIn_names l c = let rec aux l nal c = - match c, l with + match snd c, l with | _, [] -> (List.rev nal,c) - | GLambda (_,na,_,_,c), false::l -> aux l (na::nal) c - | GLetIn (_,na,_,_,c), true::l -> aux l (na::nal) c + | GLambda (na,_,_,c), false::l -> aux l (na::nal) c + | GLetIn (na,_,_,c), true::l -> aux l (na::nal) c | _, true::l -> (* let-expansion *) aux l (Anonymous :: nal) c | _, false::l -> (* eta-expansion *) @@ -355,11 +355,11 @@ let it_destRLambda_or_LetIn_names l c = x in let x = next (free_glob_vars c) in - let a = GVar (dl,x) in + let a = Loc.tag @@ GVar x in aux l (Name x :: nal) (match c with - | GApp (loc,p,l) -> GApp (loc,p,l@[a]) - | _ -> (GApp (dl,c,[a]))) + | loc, GApp (p,l) -> (loc, GApp (p,l@[a])) + | _ -> Loc.tag @@ GApp (c,[a])) in aux l [] c let detype_case computable detype detype_eqns testdep avoid data p c bl = @@ -375,12 +375,12 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl = | None -> Anonymous, None, None | Some p -> let nl,typ = it_destRLambda_or_LetIn_names k p in - let n,typ = match typ with - | GLambda (_,x,_,t,c) -> x, c + let n,typ = match snd typ with + | GLambda (x,_,t,c) -> x, c | _ -> Anonymous, typ in let aliastyp = if List.for_all (Name.equal Anonymous) nl then None - else Some (dl,(indsp,nl)) in + else Some (Loc.tag (indsp,nl)) in n, aliastyp, Some typ in let constructs = Array.init (Array.length bl) (fun i -> (indsp,i+1)) in @@ -397,25 +397,25 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl = else st with Not_found -> st - in + in Loc.tag @@ match tag, aliastyp with | LetStyle, None -> let bl' = Array.map detype bl in let (nal,d) = it_destRLambda_or_LetIn_names constagsl.(0) bl'.(0) in - GLetTuple (dl,nal,(alias,pred),tomatch,d) + GLetTuple (nal,(alias,pred),tomatch,d) | IfStyle, None -> let bl' = Array.map detype bl in let nondepbrs = Array.map3 (extract_nondep_branches testdep) bl bl' constagsl in if Array.for_all ((!=) None) nondepbrs then - GIf (dl,tomatch,(alias,pred), + GIf (tomatch,(alias,pred), Option.get nondepbrs.(0),Option.get nondepbrs.(1)) else let eqnl = detype_eqns constructs constagsl bl in - GCases (dl,tag,pred,[tomatch,(alias,aliastyp)],eqnl) + GCases (tag,pred,[tomatch,(alias,aliastyp)],eqnl) | _ -> let eqnl = detype_eqns constructs constagsl bl in - GCases (dl,tag,pred,[tomatch,(alias,aliastyp)],eqnl) + GCases (tag,pred,[tomatch,(alias,aliastyp)],eqnl) let detype_sort sigma = function | Prop Null -> GProp @@ -423,7 +423,7 @@ let detype_sort sigma = function | Type u -> GType (if !print_universes - then [dl, Pp.string_of_ppcmds (Univ.Universe.pr_with (Termops.pr_evd_level sigma) u)] + then [Loc.tag @@ Pp.string_of_ppcmds (Univ.Universe.pr_with (Termops.pr_evd_level sigma) u)] else []) type binder_kind = BProd | BLambda | BLetIn @@ -431,36 +431,36 @@ type binder_kind = BProd | BLambda | BLetIn (**********************************************************************) (* Main detyping function *) -let detype_anonymous = ref (fun loc n -> anomaly ~label:"detype" (Pp.str "index to an anonymous variable")) +let detype_anonymous = ref (fun ?loc n -> anomaly ~label:"detype" (Pp.str "index to an anonymous variable")) let set_detype_anonymous f = detype_anonymous := f let detype_level sigma l = - GType (Some (dl, Pp.string_of_ppcmds (Termops.pr_evd_level sigma l))) + GType (Some (Loc.tag @@ Pp.string_of_ppcmds (Termops.pr_evd_level sigma l))) let detype_instance sigma l = let l = EInstance.kind sigma l in if Univ.Instance.is_empty l then None else Some (List.map (detype_level sigma) (Array.to_list (Univ.Instance.to_array l))) -let rec detype flags avoid env sigma t = +let rec detype flags avoid env sigma t = Loc.tag @@ match EConstr.kind sigma (collapse_appl sigma t) with | Rel n -> (try match lookup_name_of_rel n (fst env) with - | Name id -> GVar (dl, id) - | Anonymous -> !detype_anonymous dl n + | Name id -> GVar id + | Anonymous -> snd @@ !detype_anonymous n with Not_found -> let s = "_UNBOUND_REL_"^(string_of_int n) - in GVar (dl, Id.of_string s)) + in GVar (Id.of_string s)) | Meta n -> (* Meta in constr are not user-parsable and are mapped to Evar *) (* using numbers to be unparsable *) - GEvar (dl, Id.of_string ("M" ^ string_of_int n), []) + GEvar (Id.of_string ("M" ^ string_of_int n), []) | Var id -> - (try let _ = Global.lookup_named id in GRef (dl, VarRef id, None) - with Not_found -> GVar (dl, id)) - | Sort s -> GSort (dl,detype_sort sigma (ESorts.kind sigma s)) + (try let _ = Global.lookup_named id in GRef (VarRef id, None) + with Not_found -> GVar id) + | Sort s -> GSort (detype_sort sigma (ESorts.kind sigma s)) | Cast (c1,REVERTcast,c2) when not !Flags.raw_print -> - detype flags avoid env sigma c1 + snd (detype flags avoid env sigma c1) | Cast (c1,k,c2) -> let d1 = detype flags avoid env sigma c1 in let d2 = detype flags avoid env sigma c2 in @@ -469,34 +469,34 @@ let rec detype flags avoid env sigma t = | NATIVEcast -> CastNative d2 | _ -> CastConv d2 in - GCast(dl,d1,cast) - | Prod (na,ty,c) -> detype_binder flags BProd avoid env sigma na None ty c - | Lambda (na,ty,c) -> detype_binder flags BLambda avoid env sigma na None ty c - | LetIn (na,b,ty,c) -> detype_binder flags BLetIn avoid env sigma na (Some b) ty c + GCast(d1,cast) + | Prod (na,ty,c) -> snd @@ detype_binder flags BProd avoid env sigma na None ty c + | Lambda (na,ty,c) -> snd @@ detype_binder flags BLambda avoid env sigma na None ty c + | LetIn (na,b,ty,c) -> snd @@ detype_binder flags BLetIn avoid env sigma na (Some b) ty c | App (f,args) -> let mkapp f' args' = - match f' with - | GApp (dl',f',args'') -> - GApp (dl,f',args''@args') - | _ -> GApp (dl,f',args') + match snd f' with + | GApp (f',args'') -> + GApp (f',args''@args') + | _ -> GApp (f',args') in mkapp (detype flags avoid env sigma f) (Array.map_to_list (detype flags avoid env sigma) args) - | Const (sp,u) -> GRef (dl, ConstRef sp, detype_instance sigma u) + | Const (sp,u) -> GRef (ConstRef sp, detype_instance sigma u) | Proj (p,c) -> let noparams () = let pb = Environ.lookup_projection p (snd env) in let pars = pb.Declarations.proj_npars in - let hole = GHole(Loc.ghost,Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None) in + let hole = Loc.tag @@ GHole(Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None) in let args = List.make pars hole in - GApp (dl, GRef (dl, ConstRef (Projection.constant p), None), + GApp (Loc.tag @@ GRef (ConstRef (Projection.constant p), None), (args @ [detype flags avoid env sigma c])) in if fst flags || !Flags.in_debugger || !Flags.in_toplevel then try noparams () with _ -> (* lax mode, used by debug printers only *) - GApp (dl, GRef (dl, ConstRef (Projection.constant p), None), + GApp (Loc.tag @@ GRef (ConstRef (Projection.constant p), None), [detype flags avoid env sigma c]) else if print_primproj_compatibility () && Projection.unfolded p then @@ -514,12 +514,12 @@ let rec detype flags avoid env sigma t = substl (c :: List.rev args) body' with Retyping.RetypeError _ | Not_found -> anomaly (str"Cannot detype an unfolded primitive projection.") - in detype flags avoid env sigma c' + in snd @@ detype flags avoid env sigma c' else if print_primproj_params () then try let c = Retyping.expand_projection (snd env) sigma p c [] in - detype flags avoid env sigma c + snd @@ detype flags avoid env sigma c with Retyping.RetypeError _ -> noparams () else noparams () @@ -546,14 +546,15 @@ let rec detype flags avoid env sigma t = Id.of_string ("X" ^ string_of_int (Evar.repr evk)), (Array.map_to_list (fun c -> (Id.of_string "__",c)) cl) in - GEvar (dl,id, + GEvar (id, List.map (on_snd (detype flags avoid env sigma)) l) | Ind (ind_sp,u) -> - GRef (dl, IndRef ind_sp, detype_instance sigma u) + GRef (IndRef ind_sp, detype_instance sigma u) | Construct (cstr_sp,u) -> - GRef (dl, ConstructRef cstr_sp, detype_instance sigma u) + GRef (ConstructRef cstr_sp, detype_instance sigma u) | Case (ci,p,c,bl) -> let comp = computable sigma p (List.length (ci.ci_pp_info.ind_tags)) in + snd @@ detype_case comp (detype flags avoid env sigma) (detype_eqns flags avoid env sigma ci comp) (is_nondep_branch sigma) avoid @@ -574,7 +575,7 @@ and detype_fix flags avoid env sigma (vn,_ as nvn) (names,tys,bodies) = let v = Array.map3 (fun c t i -> share_names flags (i+1) [] def_avoid def_env sigma c (lift n t)) bodies tys vn in - GRec(dl,GFix (Array.map (fun i -> Some i, GStructRec) (fst nvn), snd nvn),Array.of_list (List.rev lfi), + GRec(GFix (Array.map (fun i -> Some i, GStructRec) (fst nvn), snd nvn),Array.of_list (List.rev lfi), Array.map (fun (bl,_,_) -> bl) v, Array.map (fun (_,_,ty) -> ty) v, Array.map (fun (_,bd,_) -> bd) v) @@ -590,7 +591,7 @@ and detype_cofix flags avoid env sigma n (names,tys,bodies) = let v = Array.map2 (fun c t -> share_names flags 0 [] def_avoid def_env sigma c (lift ntys t)) bodies tys in - GRec(dl,GCoFix n,Array.of_list (List.rev lfi), + GRec(GCoFix n,Array.of_list (List.rev lfi), Array.map (fun (bl,_,_) -> bl) v, Array.map (fun (_,_,ty) -> ty) v, Array.map (fun (_,bd,_) -> bd) v) @@ -635,7 +636,7 @@ and detype_eqns flags avoid env sigma ci computable constructs consnargsl bl = try if !Flags.raw_print || not (reverse_matching ()) then raise Exit; let mat = build_tree Anonymous (snd flags) (avoid,env) sigma ci bl in - List.map (fun (pat,((avoid,env),c)) -> (dl,[],[pat],detype flags avoid env sigma c)) + List.map (fun (pat,((avoid,env),c)) -> Loc.tag ([],[pat],detype flags avoid env sigma c)) mat with e when CErrors.noncritical e -> Array.to_list @@ -644,7 +645,7 @@ and detype_eqns flags avoid env sigma ci computable constructs consnargsl bl = and detype_eqn (lax,isgoal as flags) avoid env sigma constr construct_nargs branch = let make_pat x avoid env b body ty ids = if force_wildcard () && noccurn sigma 1 b then - Loc.tag @@ PatVar (Anonymous),avoid,(add_name Anonymous body ty env),ids + Loc.tag @@ PatVar Anonymous,avoid,(add_name Anonymous body ty env),ids else let flag = if isgoal then RenamingForGoal else RenamingForCasesPattern (fst env,b) in let na,avoid' = compute_displayed_name_in sigma flag avoid x b in @@ -652,9 +653,9 @@ and detype_eqn (lax,isgoal as flags) avoid env sigma constr construct_nargs bran in let rec buildrec ids patlist avoid env l b = match EConstr.kind sigma b, l with - | _, [] -> - (dl, Id.Set.elements ids, - [Loc.tag ~loc:dl @@ PatCstr(constr, List.rev patlist,Anonymous)], + | _, [] -> Loc.tag @@ + (Id.Set.elements ids, + [Loc.tag @@ PatCstr(constr, List.rev patlist,Anonymous)], detype flags avoid env sigma b) | Lambda (x,t,b), false::l -> let pat,new_avoid,new_env,new_ids = make_pat x avoid env b None t ids in @@ -668,7 +669,7 @@ and detype_eqn (lax,isgoal as flags) avoid env sigma constr construct_nargs bran buildrec ids patlist avoid env l c | _, true::l -> - let pat = Loc.tag ~loc:dl @@ PatVar Anonymous in + let pat = Loc.tag @@ PatVar Anonymous in buildrec ids (pat::patlist) avoid env l b | _, false::l -> @@ -683,21 +684,21 @@ and detype_eqn (lax,isgoal as flags) avoid env sigma constr construct_nargs bran in buildrec Id.Set.empty [] avoid env construct_nargs branch -and detype_binder (lax,isgoal as flags) bk avoid env sigma na body ty c = +and detype_binder (lax,isgoal as flags) bk avoid env sigma na body ty c = Loc.tag @@ let flag = if isgoal then RenamingForGoal else RenamingElsewhereFor (fst env,c) in let na',avoid' = match bk with | BLetIn -> compute_displayed_let_name_in sigma flag avoid na c | _ -> compute_displayed_name_in sigma flag avoid na c in let r = detype flags avoid' (add_name na' body ty env) sigma c in match bk with - | BProd -> GProd (dl, na',Explicit,detype (lax,false) avoid env sigma ty, r) - | BLambda -> GLambda (dl, na',Explicit,detype (lax,false) avoid env sigma ty, r) + | BProd -> GProd (na',Explicit,detype (lax,false) avoid env sigma ty, r) + | BLambda -> GLambda (na',Explicit,detype (lax,false) avoid env sigma ty, r) | BLetIn -> let c = detype (lax,false) avoid env sigma (Option.get body) in (* Heuristic: we display the type if in Prop *) let s = try Retyping.get_sort_family_of (snd env) sigma ty with _ when !Flags.in_debugger || !Flags.in_toplevel -> InType (* Can fail because of sigma missing in debugger *) in let t = if s != InProp then None else Some (detype (lax,false) avoid env sigma ty) in - GLetIn (dl, na', c, t, r) + GLetIn (na', c, t, r) let detype_rel_context ?(lax=false) where avoid env sigma sign = let where = Option.map (fun c -> EConstr.it_mkLambda_or_LetIn c sign) where in @@ -741,11 +742,11 @@ let detype_closed_glob ?lax isgoal avoid env sigma t = | Name id -> Name (convert_id cl id) | Anonymous -> Anonymous in - let rec detype_closed_glob cl = function - | GVar (loc,id) -> + let rec detype_closed_glob cl cg = Loc.map (function + | GVar id -> (* if [id] is bound to a name. *) begin try - GVar(loc,Id.Map.find id cl.idents) + GVar(Id.Map.find id cl.idents) (* if [id] is bound to a typed term *) with Not_found -> try (* assumes [detype] does not raise [Not_found] exceptions *) @@ -755,38 +756,39 @@ let detype_closed_glob ?lax isgoal avoid env sigma t = [Printer.pr_constr_under_binders_env] does. *) let assums = List.map (fun id -> LocalAssum (Name id,(* dummy *) mkProp)) b in let env = push_rel_context assums env in - detype ?lax isgoal avoid env sigma c + snd @@ detype ?lax isgoal avoid env sigma c (* if [id] is bound to a [closed_glob_constr]. *) with Not_found -> try let {closure;term} = Id.Map.find id cl.untyped in - detype_closed_glob closure term + snd @@ detype_closed_glob closure term (* Otherwise [id] stands for itself *) with Not_found -> - GVar(loc,id) + GVar id end - | GLambda (loc,id,k,t,c) -> + | GLambda (id,k,t,c) -> let id = convert_name cl id in - GLambda(loc,id,k,detype_closed_glob cl t, detype_closed_glob cl c) - | GProd (loc,id,k,t,c) -> + GLambda(id,k,detype_closed_glob cl t, detype_closed_glob cl c) + | GProd (id,k,t,c) -> let id = convert_name cl id in - GProd(loc,id,k,detype_closed_glob cl t, detype_closed_glob cl c) - | GLetIn (loc,id,b,t,e) -> + GProd(id,k,detype_closed_glob cl t, detype_closed_glob cl c) + | GLetIn (id,b,t,e) -> let id = convert_name cl id in - GLetIn(loc,id,detype_closed_glob cl b, Option.map (detype_closed_glob cl) t, detype_closed_glob cl e) - | GLetTuple (loc,ids,(n,r),b,e) -> + GLetIn(id,detype_closed_glob cl b, Option.map (detype_closed_glob cl) t, detype_closed_glob cl e) + | GLetTuple (ids,(n,r),b,e) -> let ids = List.map (convert_name cl) ids in let n = convert_name cl n in - GLetTuple (loc,ids,(n,r),detype_closed_glob cl b, detype_closed_glob cl e) - | GCases (loc,sty,po,tml,eqns) -> + GLetTuple (ids,(n,r),detype_closed_glob cl b, detype_closed_glob cl e) + | GCases (sty,po,tml,eqns) -> let (tml,eqns) = Glob_ops.map_pattern_binders (fun na -> convert_name cl na) tml eqns in let (tml,eqns) = Glob_ops.map_pattern (fun c -> detype_closed_glob cl c) tml eqns in - GCases(loc,sty,po,tml,eqns) + GCases(sty,po,tml,eqns) | c -> - Glob_ops.map_glob_constr (detype_closed_glob cl) c + snd @@ Glob_ops.map_glob_constr (detype_closed_glob cl) cg + ) cg in detype_closed_glob t.closure t.term @@ -804,41 +806,41 @@ let rec subst_cases_pattern subst (loc, pat) = Loc.tag ~loc @@ let (f_subst_genarg, subst_genarg_hook) = Hook.make () -let rec subst_glob_constr subst raw = +let rec subst_glob_constr subst (loc, raw) = Loc.tag ~loc @@ match raw with - | GRef (loc,ref,u) -> + | GRef (ref,u) -> let ref',t = subst_global subst ref in if ref' == ref then raw else - detype false [] (Global.env()) Evd.empty (EConstr.of_constr t) + snd @@ detype false [] (Global.env()) Evd.empty (EConstr.of_constr t) | GVar _ -> raw | GEvar _ -> raw | GPatVar _ -> raw - | GApp (loc,r,rl) -> + | GApp (r,rl) -> let r' = subst_glob_constr subst r and rl' = List.smartmap (subst_glob_constr subst) rl in if r' == r && rl' == rl then raw else - GApp(loc,r',rl') + GApp(r',rl') - | GLambda (loc,n,bk,r1,r2) -> + | GLambda (n,bk,r1,r2) -> let r1' = subst_glob_constr subst r1 and r2' = subst_glob_constr subst r2 in if r1' == r1 && r2' == r2 then raw else - GLambda (loc,n,bk,r1',r2') + GLambda (n,bk,r1',r2') - | GProd (loc,n,bk,r1,r2) -> + | GProd (n,bk,r1,r2) -> let r1' = subst_glob_constr subst r1 and r2' = subst_glob_constr subst r2 in if r1' == r1 && r2' == r2 then raw else - GProd (loc,n,bk,r1',r2') + GProd (n,bk,r1',r2') - | GLetIn (loc,n,r1,t,r2) -> + | GLetIn (n,r1,t,r2) -> let r1' = subst_glob_constr subst r1 in - let t' = Option.smartmap (subst_glob_constr subst) t in let r2' = subst_glob_constr subst r2 in + let t' = Option.smartmap (subst_glob_constr subst) t in if r1' == r1 && t == t' && r2' == r2 then raw else - GLetIn (loc,n,r1',t',r2') + GLetIn (n,r1',t',r2') - | GCases (loc,sty,rtno,rl,branches) -> + | GCases (sty,rtno,rl,branches) -> let rtno' = Option.smartmap (subst_glob_constr subst) rtno and rl' = List.smartmap (fun (a,x as y) -> let a' = subst_glob_constr subst a in @@ -849,33 +851,33 @@ let rec subst_glob_constr subst raw = if sp == sp' then t else (loc,((sp',i),y))) topt in if a == a' && topt == topt' then y else (a',(n,topt'))) rl and branches' = List.smartmap - (fun (loc,idl,cpl,r as branch) -> + (fun (loc,(idl,cpl,r) as branch) -> let cpl' = List.smartmap (subst_cases_pattern subst) cpl and r' = subst_glob_constr subst r in if cpl' == cpl && r' == r then branch else - (loc,idl,cpl',r')) + (loc,(idl,cpl',r'))) branches in if rtno' == rtno && rl' == rl && branches' == branches then raw else - GCases (loc,sty,rtno',rl',branches') + GCases (sty,rtno',rl',branches') - | GLetTuple (loc,nal,(na,po),b,c) -> + | GLetTuple (nal,(na,po),b,c) -> let po' = Option.smartmap (subst_glob_constr subst) po and b' = subst_glob_constr subst b and c' = subst_glob_constr subst c in if po' == po && b' == b && c' == c then raw else - GLetTuple (loc,nal,(na,po'),b',c') + GLetTuple (nal,(na,po'),b',c') - | GIf (loc,c,(na,po),b1,b2) -> + | GIf (c,(na,po),b1,b2) -> let po' = Option.smartmap (subst_glob_constr subst) po and b1' = subst_glob_constr subst b1 and b2' = subst_glob_constr subst b2 and c' = subst_glob_constr subst c in if c' == c && po' == po && b1' == b1 && b2' == b2 then raw else - GIf (loc,c',(na,po'),b1',b2') + GIf (c',(na,po'),b1',b2') - | GRec (loc,fix,ida,bl,ra1,ra2) -> + | GRec (fix,ida,bl,ra1,ra2) -> let ra1' = Array.smartmap (subst_glob_constr subst) ra1 and ra2' = Array.smartmap (subst_glob_constr subst) ra2 in let bl' = Array.smartmap @@ -885,11 +887,11 @@ let rec subst_glob_constr subst raw = if ty'==ty && obd'==obd then dcl else (na,k,obd',ty'))) bl in if ra1' == ra1 && ra2' == ra2 && bl'==bl then raw else - GRec (loc,fix,ida,bl',ra1',ra2') + GRec (fix,ida,bl',ra1',ra2') | GSort _ -> raw - | GHole (loc, knd, naming, solve) -> + | GHole (knd, naming, solve) -> let nknd = match knd with | Evar_kinds.ImplicitArg (ref, i, b) -> let nref, _ = subst_global subst ref in @@ -898,12 +900,12 @@ let rec subst_glob_constr subst raw = in let nsolve = Option.smartmap (Hook.get f_subst_genarg subst) solve in if nsolve == solve && nknd == knd then raw - else GHole (loc, nknd, naming, nsolve) + else GHole (nknd, naming, nsolve) - | GCast (loc,r1,k) -> + | GCast (r1,k) -> let r1' = subst_glob_constr subst r1 in let k' = Miscops.smartmap_cast_type (subst_glob_constr subst) k in - if r1' == r1 && k' == k then raw else GCast (loc,r1',k') + if r1' == r1 && k' == k then raw else GCast (r1',k') (* Utilities to transform kernel cases to simple pattern-matching problem *) @@ -914,7 +916,7 @@ let simple_cases_matrix_of_branches ind brs = let p = Loc.tag @@ PatCstr ((ind,i+1),List.map mkPatVar nal,Anonymous) in let map name = try Some (Nameops.out_name name) with Failure _ -> None in let ids = List.map_filter map nal in - (Loc.ghost,ids,[p],c)) + Loc.tag @@ (ids,[p],c)) brs let return_type_of_predicate ind nrealargs_tags pred = diff --git a/pretyping/detyping.mli b/pretyping/detyping.mli index 4c6f9129f..84da3652f 100644 --- a/pretyping/detyping.mli +++ b/pretyping/detyping.mli @@ -38,7 +38,7 @@ val detype : ?lax:bool -> bool -> Id.t list -> env -> evar_map -> constr -> glob val detype_case : bool -> (constr -> glob_constr) -> (constructor array -> bool list array -> constr array -> - (Loc.t * Id.t list * cases_pattern list * glob_constr) list) -> + (Id.t list * cases_pattern list * glob_constr) Loc.located list) -> (constr -> bool list -> bool) -> Id.t list -> inductive * case_style * bool list array * bool list -> constr option -> constr -> constr array -> glob_constr @@ -54,7 +54,9 @@ val detype_closed_glob : ?lax:bool -> bool -> Id.t list -> env -> evar_map -> cl val lookup_name_as_displayed : env -> evar_map -> constr -> Id.t -> int option val lookup_index_as_renamed : env -> evar_map -> constr -> int -> int option -val set_detype_anonymous : (Loc.t -> int -> glob_constr) -> unit +(* XXX: This is a hack and should go away *) +val set_detype_anonymous : (?loc:Loc.t -> int -> glob_constr) -> unit + val force_wildcard : unit -> bool val synthetize_type : unit -> bool diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml index 4cccaaf8f..25ece5b8e 100644 --- a/pretyping/glob_ops.ml +++ b/pretyping/glob_ops.ml @@ -22,10 +22,10 @@ let cases_predicate_names tml = | (tm,(na,None)) -> [na] | (tm,(na,Some (_,(_,nal)))) -> na::nal) tml) -let mkGApp loc p t = - match p with - | GApp (loc,f,l) -> GApp (loc,f,l@[t]) - | _ -> GApp (loc,p,[t]) +let mkGApp loc p t = Loc.tag ~loc @@ + match snd p with + | GApp (f,l) -> GApp (f,l@[t]) + | _ -> GApp (p,[t]) let map_glob_decl_left_to_right f (na,k,obd,ty) = let comp1 = Option.map f obd in @@ -59,46 +59,46 @@ let cast_type_eq eq t1 t2 = match t1, t2 with | CastNative t1, CastNative t2 -> eq t1 t2 | _ -> false -let rec glob_constr_eq c1 c2 = match c1, c2 with -| GRef (_, gr1, _), GRef (_, gr2, _) -> eq_gr gr1 gr2 -| GVar (_, id1), GVar (_, id2) -> Id.equal id1 id2 -| GEvar (_, id1, arg1), GEvar (_, id2, arg2) -> +let rec glob_constr_eq (_loc1, c1) (_loc2, c2) = match c1, c2 with +| GRef (gr1, _), GRef (gr2, _) -> eq_gr gr1 gr2 +| GVar id1, GVar id2 -> Id.equal id1 id2 +| GEvar (id1, arg1), GEvar (id2, arg2) -> Id.equal id1 id2 && List.equal instance_eq arg1 arg2 -| GPatVar (_, (b1, pat1)), GPatVar (_, (b2, pat2)) -> +| GPatVar (b1, pat1), GPatVar (b2, pat2) -> (b1 : bool) == b2 && Id.equal pat1 pat2 -| GApp (_, f1, arg1), GApp (_, f2, arg2) -> +| GApp (f1, arg1), GApp (f2, arg2) -> glob_constr_eq f1 f2 && List.equal glob_constr_eq arg1 arg2 -| GLambda (_, na1, bk1, t1, c1), GLambda (_, na2, bk2, t2, c2) -> +| GLambda (na1, bk1, t1, c1), GLambda (na2, bk2, t2, c2) -> Name.equal na1 na2 && binding_kind_eq bk1 bk2 && glob_constr_eq t1 t2 && glob_constr_eq c1 c2 -| GProd (_, na1, bk1, t1, c1), GProd (_, na2, bk2, t2, c2) -> +| GProd (na1, bk1, t1, c1), GProd (na2, bk2, t2, c2) -> Name.equal na1 na2 && binding_kind_eq bk1 bk2 && glob_constr_eq t1 t2 && glob_constr_eq c1 c2 -| GLetIn (_, na1, b1, t1, c1), GLetIn (_, na2, b2, t2, c2) -> +| GLetIn (na1, b1, t1, c1), GLetIn (na2, b2, t2, c2) -> Name.equal na1 na2 && glob_constr_eq b1 b2 && Option.equal glob_constr_eq t1 t2 && glob_constr_eq c1 c2 -| GCases (_, st1, c1, tp1, cl1), GCases (_, st2, c2, tp2, cl2) -> +| GCases (st1, c1, tp1, cl1), GCases (st2, c2, tp2, cl2) -> case_style_eq st1 st2 && Option.equal glob_constr_eq c1 c2 && List.equal tomatch_tuple_eq tp1 tp2 && List.equal cases_clause_eq cl1 cl2 -| GLetTuple (_, na1, (n1, p1), c1, t1), GLetTuple (_, na2, (n2, p2), c2, t2) -> +| GLetTuple (na1, (n1, p1), c1, t1), GLetTuple (na2, (n2, p2), c2, t2) -> List.equal Name.equal na1 na2 && Name.equal n1 n2 && Option.equal glob_constr_eq p1 p2 && glob_constr_eq c1 c2 && glob_constr_eq t1 t2 -| GIf (_, m1, (pat1, p1), c1, t1), GIf (_, m2, (pat2, p2), c2, t2) -> +| GIf (m1, (pat1, p1), c1, t1), GIf (m2, (pat2, p2), c2, t2) -> glob_constr_eq m1 m2 && Name.equal pat1 pat2 && Option.equal glob_constr_eq p1 p2 && glob_constr_eq c1 c2 && glob_constr_eq t1 t2 -| GRec (_, kn1, id1, decl1, c1, t1), GRec (_, kn2, id2, decl2, c2, t2) -> +| GRec (kn1, id1, decl1, c1, t1), GRec (kn2, id2, decl2, c2, t2) -> fix_kind_eq kn1 kn2 && Array.equal Id.equal id1 id2 && Array.equal (fun l1 l2 -> List.equal glob_decl_eq l1 l2) decl1 decl2 && Array.equal glob_constr_eq c1 c2 && Array.equal glob_constr_eq t1 t2 -| GSort (_, s1), GSort (_, s2) -> Miscops.glob_sort_eq s1 s2 -| GHole (_, kn1, nam1, gn1), GHole (_, kn2, nam2, gn2) -> +| GSort s1, GSort s2 -> Miscops.glob_sort_eq s1 s2 +| GHole (kn1, nam1, gn1), GHole (kn2, nam2, gn2) -> Option.equal (==) gn1 gn2 (** Only thing sensible *) && Miscops.intro_pattern_naming_eq nam1 nam2 -| GCast (_, c1, t1), GCast (_, c2, t2) -> +| GCast (c1, t1), GCast (c2, t2) -> glob_constr_eq c1 c2 && cast_type_eq glob_constr_eq t1 t2 | _ -> false @@ -109,7 +109,7 @@ and tomatch_tuple_eq (c1, p1) (c2, p2) = let eq_pred (n1, o1) (n2, o2) = Name.equal n1 n2 && Option.equal eqp o1 o2 in glob_constr_eq c1 c2 && eq_pred p1 p2 -and cases_clause_eq (_, id1, p1, c1) (_, id2, p2, c2) = +and cases_clause_eq (_, (id1, p1, c1)) (_, (id2, p2, c2)) = List.equal Id.equal id1 id2 && List.equal cases_pattern_eq p1 p2 && glob_constr_eq c1 c2 @@ -137,80 +137,82 @@ and fix_recursion_order_eq o1 o2 = match o1, o2 with and instance_eq (x1,c1) (x2,c2) = Id.equal x1 x2 && glob_constr_eq c1 c2 -let map_glob_constr_left_to_right f = function - | GApp (loc,g,args) -> +let map_glob_constr_left_to_right f = Loc.map (function + | GApp (g,args) -> let comp1 = f g in let comp2 = Util.List.map_left f args in - GApp (loc,comp1,comp2) - | GLambda (loc,na,bk,ty,c) -> + GApp (comp1,comp2) + | GLambda (na,bk,ty,c) -> let comp1 = f ty in let comp2 = f c in - GLambda (loc,na,bk,comp1,comp2) - | GProd (loc,na,bk,ty,c) -> + GLambda (na,bk,comp1,comp2) + | GProd (na,bk,ty,c) -> let comp1 = f ty in let comp2 = f c in - GProd (loc,na,bk,comp1,comp2) - | GLetIn (loc,na,b,t,c) -> + GProd (na,bk,comp1,comp2) + | GLetIn (na,b,t,c) -> let comp1 = f b in let compt = Option.map f t in let comp2 = f c in - GLetIn (loc,na,comp1,compt,comp2) - | GCases (loc,sty,rtntypopt,tml,pl) -> + GLetIn (na,comp1,compt,comp2) + | GCases (sty,rtntypopt,tml,pl) -> let comp1 = Option.map f rtntypopt in let comp2 = Util.List.map_left (fun (tm,x) -> (f tm,x)) tml in - let comp3 = Util.List.map_left (fun (loc,idl,p,c) -> (loc,idl,p,f c)) pl in - GCases (loc,sty,comp1,comp2,comp3) - | GLetTuple (loc,nal,(na,po),b,c) -> + let comp3 = Util.List.map_left (fun (loc,(idl,p,c)) -> (loc,(idl,p,f c))) pl in + GCases (sty,comp1,comp2,comp3) + | GLetTuple (nal,(na,po),b,c) -> let comp1 = Option.map f po in let comp2 = f b in let comp3 = f c in - GLetTuple (loc,nal,(na,comp1),comp2,comp3) - | GIf (loc,c,(na,po),b1,b2) -> + GLetTuple (nal,(na,comp1),comp2,comp3) + | GIf (c,(na,po),b1,b2) -> let comp1 = Option.map f po in let comp2 = f b1 in let comp3 = f b2 in - GIf (loc,f c,(na,comp1),comp2,comp3) - | GRec (loc,fk,idl,bl,tyl,bv) -> + GIf (f c,(na,comp1),comp2,comp3) + | GRec (fk,idl,bl,tyl,bv) -> let comp1 = Array.map (Util.List.map_left (map_glob_decl_left_to_right f)) bl in let comp2 = Array.map f tyl in let comp3 = Array.map f bv in - GRec (loc,fk,idl,comp1,comp2,comp3) - | GCast (loc,c,k) -> + GRec (fk,idl,comp1,comp2,comp3) + | GCast (c,k) -> let comp1 = f c in let comp2 = Miscops.map_cast_type f k in - GCast (loc,comp1,comp2) + GCast (comp1,comp2) | (GVar _ | GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) as x -> x + ) let map_glob_constr = map_glob_constr_left_to_right let fold_return_type f acc (na,tyopt) = Option.fold_left f acc tyopt -let fold_glob_constr f acc = function +let fold_glob_constr f acc = Loc.with_unloc (function | GVar _ -> acc - | GApp (_,c,args) -> List.fold_left f (f acc c) args - | GLambda (_,_,_,b,c) | GProd (_,_,_,b,c) -> + | GApp (c,args) -> List.fold_left f (f acc c) args + | GLambda (_,_,b,c) | GProd (_,_,b,c) -> f (f acc b) c - | GLetIn (_,_,b,t,c) -> + | GLetIn (_,b,t,c) -> f (Option.fold_left f (f acc b) t) c - | GCases (_,_,rtntypopt,tml,pl) -> - let fold_pattern acc (_,idl,p,c) = f acc c in + | GCases (_,rtntypopt,tml,pl) -> + let fold_pattern acc (_,(idl,p,c)) = f acc c in List.fold_left fold_pattern (List.fold_left f (Option.fold_left f acc rtntypopt) (List.map fst tml)) pl - | GLetTuple (_,_,rtntyp,b,c) -> + | GLetTuple (_,rtntyp,b,c) -> f (f (fold_return_type f acc rtntyp) b) c - | GIf (_,c,rtntyp,b1,b2) -> + | GIf (c,rtntyp,b1,b2) -> f (f (f (fold_return_type f acc rtntyp) c) b1) b2 - | GRec (_,_,_,bl,tyl,bv) -> + | GRec (_,_,bl,tyl,bv) -> let acc = Array.fold_left (List.fold_left (fun acc (na,k,bbd,bty) -> f (Option.fold_left f acc bbd) bty)) acc bl in Array.fold_left f (Array.fold_left f acc tyl) bv - | GCast (_,c,k) -> + | GCast (c,k) -> let acc = match k with | CastConv t | CastVM t | CastNative t -> f acc t | CastCoerce -> acc in f acc c | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> acc + ) let iter_glob_constr f = fold_glob_constr (fun () -> f) () @@ -219,25 +221,25 @@ let same_id na id = match na with | Name id' -> Id.equal id id' let occur_glob_constr id = - let rec occur = function - | GVar (loc,id') -> Id.equal id id' - | GApp (loc,f,args) -> (occur f) || (List.exists occur args) - | GLambda (loc,na,bk,ty,c) -> + let rec occur gt = Loc.with_unloc (function + | GVar (id') -> Id.equal id id' + | GApp (f,args) -> (occur f) || (List.exists occur args) + | GLambda (na,bk,ty,c) -> (occur ty) || (not (same_id na id) && (occur c)) - | GProd (loc,na,bk,ty,c) -> + | GProd (na,bk,ty,c) -> (occur ty) || (not (same_id na id) && (occur c)) - | GLetIn (loc,na,b,t,c) -> + | GLetIn (na,b,t,c) -> (Option.fold_left (fun b t -> occur t || b) (occur b) t) || (not (same_id na id) && (occur c)) - | GCases (loc,sty,rtntypopt,tml,pl) -> + | GCases (sty,rtntypopt,tml,pl) -> (occur_option rtntypopt) || (List.exists (fun (tm,_) -> occur tm) tml) || (List.exists occur_pattern pl) - | GLetTuple (loc,nal,rtntyp,b,c) -> + | GLetTuple (nal,rtntyp,b,c) -> occur_return_type rtntyp id || (occur b) || (not (List.mem_f Name.equal (Name id) nal) && (occur c)) - | GIf (loc,c,rtntyp,b1,b2) -> + | GIf (c,rtntyp,b1,b2) -> occur_return_type rtntyp id || (occur c) || (occur b1) || (occur b2) - | GRec (loc,fk,idl,bl,tyl,bv) -> + | GRec (fk,idl,bl,tyl,bv) -> not (Array.for_all4 (fun fid bl ty bd -> let rec occur_fix = function [] -> not (occur ty) && (Id.equal fid id || not(occur bd)) @@ -249,11 +251,11 @@ let occur_glob_constr id = (match na with Name id' -> Id.equal id id' | _ -> not (occur_fix bl)) in occur_fix bl) idl bl tyl bv) - | GCast (loc,c,k) -> (occur c) || (match k with CastConv t + | GCast (c,k) -> (occur c) || (match k with CastConv t | CastVM t | CastNative t -> occur t | CastCoerce -> false) | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> false - - and occur_pattern (loc,idl,p,c) = not (Id.List.mem id idl) && (occur c) + ) gt + and occur_pattern (loc,(idl,p,c)) = not (Id.List.mem id idl) && (occur c) and occur_option = function None -> false | Some p -> occur p @@ -268,33 +270,33 @@ let add_name_to_ids set na = | Name id -> Id.Set.add id set let free_glob_vars = - let rec vars bounded vs = function - | GVar (loc,id') -> if Id.Set.mem id' bounded then vs else Id.Set.add id' vs - | GApp (loc,f,args) -> List.fold_left (vars bounded) vs (f::args) - | GLambda (loc,na,_,ty,c) | GProd (loc,na,_,ty,c) -> + let rec vars bounded vs = Loc.with_unloc @@ (function + | GVar (id') -> if Id.Set.mem id' bounded then vs else Id.Set.add id' vs + | GApp (f,args) -> List.fold_left (vars bounded) vs (f::args) + | GLambda (na,_,ty,c) | GProd (na,_,ty,c) -> let vs' = vars bounded vs ty in let bounded' = add_name_to_ids bounded na in vars bounded' vs' c - | GLetIn (loc,na,b,ty,c) -> + | GLetIn (na,b,ty,c) -> let vs' = vars bounded vs b in let vs'' = Option.fold_left (vars bounded) vs' ty in let bounded' = add_name_to_ids bounded na in vars bounded' vs'' c - | GCases (loc,sty,rtntypopt,tml,pl) -> + | GCases (sty,rtntypopt,tml,pl) -> let vs1 = vars_option bounded vs rtntypopt in let vs2 = List.fold_left (fun vs (tm,_) -> vars bounded vs tm) vs1 tml in List.fold_left (vars_pattern bounded) vs2 pl - | GLetTuple (loc,nal,rtntyp,b,c) -> + | GLetTuple (nal,rtntyp,b,c) -> let vs1 = vars_return_type bounded vs rtntyp in let vs2 = vars bounded vs1 b in let bounded' = List.fold_left add_name_to_ids bounded nal in vars bounded' vs2 c - | GIf (loc,c,rtntyp,b1,b2) -> + | GIf (c,rtntyp,b1,b2) -> let vs1 = vars_return_type bounded vs rtntyp in let vs2 = vars bounded vs1 c in let vs3 = vars bounded vs2 b1 in vars bounded vs3 b2 - | GRec (loc,fk,idl,bl,tyl,bv) -> + | GRec (fk,idl,bl,tyl,bv) -> let bounded' = Array.fold_right Id.Set.add idl bounded in let vars_fix i vs fid = let vs1,bounded1 = @@ -312,11 +314,12 @@ let free_glob_vars = vars bounded1 vs2 bv.(i) in Array.fold_left_i vars_fix vs idl - | GCast (loc,c,k) -> let v = vars bounded vs c in + | GCast (c,k) -> let v = vars bounded vs c in (match k with CastConv t | CastVM t | CastNative t -> vars bounded v t | _ -> v) | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> vs + ) - and vars_pattern bounded vs (loc,idl,p,c) = + and vars_pattern bounded vs (loc,(idl,p,c)) = let bounded' = List.fold_right Id.Set.add idl bounded in vars bounded' vs c @@ -332,7 +335,7 @@ let free_glob_vars = let glob_visible_short_qualid c = let rec aux acc = function - | GRef (_,c,_) -> + | _, GRef (c,_) -> let qualid = Nametab.shortest_qualid_of_global Id.Set.empty c in let dir,id = Libnames.repr_qualid qualid in if DirPath.is_empty dir then id :: acc else acc @@ -351,26 +354,26 @@ let add_and_check_ident id set = Id.Set.add id set let bound_glob_vars = - let rec vars bound = function - | GLambda (_,na,_,_,_) | GProd (_,na,_,_,_) | GLetIn (_,na,_,_,_) as c -> + let rec vars bound = Loc.with_loc (fun ~loc -> function + | GLambda (na,_,_,_) | GProd (na,_,_,_) | GLetIn (na,_,_,_) as c -> let bound = name_fold add_and_check_ident na bound in - fold_glob_constr vars bound c - | GCases (loc,sty,rtntypopt,tml,pl) -> + fold_glob_constr vars bound (loc, c) + | GCases (sty,rtntypopt,tml,pl) -> let bound = vars_option bound rtntypopt in let bound = List.fold_left (fun bound (tm,_) -> vars bound tm) bound tml in List.fold_left vars_pattern bound pl - | GLetTuple (loc,nal,rtntyp,b,c) -> + | GLetTuple (nal,rtntyp,b,c) -> let bound = vars_return_type bound rtntyp in let bound = vars bound b in let bound = List.fold_right (name_fold add_and_check_ident) nal bound in vars bound c - | GIf (loc,c,rtntyp,b1,b2) -> + | GIf (c,rtntyp,b1,b2) -> let bound = vars_return_type bound rtntyp in let bound = vars bound c in let bound = vars bound b1 in vars bound b2 - | GRec (loc,fk,idl,bl,tyl,bv) -> + | GRec (fk,idl,bl,tyl,bv) -> let bound = Array.fold_right Id.Set.add idl bound in let vars_fix i bound fid = let bound = @@ -388,9 +391,10 @@ let bound_glob_vars = in Array.fold_left_i vars_fix bound idl | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GVar _) -> bound - | GApp _ | GCast _ as c -> fold_glob_constr vars bound c + | GApp _ | GCast _ as c -> fold_glob_constr vars bound (loc, c) + ) - and vars_pattern bound (loc,idl,p,c) = + and vars_pattern bound (loc,(idl,p,c)) = let bound = List.fold_right add_and_check_ident idl bound in vars bound c @@ -435,14 +439,14 @@ let rec map_case_pattern_binders f = Loc.map (function else PatCstr(c,rps,rna) ) -let map_cases_branch_binders f ((loc,il,cll,rhs) as x) : cases_clause = +let map_cases_branch_binders f ((loc,(il,cll,rhs)) as x) : cases_clause = (* spiwack: not sure if I must do something with the list of idents. It is intended to be a superset of the free variable of the right-hand side, if I understand correctly. But I'm not sure when or how they are used. *) let r = List.smartmap (fun cl -> map_case_pattern_binders f cl) cll in if r == cll then x - else loc,il,r,rhs + else loc,(il,r,rhs) let map_pattern_binders f tomatch branches = CList.smartmap (fun tm -> map_tomatch_binders f tm) tomatch, @@ -452,29 +456,14 @@ let map_pattern_binders f tomatch branches = let map_tomatch f (c,pp) : tomatch_tuple = f c , pp -let map_cases_branch f (loc,il,cll,rhs) : cases_clause = - loc , il , cll , f rhs +let map_cases_branch f (loc,(il,cll,rhs)) : cases_clause = + loc , (il , cll , f rhs) let map_pattern f tomatch branches = List.map (fun tm -> map_tomatch f tm) tomatch, List.map (fun br -> map_cases_branch f br) branches -let loc_of_glob_constr = function - | GRef (loc,_,_) -> loc - | GVar (loc,_) -> loc - | GEvar (loc,_,_) -> loc - | GPatVar (loc,_) -> loc - | GApp (loc,_,_) -> loc - | GLambda (loc,_,_,_,_) -> loc - | GProd (loc,_,_,_,_) -> loc - | GLetIn (loc,_,_,_,_) -> loc - | GCases (loc,_,_,_,_) -> loc - | GLetTuple (loc,_,_,_,_) -> loc - | GIf (loc,_,_,_,_) -> loc - | GRec (loc,_,_,_,_,_) -> loc - | GSort (loc,_) -> loc - | GHole (loc,_,_,_) -> loc - | GCast (loc,_,_) -> loc +let loc_of_glob_constr (loc, _) = loc (**********************************************************************) (* Alpha-renaming *) @@ -506,73 +495,74 @@ let rename_var l id = if List.exists (fun (_,id') -> Id.equal id id') l then raise UnsoundRenaming else id -let rec rename_glob_vars l = function - | GVar (loc,id) as r -> +let rec rename_glob_vars l = Loc.map_with_loc (fun ~loc -> function + | GVar id as r -> let id' = rename_var l id in - if id == id' then r else GVar (loc,id') - | GRef (_,VarRef id,_) as r -> + if id == id' then r else GVar id' + | GRef (VarRef id,_) as r -> if List.exists (fun (_,id') -> Id.equal id id') l then raise UnsoundRenaming else r - | GProd (loc,na,bk,t,c) -> + | GProd (na,bk,t,c) -> let na',l' = update_subst na l in - GProd (loc,na,bk,rename_glob_vars l t,rename_glob_vars l' c) - | GLambda (loc,na,bk,t,c) -> + GProd (na,bk,rename_glob_vars l t,rename_glob_vars l' c) + | GLambda (na,bk,t,c) -> let na',l' = update_subst na l in - GLambda (loc,na',bk,rename_glob_vars l t,rename_glob_vars l' c) - | GLetIn (loc,na,b,t,c) -> + GLambda (na',bk,rename_glob_vars l t,rename_glob_vars l' c) + | GLetIn (na,b,t,c) -> let na',l' = update_subst na l in - GLetIn (loc,na',rename_glob_vars l b,Option.map (rename_glob_vars l) t,rename_glob_vars l' c) + GLetIn (na',rename_glob_vars l b,Option.map (rename_glob_vars l) t,rename_glob_vars l' c) (* Lazy strategy: we fail if a collision with renaming occurs, rather than renaming further *) - | GCases (loc,ci,po,tomatchl,cls) -> + | GCases (ci,po,tomatchl,cls) -> let test_pred_pat (na,ino) = test_na l na; Option.iter (fun (_,(_,nal)) -> List.iter (test_na l) nal) ino in let test_clause idl = List.iter (test_id l) idl in let po = Option.map (rename_glob_vars l) po in let tomatchl = Util.List.map_left (fun (tm,x) -> test_pred_pat x; (rename_glob_vars l tm,x)) tomatchl in - let cls = Util.List.map_left (fun (loc,idl,p,c) -> test_clause idl; (loc,idl,p,rename_glob_vars l c)) cls in - GCases (loc,ci,po,tomatchl,cls) - | GLetTuple (loc,nal,(na,po),c,b) -> + let cls = Util.List.map_left (fun (loc,(idl,p,c)) -> test_clause idl; (loc,(idl,p,rename_glob_vars l c))) cls in + GCases (ci,po,tomatchl,cls) + | GLetTuple (nal,(na,po),c,b) -> List.iter (test_na l) (na::nal); - GLetTuple (loc,nal,(na,Option.map (rename_glob_vars l) po), + GLetTuple (nal,(na,Option.map (rename_glob_vars l) po), rename_glob_vars l c,rename_glob_vars l b) - | GIf (loc,c,(na,po),b1,b2) -> + | GIf (c,(na,po),b1,b2) -> test_na l na; - GIf (loc,rename_glob_vars l c,(na,Option.map (rename_glob_vars l) po), + GIf (rename_glob_vars l c,(na,Option.map (rename_glob_vars l) po), rename_glob_vars l b1,rename_glob_vars l b2) - | GRec (loc,k,idl,decls,bs,ts) -> + | GRec (k,idl,decls,bs,ts) -> Array.iter (test_id l) idl; - GRec (loc,k,idl, + GRec (k,idl, Array.map (List.map (fun (na,k,bbd,bty) -> test_na l na; (na,k,Option.map (rename_glob_vars l) bbd,rename_glob_vars l bty))) decls, Array.map (rename_glob_vars l) bs, Array.map (rename_glob_vars l) ts) - | r -> map_glob_constr (rename_glob_vars l) r + (* XXX: This located use case should be improved. *) + | r -> snd @@ map_glob_constr (rename_glob_vars l) (loc, r) + ) (**********************************************************************) (* Conversion from glob_constr to cases pattern, if possible *) -let rec cases_pattern_of_glob_constr na = function - | GVar (loc,id) -> +let rec cases_pattern_of_glob_constr na = Loc.map (function + | GVar id -> begin match na with | Name _ -> (* Unable to manage the presence of both an alias and a variable *) raise Not_found - | Anonymous -> Loc.tag ~loc @@ PatVar (Name id) + | Anonymous -> PatVar (Name id) end - | GHole (loc,_,_,_) -> Loc.tag ~loc @@ PatVar na - | GRef (loc,ConstructRef cstr,_) -> - Loc.tag ~loc @@ PatCstr (cstr,[],na) - | GApp (loc,GRef (_,ConstructRef cstr,_),l) -> - Loc.tag ~loc @@ PatCstr (cstr,List.map (cases_pattern_of_glob_constr Anonymous) l,na) + | GHole (_,_,_) -> PatVar na + | GRef (ConstructRef cstr,_) -> PatCstr (cstr,[],na) + | GApp ((_loc, GRef (ConstructRef cstr,_)),l) -> + PatCstr (cstr,List.map (cases_pattern_of_glob_constr Anonymous) l,na) | _ -> raise Not_found + ) (* Turn a closed cases pattern into a glob_constr *) -let rec glob_constr_of_closed_cases_pattern_aux x = Loc.with_loc (fun ~loc -> function - | PatCstr (cstr,[],Anonymous) -> - GRef (loc,ConstructRef cstr,None) - | PatCstr (cstr,l,Anonymous) -> - let ref = GRef (loc,ConstructRef cstr,None) in - GApp (loc,ref, List.map glob_constr_of_closed_cases_pattern_aux l) +let rec glob_constr_of_closed_cases_pattern_aux x = Loc.map_with_loc (fun ~loc -> function + | PatCstr (cstr,[],Anonymous) -> GRef (ConstructRef cstr,None) + | PatCstr (cstr,l,Anonymous) -> + let ref = Loc.tag ~loc @@ GRef (ConstructRef cstr,None) in + GApp (ref, List.map glob_constr_of_closed_cases_pattern_aux l) | _ -> raise Not_found ) x diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index 48ae93f3e..6696e174b 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -324,46 +324,46 @@ let warn_cast_in_pattern = CWarnings.create ~name:"cast-in-pattern" ~category:"automation" (fun () -> Pp.strbrk "Casts are ignored in patterns") -let rec pat_of_raw metas vars = function - | GVar (_,id) -> +let rec pat_of_raw metas vars = Loc.with_loc (fun ~loc -> function + | GVar id -> (try PRel (List.index Name.equal (Name id) vars) with Not_found -> PVar id) - | GPatVar (_,(false,n)) -> + | GPatVar (false,n) -> metas := n::!metas; PMeta (Some n) - | GRef (_,gr,_) -> + | GRef (gr,_) -> PRef (canonical_gr gr) (* Hack to avoid rewriting a complete interpretation of patterns *) - | GApp (_, GPatVar (_,(true,n)), cl) -> + | GApp ((_, GPatVar (true,n)), cl) -> metas := n::!metas; PSoApp (n, List.map (pat_of_raw metas vars) cl) - | GApp (_,c,cl) -> + | GApp (c,cl) -> PApp (pat_of_raw metas vars c, Array.of_list (List.map (pat_of_raw metas vars) cl)) - | GLambda (_,na,bk,c1,c2) -> + | GLambda (na,bk,c1,c2) -> name_iter (fun n -> metas := n::!metas) na; PLambda (na, pat_of_raw metas vars c1, pat_of_raw metas (na::vars) c2) - | GProd (_,na,bk,c1,c2) -> + | GProd (na,bk,c1,c2) -> name_iter (fun n -> metas := n::!metas) na; PProd (na, pat_of_raw metas vars c1, pat_of_raw metas (na::vars) c2) - | GLetIn (_,na,c1,t,c2) -> + | GLetIn (na,c1,t,c2) -> name_iter (fun n -> metas := n::!metas) na; PLetIn (na, pat_of_raw metas vars c1, Option.map (pat_of_raw metas vars) t, pat_of_raw metas (na::vars) c2) - | GSort (_,s) -> + | GSort s -> PSort s | GHole _ -> PMeta None - | GCast (_,c,_) -> + | GCast (c,_) -> warn_cast_in_pattern (); pat_of_raw metas vars c - | GIf (_,c,(_,None),b1,b2) -> + | GIf (c,(_,None),b1,b2) -> PIf (pat_of_raw metas vars c, pat_of_raw metas vars b1,pat_of_raw metas vars b2) - | GLetTuple (loc,nal,(_,None),b,c) -> - let mkGLambda c na = - GLambda (loc,na,Explicit,GHole (loc,Evar_kinds.InternalHole, IntroAnonymous, None),c) in + | GLetTuple (nal,(_,None),b,c) -> + let mkGLambda c na = Loc.tag ~loc @@ + GLambda (na,Explicit, Loc.tag @@ GHole (Evar_kinds.InternalHole, IntroAnonymous, None),c) in let c = List.fold_left mkGLambda c nal in let cip = { cip_style = LetStyle; @@ -374,9 +374,9 @@ let rec pat_of_raw metas vars = function let tags = List.map (fun _ -> false) nal (* Approximation which can be without let-ins... *) in PCase (cip, PMeta None, pat_of_raw metas vars b, [0,tags,pat_of_raw metas vars c]) - | GCases (loc,sty,p,[c,(na,indnames)],brs) -> + | GCases (sty,p,[c,(na,indnames)],brs) -> let get_ind = function - | (_,_,[_, PatCstr((ind,_),_,_)],_)::_ -> Some ind + | (_,(_,[_, PatCstr((ind,_),_,_)],_))::_ -> Some ind | _ -> None in let ind_tags,ind = match indnames with @@ -389,7 +389,7 @@ let rec pat_of_raw metas vars = function | Some p, Some (_,(_,nal)) -> let nvars = na :: List.rev nal @ vars in rev_it_mkPLambda nal (mkPLambda na (pat_of_raw metas nvars p)) - | (None | Some (GHole _)), _ -> PMeta None + | (None | Some (_, GHole _)), _ -> PMeta None | Some p, None -> user_err ~loc (strbrk "Clause \"in\" expected in patterns over \"match\" expressions with an explicit \"return\" clause.") in @@ -404,7 +404,8 @@ let rec pat_of_raw metas vars = function one non-trivial branch. These facts are used in [Constrextern]. *) PCase (info, pred, pat_of_raw metas vars c, brs) - | r -> err ~loc:(loc_of_glob_constr r) (Pp.str "Non supported pattern.") + | r -> err ~loc (Pp.str "Non supported pattern.") + ) and pats_of_glob_branches loc metas vars ind brs = let get_arg = function @@ -415,8 +416,8 @@ and pats_of_glob_branches loc metas vars ind brs = in let rec get_pat indexes = function | [] -> false, [] - | [(_,_,[_, PatVar(Anonymous)],GHole _)] -> true, [] (* ends with _ => _ *) - | (_,_,[_, PatCstr((indsp,j),lv,_)],br) :: brs -> + | [(_,(_,[_, PatVar(Anonymous)],(_,GHole _)))] -> true, [] (* ends with _ => _ *) + | (_,(_,[_, PatCstr((indsp,j),lv,_)],br)) :: brs -> let () = match ind with | Some sp when eq_ind sp indsp -> () | _ -> @@ -431,7 +432,7 @@ and pats_of_glob_branches loc metas vars ind brs = let ext,pats = get_pat (Int.Set.add (j-1) indexes) brs in let tags = List.map (fun _ -> false) lv (* approximation, w/o let-in *) in ext, ((j-1, tags, pat) :: pats) - | (loc,_,_,_) :: _ -> err ~loc (Pp.str "Non supported pattern.") + | (loc,(_,_,_)) :: _ -> err ~loc (Pp.str "Non supported pattern.") in get_pat Int.Set.empty brs diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index ae87cd8c0..5f9f4bb08 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -567,23 +567,23 @@ let (f_genarg_interp, genarg_interp_hook) = Hook.make () (* in environment [env], with existential variables [evdref] and *) (* the type constraint tycon *) -let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdref (lvar : ltac_var_map) t = +let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdref (lvar : ltac_var_map) (loc, t) = let inh_conv_coerce_to_tycon = inh_conv_coerce_to_tycon resolve_tc in let pretype_type = pretype_type k0 resolve_tc in let pretype = pretype k0 resolve_tc in let open Context.Rel.Declaration in match t with - | GRef (loc,ref,u) -> + | GRef (ref,u) -> inh_conv_coerce_to_tycon loc env evdref (pretype_ref loc evdref env ref u) tycon - | GVar (loc, id) -> + | GVar id -> inh_conv_coerce_to_tycon loc env evdref (pretype_id (fun e r l t -> pretype tycon e r l t) k0 loc env evdref lvar id) tycon - | GEvar (loc, id, inst) -> + | GEvar (id, inst) -> (* Ne faudrait-il pas s'assurer que hyps est bien un sous-contexte du contexte courant, et qu'il n'y a pas de Rel "caché" *) let evk = @@ -596,7 +596,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let j = (Retyping.get_judgment_of env.ExtraEnv.env !evdref c) in inh_conv_coerce_to_tycon loc env evdref j tycon - | GPatVar (loc,(someta,n)) -> + | GPatVar (someta,n) -> let env = ltac_interp_name_env k0 lvar env !evdref in let ty = match tycon with @@ -605,7 +605,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let k = Evar_kinds.MatchingVar (someta,n) in { uj_val = e_new_evar env evdref ~src:(loc,k) ty; uj_type = ty } - | GHole (loc, k, naming, None) -> + | GHole (k, naming, None) -> let env = ltac_interp_name_env k0 lvar env !evdref in let ty = match tycon with @@ -614,7 +614,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre new_type_evar env evdref loc in { uj_val = e_new_evar env evdref ~src:(loc,k) ~naming ty; uj_type = ty } - | GHole (loc, k, _naming, Some arg) -> + | GHole (k, _naming, Some arg) -> let env = ltac_interp_name_env k0 lvar env !evdref in let ty = match tycon with @@ -626,7 +626,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let () = evdref := sigma in { uj_val = c; uj_type = ty } - | GRec (loc,fixkind,names,bl,lar,vdef) -> + | GRec (fixkind,names,bl,lar,vdef) -> let rec type_bl env ctxt = function [] -> ctxt | (na,bk,None,ty)::bl -> @@ -711,11 +711,11 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre in inh_conv_coerce_to_tycon loc env evdref fixj tycon - | GSort (loc,s) -> + | GSort s -> let j = pretype_sort loc evdref s in inh_conv_coerce_to_tycon loc env evdref j tycon - | GApp (loc,f,args) -> + | GApp (f,args) -> let fj = pretype empty_tycon env evdref lvar f in let floc = loc_of_glob_constr f in let length = List.length args in @@ -794,7 +794,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre in inh_conv_coerce_to_tycon loc env evdref resj tycon - | GLambda(loc,name,bk,c1,c2) -> + | GLambda(name,bk,c1,c2) -> let tycon' = evd_comb1 (fun evd tycon -> match tycon with @@ -816,7 +816,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let resj = judge_of_abstraction env.ExtraEnv.env (orelse_name name name') j j' in inh_conv_coerce_to_tycon loc env evdref resj tycon - | GProd(loc,name,bk,c1,c2) -> + | GProd(name,bk,c1,c2) -> let j = pretype_type empty_valcon env evdref lvar c1 in (* The name specified by ltac is used also to create bindings. So the substitution must also be applied on variables before they are @@ -840,7 +840,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre iraise (e, info) in inh_conv_coerce_to_tycon loc env evdref resj tycon - | GLetIn(loc,name,c1,t,c2) -> + | GLetIn(name,c1,t,c2) -> let tycon1 = match t with | Some t -> @@ -861,7 +861,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre { uj_val = mkLetIn (name, j.uj_val, t, j'.uj_val) ; uj_type = subst1 j.uj_val j'.uj_type } - | GLetTuple (loc,nal,(na,po),c,d) -> + | GLetTuple (nal,(na,po),c,d) -> let cj = pretype empty_tycon env evdref lvar c in let (IndType (indf,realargs)) = try find_rectype env.ExtraEnv.env !evdref cj.uj_type @@ -954,7 +954,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre obj ind p cj.uj_val fj.uj_val in { uj_val = v; uj_type = ccl }) - | GIf (loc,c,(na,po),b1,b2) -> + | GIf (c,(na,po),b1,b2) -> let cj = pretype empty_tycon env evdref lvar c in let (IndType (indf,realargs)) = try find_rectype env.ExtraEnv.env !evdref cj.uj_type @@ -1022,12 +1022,12 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let cj = { uj_val = v; uj_type = p } in inh_conv_coerce_to_tycon loc env evdref cj tycon - | GCases (loc,sty,po,tml,eqns) -> + | GCases (sty,po,tml,eqns) -> Cases.compile_cases loc sty ((fun vtyc env evdref -> pretype vtyc (make_env env !evdref) evdref lvar),evdref) tycon env.ExtraEnv.env (* loc *) (po,tml,eqns) - | GCast (loc,c,k) -> + | GCast (c,k) -> let cj = match k with | CastCoerce -> @@ -1097,7 +1097,7 @@ and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update = (* [pretype_type valcon env evdref lvar c] coerces [c] into a type *) and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function - | GHole (loc, knd, naming, None) -> + | loc, GHole (knd, naming, None) -> let rec is_Type c = match EConstr.kind !evdref c with | Sort s -> begin match ESorts.kind !evdref s with diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index 851554b83..2c331ba56 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -253,16 +253,16 @@ open Decl_kinds open Evar_kinds let mkPattern c = snd (Patternops.pattern_of_glob_constr c) -let mkGApp f args = GApp (Loc.ghost, f, args) -let mkGHole = - GHole (Loc.ghost, QuestionMark (Define false), Misctypes.IntroAnonymous, None) -let mkGProd id c1 c2 = - GProd (Loc.ghost, Name (Id.of_string id), Explicit, c1, c2) -let mkGArrow c1 c2 = - GProd (Loc.ghost, Anonymous, Explicit, c1, c2) -let mkGVar id = GVar (Loc.ghost, Id.of_string id) -let mkGPatVar id = GPatVar(Loc.ghost, (false, Id.of_string id)) -let mkGRef r = GRef (Loc.ghost, Lazy.force r, None) +let mkGApp f args = Loc.tag @@ GApp (f, args) +let mkGHole = Loc.tag @@ + GHole (QuestionMark (Define false), Misctypes.IntroAnonymous, None) +let mkGProd id c1 c2 = Loc.tag @@ + GProd (Name (Id.of_string id), Explicit, c1, c2) +let mkGArrow c1 c2 = Loc.tag @@ + GProd (Anonymous, Explicit, c1, c2) +let mkGVar id = Loc.tag @@ GVar (Id.of_string id) +let mkGPatVar id = Loc.tag @@ GPatVar((false, Id.of_string id)) +let mkGRef r = Loc.tag @@ GRef (Lazy.force r, None) let mkGAppRef r args = mkGApp (mkGRef r) args (** forall x : _, _ x x *) diff --git a/vernac/command.ml b/vernac/command.ml index 1f1464856..446afb578 100644 --- a/vernac/command.ml +++ b/vernac/command.ml @@ -422,13 +422,13 @@ let prepare_param = function let rec check_anonymous_type ind = let open Glob_term in - match ind with - | GSort (_, GType []) -> true - | GProd (_, _, _, _, e) - | GLetIn (_, _, _, _, e) - | GLambda (_, _, _, _, e) - | GApp (_, e, _) - | GCast (_, e, _) -> check_anonymous_type e + match snd ind with + | GSort (GType []) -> true + | GProd ( _, _, _, e) + | GLetIn (_, _, _, e) + | GLambda (_, _, _, e) + | GApp (e, _) + | GCast (e, _) -> check_anonymous_type e | _ -> false let make_conclusion_flexible evdref ty poly = -- cgit v1.2.3 From a9d151a31937724543d5269e72b0262c8764c46e Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Tue, 17 Jan 2017 14:44:28 +0100 Subject: [location] More located use. --- interp/constrexpr_ops.ml | 4 ++-- interp/constrextern.ml | 2 +- interp/constrintern.ml | 2 +- interp/topconstr.ml | 8 ++++---- intf/constrexpr.mli | 6 +++--- parsing/g_constr.ml4 | 2 +- printing/ppconstr.ml | 8 ++++---- vernac/record.ml | 2 +- 8 files changed, 17 insertions(+), 17 deletions(-) (limited to 'intf') diff --git a/interp/constrexpr_ops.ml b/interp/constrexpr_ops.ml index 4f23dd2ab..61115c00b 100644 --- a/interp/constrexpr_ops.ml +++ b/interp/constrexpr_ops.ml @@ -239,7 +239,7 @@ let local_binder_loc = function | CLocalDef ((loc,_),t,None) -> Loc.merge loc (constr_loc t) | CLocalDef ((loc,_),b,Some t) -> Loc.merge loc (Loc.merge (constr_loc b) (constr_loc t)) | CLocalAssum ([],_,_) -> assert false - | CLocalPattern (loc,_,_) -> loc + | CLocalPattern (loc,_) -> loc let local_binders_loc bll = match bll with | [] -> Loc.ghost @@ -283,7 +283,7 @@ let expand_binders ~loc mkC bl c = let env = List.fold_left add_name_in_env env nl in (env, mkC ~loc (nl,bk,t) c) | CLocalAssum ([],_,_) -> loop loc bl c - | CLocalPattern (loc1, p, ty) -> + | CLocalPattern (loc1, (p, ty)) -> 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 diff --git a/interp/constrextern.ml b/interp/constrextern.ml index bbc98dd28..8d9f8552d 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -866,7 +866,7 @@ and extern_local_binder scopes vars = function if !Flags.raw_print then Some (extern_typ scopes vars ty) else None in let p = extern_cases_pattern vars p in let (assums,ids,l) = extern_local_binder scopes vars l in - (assums,ids, CLocalPattern(Loc.ghost,p,ty) :: l) + (assums,ids, CLocalPattern(Loc.tag @@ (p,ty)) :: l) and extern_eqn inctx scopes vars (loc,(ids,pl,c)) = Loc.tag ~loc ([loc,List.map (extern_cases_pattern_in_scope scopes vars) pl], diff --git a/interp/constrintern.ml b/interp/constrintern.ml index cc7203ac0..d1b931a22 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -475,7 +475,7 @@ let intern_local_binder_aux ?(global_level=false) intern lvar (env,bl) = functio let ty = Option.map (intern env) ty in (push_name_env lvar (impls_term_list term) env locna, GLocalDef (loc,na,Explicit,term,ty) :: bl) - | CLocalPattern (loc,p,ty) -> + | CLocalPattern (loc,(p,ty)) -> let tyc = match ty with | Some ty -> ty diff --git a/interp/topconstr.ml b/interp/topconstr.ml index c3e341d74..c8fbdaf28 100644 --- a/interp/topconstr.ml +++ b/interp/topconstr.ml @@ -97,7 +97,7 @@ let rec fold_local_binders g f n acc b = function f n (fold_local_binders g f n' acc b l) t | CLocalDef ((_,na),c,t)::l -> Option.fold_left (f n) (f n (fold_local_binders g f (name_fold g na n) acc b l) c) t - | CLocalPattern (_,pat,t)::l -> + | CLocalPattern (_,(pat,t))::l -> let acc = fold_local_binders g f (cases_pattern_fold_names g n pat) acc b l in Option.fold_left (f n) acc t | [] -> @@ -180,7 +180,7 @@ let split_at_annot bl na = (List.rev ans, CLocalAssum (r, k, t) :: rest) end | CLocalDef _ as x :: rest -> aux (x :: acc) rest - | CLocalPattern (loc,_,_) :: rest -> + | CLocalPattern (loc,_) :: rest -> Loc.raise ~loc (Stream.Error "pattern with quote not allowed after fix") | [] -> user_err ~loc @@ -204,9 +204,9 @@ let map_local_binders f g e bl = (map_binder g e nal, CLocalAssum(nal,k,f e ty)::bl) | CLocalDef((loc,na),c,ty) -> (name_fold g na e, CLocalDef((loc,na),f e c,Option.map (f e) ty)::bl) - | CLocalPattern (loc,pat,t) -> + | CLocalPattern (loc,(pat,t)) -> let ids = ids_of_pattern pat in - (Id.Set.fold g ids e, CLocalPattern (loc,pat,Option.map (f e) t)::bl) in + (Id.Set.fold g ids e, CLocalPattern (loc,(pat,Option.map (f e) t))::bl) in let (e,rbl) = List.fold_left h (e,[]) bl in (e, List.rev rbl) diff --git a/intf/constrexpr.mli b/intf/constrexpr.mli index c1ea71df4..92d0020a5 100644 --- a/intf/constrexpr.mli +++ b/intf/constrexpr.mli @@ -125,9 +125,9 @@ and recursion_order_expr = (** Anonymous defs allowed ?? *) and local_binder_expr = - | CLocalAssum of Name.t located list * binder_kind * constr_expr - | CLocalDef of Name.t located * constr_expr * constr_expr option - | CLocalPattern of Loc.t * cases_pattern_expr * constr_expr option + | CLocalAssum of Name.t located list * binder_kind * constr_expr + | CLocalDef of Name.t located * constr_expr * constr_expr option + | CLocalPattern of (cases_pattern_expr * constr_expr option) Loc.located and constr_notation_substitution = constr_expr list * (** for constr subterms *) diff --git a/parsing/g_constr.ml4 b/parsing/g_constr.ml4 index 9bf00b0b1..33e7236f5 100644 --- a/parsing/g_constr.ml4 +++ b/parsing/g_constr.ml4 @@ -486,7 +486,7 @@ GEXTEND Gram | _, CPatCast (p, ty) -> (p, Some ty) | _ -> (p, None) in - [CLocalPattern (!@loc, p, ty)] + [CLocalPattern (Loc.tag ~loc:!@loc (p, ty))] ] ] ; typeclass_constraint: diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml index a99c0951f..117e1900d 100644 --- a/printing/ppconstr.ml +++ b/printing/ppconstr.ml @@ -319,7 +319,7 @@ let tag_var = tag Tag.variable let begin_of_binder = function | CLocalDef((loc,_),_,_) -> fst (Loc.unloc loc) | CLocalAssum((loc,_)::_,_,_) -> fst (Loc.unloc loc) - | CLocalPattern(loc,_,_) -> fst (Loc.unloc loc) + | CLocalPattern(loc,(_,_)) -> fst (Loc.unloc loc) | _ -> assert false let begin_of_binders = function @@ -366,7 +366,7 @@ let tag_var = tag Tag.variable surround (pr_lname na ++ pr_opt_no_spc (fun t -> str " :" ++ ws 1 ++ pr_c t) topt ++ str" :=" ++ spc() ++ pr_c c) - | CLocalPattern (loc,p,tyo) -> + | CLocalPattern (loc,(p,tyo)) -> let p = pr_patt lsimplepatt p in match tyo with | None -> @@ -400,7 +400,7 @@ let tag_var = tag Tag.variable (_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 + 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 @@ -416,7 +416,7 @@ let tag_var = tag Tag.variable (_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 + CLocalPattern (loc,(p,None)) :: bl, c | CLambdaN ((nal,bk,t)::bl,c) -> let bl,c = extract_lam_binders (loc, CLambdaN(bl,c)) in CLocalAssum (nal,bk,t) :: bl, c diff --git a/vernac/record.ml b/vernac/record.ml index 37ce231f9..95f5ad7cc 100644 --- a/vernac/record.ml +++ b/vernac/record.ml @@ -112,7 +112,7 @@ let typecheck_params_and_fields def id pl t ps nots fs = List.iter (function CLocalDef (b, _, _) -> error default_binder_kind b | CLocalAssum (ls, bk, ce) -> List.iter (error bk) ls - | CLocalPattern (loc,_,_) -> + | CLocalPattern (loc,(_,_)) -> Loc.raise ~loc (Stream.Error "pattern with quote not allowed in record parameters.")) ps in let impls_env, ((env1,newps), imps) = interp_context_evars env0 evars ps in -- cgit v1.2.3 From bf13037e9ca39da28fb648e5488ce56ef8a1f1e2 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Tue, 17 Jan 2017 15:06:26 +0100 Subject: [location] Use located in misctypes. --- ide/texmacspp.ml | 2 +- interp/smartlocate.ml | 6 +++--- interp/stdarg.ml | 2 +- intf/genredexpr.mli | 2 +- intf/misctypes.mli | 12 ++++++------ intf/vernacexpr.mli | 2 +- parsing/g_prim.ml4 | 2 +- parsing/pcoq.mli | 2 +- plugins/funind/g_indfun.ml4 | 4 ++-- plugins/funind/invfun.ml | 4 ++-- plugins/funind/recdef.ml | 10 ++++------ plugins/ltac/g_tactic.ml4 | 4 ++-- plugins/ltac/pptactic.ml | 10 +++++----- plugins/ltac/tacintern.ml | 6 +++--- plugins/ltac/tacinterp.ml | 6 +++--- plugins/ltac/tacsubst.ml | 4 ++-- pretyping/miscops.ml | 2 +- printing/pputils.ml | 2 +- printing/prettyp.ml | 4 ++-- proofs/clenv.ml | 6 +++--- vernac/vernacentries.ml | 4 ++-- 21 files changed, 47 insertions(+), 49 deletions(-) (limited to 'intf') diff --git a/ide/texmacspp.ml b/ide/texmacspp.ml index 77dca0d06..f86b260df 100644 --- a/ide/texmacspp.ml +++ b/ide/texmacspp.ml @@ -649,7 +649,7 @@ let rec tmpp v loc = match r with | AN (Qualid (_, q)) -> ["qualid", string_of_qualid q] | AN (Ident (_, id)) -> ["id", Id.to_string id] - | ByNotation (_, s, _) -> ["notation", s] in + | ByNotation (_, (s, _)) -> ["notation", s] in xmlCanonicalStructure attr loc | VernacCoercion _ as x -> xmlTODO loc x | VernacIdentityCoercion _ as x -> xmlTODO loc x diff --git a/interp/smartlocate.ml b/interp/smartlocate.ml index d863e0561..64d260cc1 100644 --- a/interp/smartlocate.ml +++ b/interp/smartlocate.ml @@ -66,16 +66,16 @@ let global_with_alias ?head r = let smart_global ?head = function | AN r -> global_with_alias ?head r - | ByNotation (loc,ntn,sc) -> + | ByNotation (loc,(ntn,sc)) -> Notation.interp_notation_as_global_reference loc (fun _ -> true) ntn sc let smart_global_inductive = function | AN r -> global_inductive_with_alias r - | ByNotation (loc,ntn,sc) -> + | ByNotation (loc,(ntn,sc)) -> destIndRef (Notation.interp_notation_as_global_reference loc isIndRef ntn sc) let loc_of_smart_reference = function | AN r -> loc_of_reference r - | ByNotation (loc,_,_) -> loc + | ByNotation (loc,(_,_)) -> loc diff --git a/interp/stdarg.ml b/interp/stdarg.ml index 341ff5662..c0dd9e45c 100644 --- a/interp/stdarg.ml +++ b/interp/stdarg.ml @@ -34,7 +34,7 @@ let wit_pre_ident : string uniform_genarg_type = let loc_of_or_by_notation f = function | AN c -> f c - | ByNotation (loc,s,_) -> loc + | ByNotation (loc,(s,_)) -> loc let wit_int_or_var = make0 ~dyn:(val_tag (topwit wit_int)) "int_or_var" diff --git a/intf/genredexpr.mli b/intf/genredexpr.mli index 16f0c0c92..2a542e0ff 100644 --- a/intf/genredexpr.mli +++ b/intf/genredexpr.mli @@ -52,7 +52,7 @@ type ('a,'b,'c) red_expr_gen = type ('a,'b,'c) may_eval = | ConstrTerm of 'a | ConstrEval of ('a,'b,'c) red_expr_gen * 'a - | ConstrContext of (Loc.t * Id.t) * 'a + | ConstrContext of Id.t Loc.located * 'a | ConstrTypeOf of 'a open Libnames diff --git a/intf/misctypes.mli b/intf/misctypes.mli index 33dc2a335..0157800cd 100644 --- a/intf/misctypes.mli +++ b/intf/misctypes.mli @@ -27,12 +27,12 @@ and intro_pattern_naming_expr = and 'constr intro_pattern_action_expr = | IntroWildcard | IntroOrAndPattern of 'constr or_and_intro_pattern_expr - | IntroInjection of (Loc.t * 'constr intro_pattern_expr) list - | IntroApplyOn of (Loc.t * 'constr) * (Loc.t * 'constr intro_pattern_expr) + | IntroInjection of ('constr intro_pattern_expr) Loc.located list + | IntroApplyOn of 'constr Loc.located * 'constr intro_pattern_expr Loc.located | IntroRewrite of bool and 'constr or_and_intro_pattern_expr = - | IntroOrPattern of (Loc.t * 'constr intro_pattern_expr) list list - | IntroAndPattern of (Loc.t * 'constr intro_pattern_expr) list + | IntroOrPattern of ('constr intro_pattern_expr) Loc.located list list + | IntroAndPattern of ('constr intro_pattern_expr) Loc.located list (** Move destination for hypothesis *) @@ -79,7 +79,7 @@ type 'a cast_type = type quantified_hypothesis = AnonHyp of int | NamedHyp of Id.t -type 'a explicit_bindings = (Loc.t * quantified_hypothesis * 'a) list +type 'a explicit_bindings = (quantified_hypothesis * 'a) Loc.located list type 'a bindings = | ImplicitBindings of 'a list @@ -99,7 +99,7 @@ type 'a and_short_name = 'a * Id.t Loc.located option type 'a or_by_notation = | AN of 'a - | ByNotation of (Loc.t * string * string option) + | ByNotation of (string * string option) Loc.located (* NB: the last string in [ByNotation] is actually a [Notation.delimiters], but this formulation avoids a useless dependency. *) diff --git a/intf/vernacexpr.mli b/intf/vernacexpr.mli index f018d59e6..211d41bb0 100644 --- a/intf/vernacexpr.mli +++ b/intf/vernacexpr.mli @@ -482,7 +482,7 @@ and vernac_implicit_status = Implicit | MaximallyImplicit | NotImplicit and vernac_argument_status = { name : Name.t; recarg_like : bool; - notation_scope : (Loc.t * string) option; + notation_scope : string Loc.located option; implicit_status : vernac_implicit_status; } diff --git a/parsing/g_prim.ml4 b/parsing/g_prim.ml4 index 2af4ed533..ed6a8df4e 100644 --- a/parsing/g_prim.ml4 +++ b/parsing/g_prim.ml4 @@ -81,7 +81,7 @@ GEXTEND Gram ] ] ; by_notation: - [ [ s = ne_string; sc = OPT ["%"; key = IDENT -> key ] -> (!@loc, s, sc) ] ] + [ [ s = ne_string; sc = OPT ["%"; key = IDENT -> key ] -> Loc.tag ~loc:!@loc (s, sc) ] ] ; smart_global: [ [ c = reference -> Misctypes.AN c diff --git a/parsing/pcoq.mli b/parsing/pcoq.mli index 4248db697..959e8ddf5 100644 --- a/parsing/pcoq.mli +++ b/parsing/pcoq.mli @@ -206,7 +206,7 @@ module Prim : val qualid : qualid located Gram.entry val fullyqualid : Id.t list located Gram.entry val reference : reference Gram.entry - val by_notation : (Loc.t * string * string option) Gram.entry + val by_notation : (string * string option) Loc.located Gram.entry val smart_global : reference or_by_notation Gram.entry val dirpath : DirPath.t Gram.entry val ne_string : string Gram.entry diff --git a/plugins/funind/g_indfun.ml4 b/plugins/funind/g_indfun.ml4 index 0dccd25d7..129c8dc16 100644 --- a/plugins/funind/g_indfun.ml4 +++ b/plugins/funind/g_indfun.ml4 @@ -24,8 +24,8 @@ open Pltac DECLARE PLUGIN "recdef_plugin" let pr_binding prc = function - | loc, NamedHyp id, c -> hov 1 (Ppconstr.pr_id id ++ str " := " ++ cut () ++ prc c) - | loc, AnonHyp n, c -> hov 1 (int n ++ str " := " ++ cut () ++ prc c) + | loc, (NamedHyp id, c) -> hov 1 (Ppconstr.pr_id id ++ str " := " ++ cut () ++ prc c) + | loc, (AnonHyp n, c) -> hov 1 (int n ++ str " := " ++ cut () ++ prc c) let pr_bindings prc prlc = function | ImplicitBindings l -> diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index 94ec0a898..8da4f9233 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -31,8 +31,8 @@ module RelDecl = Context.Rel.Declaration let pr_binding prc = function - | loc, NamedHyp id, c -> hov 1 (Ppconstr.pr_id id ++ str " := " ++ Pp.cut () ++ prc c) - | loc, AnonHyp n, c -> hov 1 (int n ++ str " := " ++ Pp.cut () ++ prc c) + | loc, (NamedHyp id, c) -> hov 1 (Ppconstr.pr_id id ++ str " := " ++ Pp.cut () ++ prc c) + | loc, (AnonHyp n, c) -> hov 1 (int n ++ str " := " ++ Pp.cut () ++ prc c) let pr_bindings prc prlc = function | ImplicitBindings l -> diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index c796fe7a2..25daedd0e 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -882,9 +882,8 @@ let rec make_rewrite_list expr_info max = function Proofview.V82.of_tactic (general_rewrite_bindings false Locus.AllOccurrences true (* dep proofs also: *) true (mkVar hp, - ExplicitBindings[Loc.ghost,NamedHyp def, - expr_info.f_constr;Loc.ghost,NamedHyp k, - f_S max]) false) g) ) + ExplicitBindings[Loc.tag @@ (NamedHyp def, expr_info.f_constr); + Loc.tag @@ (NamedHyp k, f_S max)]) false) g) ) ) [make_rewrite_list expr_info max l; observe_tclTHENLIST (str "make_rewrite_list")[ (* x < S max proof *) @@ -910,9 +909,8 @@ let make_rewrite expr_info l hp max = (Proofview.V82.of_tactic (general_rewrite_bindings false Locus.AllOccurrences true (* dep proofs also: *) true (mkVar hp, - ExplicitBindings[Loc.ghost,NamedHyp def, - expr_info.f_constr;Loc.ghost,NamedHyp k, - f_S (f_S max)]) false)) g) + ExplicitBindings[Loc.tag @@ (NamedHyp def, expr_info.f_constr); + Loc.tag @@ (NamedHyp k, f_S (f_S max))]) false)) g) [observe_tac(str "make_rewrite finalize") ( (* tclORELSE( h_reflexivity) *) (observe_tclTHENLIST (str "make_rewrite")[ diff --git a/plugins/ltac/g_tactic.ml4 b/plugins/ltac/g_tactic.ml4 index 1674bb4ca..d1e5c810f 100644 --- a/plugins/ltac/g_tactic.ml4 +++ b/plugins/ltac/g_tactic.ml4 @@ -336,8 +336,8 @@ GEXTEND Gram | pat = naming_intropattern -> !@loc, IntroNaming pat ] ] ; simple_binding: - [ [ "("; id = ident; ":="; c = lconstr; ")" -> (!@loc, NamedHyp id, c) - | "("; n = natural; ":="; c = lconstr; ")" -> (!@loc, AnonHyp n, c) ] ] + [ [ "("; id = ident; ":="; c = lconstr; ")" -> Loc.tag ~loc:!@loc (NamedHyp id, c) + | "("; n = natural; ":="; c = lconstr; ")" -> Loc.tag ~loc:!@loc (AnonHyp n, c) ] ] ; bindings: [ [ test_lpar_idnum_coloneq; bl = LIST1 simple_binding -> diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml index aec2e37fd..0dd6819fd 100644 --- a/plugins/ltac/pptactic.ml +++ b/plugins/ltac/pptactic.ml @@ -149,7 +149,7 @@ type 'a extra_genarg_printer = let pr_or_by_notation f = function | AN v -> f v - | ByNotation (_,s,sc) -> qs s ++ pr_opt (fun sc -> str "%" ++ str sc) sc + | ByNotation (_,(s,sc)) -> qs s ++ pr_opt (fun sc -> str "%" ++ str sc) sc let pr_located pr (loc,x) = pr x @@ -162,8 +162,8 @@ type 'a extra_genarg_printer = | NamedHyp id -> pr_id id let pr_binding prc = function - | loc, NamedHyp id, c -> hov 1 (pr_id id ++ str " := " ++ cut () ++ prc c) - | loc, AnonHyp n, c -> hov 1 (int n ++ str " := " ++ cut () ++ prc c) + | loc, (NamedHyp id, c) -> hov 1 (pr_id id ++ str " := " ++ cut () ++ prc c) + | loc, (AnonHyp n, c) -> hov 1 (int n ++ str " := " ++ cut () ++ prc c) let pr_bindings prc prlc = function | ImplicitBindings l -> @@ -368,8 +368,8 @@ type 'a extra_genarg_printer = let pr_esubst prc l = let pr_qhyp = function - (_,AnonHyp n,c) -> str "(" ++ int n ++ str" := " ++ prc c ++ str ")" - | (_,NamedHyp id,c) -> + (_,(AnonHyp n,c)) -> str "(" ++ int n ++ str" := " ++ prc c ++ str ")" + | (_,(NamedHyp id,c)) -> str "(" ++ pr_id id ++ str" := " ++ prc c ++ str ")" in prlist_with_sep spc pr_qhyp l diff --git a/plugins/ltac/tacintern.ml b/plugins/ltac/tacintern.ml index e7d4c1be9..ad9814b84 100644 --- a/plugins/ltac/tacintern.ml +++ b/plugins/ltac/tacintern.ml @@ -212,8 +212,8 @@ let intern_constr = intern_constr_gen false false let intern_type = intern_constr_gen false true (* Globalize bindings *) -let intern_binding ist (loc,b,c) = - (loc,intern_binding_name ist b,intern_constr ist c) +let intern_binding ist (loc,(b,c)) = + (loc,(intern_binding_name ist b,intern_constr ist c)) let intern_bindings ist = function | NoBindings -> NoBindings @@ -291,7 +291,7 @@ let intern_evaluable_global_reference ist r = let intern_evaluable_reference_or_by_notation ist = function | AN r -> intern_evaluable_global_reference ist r - | ByNotation (loc,ntn,sc) -> + | ByNotation (loc,(ntn,sc)) -> evaluable_of_global_reference ist.genv (Notation.interp_notation_as_global_reference loc (function ConstRef _ | VarRef _ -> true | _ -> false) ntn sc) diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml index a8d8eda1d..aac63fcb2 100644 --- a/plugins/ltac/tacinterp.ml +++ b/plugins/ltac/tacinterp.ml @@ -987,9 +987,9 @@ let interp_declared_or_quantified_hypothesis ist env sigma = function (coerce_to_decl_or_quant_hyp env sigma) ist (Some (env,sigma)) (dloc,id) with Not_found -> NamedHyp id -let interp_binding ist env sigma (loc,b,c) = +let interp_binding ist env sigma (loc,(b,c)) = let sigma, c = interp_open_constr ist env sigma c in - sigma, (loc,interp_binding_name ist sigma b,c) + sigma, (loc,(interp_binding_name ist sigma b,c)) let interp_bindings ist env sigma = function | NoBindings -> @@ -1014,7 +1014,7 @@ let interp_open_constr_with_bindings ist env sigma (c,bl) = let loc_of_bindings = function | NoBindings -> Loc.ghost | ImplicitBindings l -> loc_of_glob_constr (fst (List.last l)) -| ExplicitBindings l -> pi1 (List.last l) +| ExplicitBindings l -> fst (List.last l) let interp_open_constr_with_bindings_loc ist ((c,_),bl as cb) = let loc1 = loc_of_glob_constr c in diff --git a/plugins/ltac/tacsubst.ml b/plugins/ltac/tacsubst.ml index fe3a9f3b2..9582ebd89 100644 --- a/plugins/ltac/tacsubst.ml +++ b/plugins/ltac/tacsubst.ml @@ -32,8 +32,8 @@ let subst_glob_constr_and_expr subst (c, e) = let subst_glob_constr = subst_glob_constr_and_expr (* shortening *) -let subst_binding subst (loc,b,c) = - (loc,subst_quantified_hypothesis subst b,subst_glob_constr subst c) +let subst_binding subst (loc,(b,c)) = + (loc,(subst_quantified_hypothesis subst b,subst_glob_constr subst c)) let subst_bindings subst = function | NoBindings -> NoBindings diff --git a/pretyping/miscops.ml b/pretyping/miscops.ml index 7fe81c9a4..f53677abb 100644 --- a/pretyping/miscops.ml +++ b/pretyping/miscops.ml @@ -62,7 +62,7 @@ let map_red_expr_gen f g h = function (** Mapping bindings *) let map_explicit_bindings f l = - let map (loc, hyp, x) = (loc, hyp, f x) in + let map (loc, (hyp, x)) = (loc, (hyp, f x)) in List.map map l let map_bindings f = function diff --git a/printing/pputils.ml b/printing/pputils.ml index 50630fb9b..e90b54685 100644 --- a/printing/pputils.ml +++ b/printing/pputils.ml @@ -105,7 +105,7 @@ let pr_red_expr (pr_constr,pr_lconstr,pr_ref,pr_pattern) keyword = function let pr_or_by_notation f = function | AN v -> f v - | ByNotation (_,s,sc) -> qs s ++ pr_opt (fun sc -> str "%" ++ str sc) sc + | ByNotation (_,(s,sc)) -> qs s ++ pr_opt (fun sc -> str "%" ++ str sc) sc let hov_if_not_empty n p = if Pp.ismt p then p else hov n p diff --git a/printing/prettyp.ml b/printing/prettyp.ml index aa422e36a..96b0f49d4 100644 --- a/printing/prettyp.ml +++ b/printing/prettyp.ml @@ -750,7 +750,7 @@ let print_any_name = function ~hdr:"print_name" (pr_qualid qid ++ spc () ++ str "not a defined object.") let print_name = function - | ByNotation (loc,ntn,sc) -> + | ByNotation (loc,(ntn,sc)) -> print_any_name (Term (Notation.interp_notation_as_global_reference loc (fun _ -> true) ntn sc)) @@ -798,7 +798,7 @@ let print_about_any loc k = hov 0 (pr_located_qualid k) let print_about = function - | ByNotation (loc,ntn,sc) -> + | ByNotation (loc,(ntn,sc)) -> print_about_any loc (Term (Notation.interp_notation_as_global_reference loc (fun _ -> true) ntn sc)) diff --git a/proofs/clenv.ml b/proofs/clenv.ml index f9ebc4233..17a9651cd 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -421,7 +421,7 @@ let qhyp_eq h1 h2 = match h1, h2 with | _ -> false let check_bindings bl = - match List.duplicates qhyp_eq (List.map pi2 bl) with + match List.duplicates qhyp_eq (List.map (fun x -> fst (snd x)) bl) with | NamedHyp s :: _ -> user_err (str "The variable " ++ pr_id s ++ @@ -517,7 +517,7 @@ let clenv_match_args bl clenv = let mvs = clenv_independent clenv in check_bindings bl; List.fold_left - (fun clenv (loc,b,c) -> + (fun clenv (loc,(b,c)) -> let k = meta_of_binder clenv loc mvs b in if meta_defined clenv.evd k then if EConstr.eq_constr clenv.evd (EConstr.of_constr (fst (meta_fvalue clenv.evd k)).rebus) c then clenv @@ -716,7 +716,7 @@ let solve_evar_clause env sigma hyp_only clause = function error_not_right_number_missing_arguments len | ExplicitBindings lbind -> let () = check_bindings lbind in - let fold sigma (_, binder, c) = + let fold sigma (_, (binder, c)) = let ev = evar_of_binder clause.cl_holes binder in define_with_type sigma env ev c in diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml index 0a5a000fe..86406dbe5 100644 --- a/vernac/vernacentries.ml +++ b/vernac/vernacentries.ml @@ -1772,8 +1772,8 @@ let vernac_search ~loc s gopt r = let vernac_locate = let open Feedback in function | LocateAny (AN qid) -> msg_notice (print_located_qualid qid) | LocateTerm (AN qid) -> msg_notice (print_located_term qid) - | LocateAny (ByNotation (_, ntn, sc)) (** TODO : handle Ltac notations *) - | LocateTerm (ByNotation (_, ntn, sc)) -> + | LocateAny (ByNotation (_, (ntn, sc))) (** TODO : handle Ltac notations *) + | LocateTerm (ByNotation (_, (ntn, sc))) -> msg_notice (Notation.locate_notation (Constrextern.without_symbols pr_lglob_constr) ntn sc) -- cgit v1.2.3 From 30d3515546cf244837c6340b6b87c5f51e68cbf4 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Tue, 17 Jan 2017 23:40:35 +0100 Subject: [location] Remove Loc.ghost. Now it is a private field, locations are optional. --- .merlin | 2 + checker/modops.ml | 6 +- dev/top_printers.ml | 30 +-- engine/evarutil.ml | 2 +- engine/evd.ml | 12 +- engine/evd.mli | 2 +- engine/proofview.ml | 2 +- engine/uState.ml | 6 +- grammar/argextend.mlp | 2 +- grammar/tacextend.mlp | 6 +- ide/coqOps.ml | 22 +- ide/ide_slave.ml | 4 +- ide/texmacspp.ml | 449 ++++++++++++++++---------------- ide/texmacspp.mli | 2 +- ide/xmlprotocol.ml | 4 +- interp/constrexpr_ops.ml | 45 ++-- interp/constrexpr_ops.mli | 6 +- interp/constrextern.ml | 112 ++++---- interp/constrextern.mli | 6 +- interp/constrintern.ml | 99 +++---- interp/implicit_quantifiers.ml | 2 +- interp/notation.ml | 64 ++--- interp/notation.mli | 12 +- interp/notation_ops.ml | 74 +++--- interp/notation_ops.mli | 4 +- interp/smartlocate.ml | 4 +- interp/topconstr.ml | 2 +- intf/glob_term.mli | 9 +- lib/aux_file.ml | 20 +- lib/aux_file.mli | 8 +- lib/cWarnings.ml | 19 +- lib/feedback.ml | 2 +- lib/feedback.mli | 2 +- lib/loc.ml | 17 +- lib/loc.mli | 17 +- lib/stateid.ml | 2 +- lib/stateid.mli | 2 +- library/declare.ml | 4 +- parsing/g_constr.ml4 | 10 +- parsing/g_vernac.ml4 | 16 +- plugins/firstorder/g_ground.ml4 | 2 +- plugins/funind/glob_term_to_relation.ml | 22 +- plugins/funind/indfun.ml | 22 +- plugins/funind/indfun_common.ml | 2 +- plugins/funind/invfun.ml | 2 +- plugins/funind/merge.ml | 12 +- plugins/funind/recdef.ml | 4 +- plugins/ltac/coretactics.ml4 | 5 +- plugins/ltac/evar_tactics.ml | 2 +- plugins/ltac/extraargs.ml4 | 6 +- plugins/ltac/extratactics.ml4 | 8 +- plugins/ltac/g_ltac.ml4 | 4 +- plugins/ltac/g_obligations.ml4 | 2 +- plugins/ltac/pptactic.ml | 42 +-- plugins/ltac/pptactic.mli | 2 +- plugins/ltac/profile_ltac.ml | 4 +- plugins/ltac/rewrite.ml | 41 ++- plugins/ltac/tacentries.ml | 22 +- plugins/ltac/tacentries.mli | 2 +- plugins/ltac/tacexpr.mli | 2 +- plugins/ltac/tacintern.ml | 10 +- plugins/ltac/tacinterp.ml | 52 ++-- plugins/ltac/tacinterp.mli | 2 +- plugins/ltac/tacsubst.ml | 2 +- plugins/ltac/tactic_debug.ml | 19 +- plugins/ltac/tactic_debug.mli | 2 +- plugins/ltac/tauto.ml | 11 +- plugins/micromega/coq_micromega.ml | 4 +- plugins/quote/g_quote.ml4 | 5 +- plugins/setoid_ring/newring.ml | 14 +- plugins/ssrmatching/ssrmatching.ml4 | 34 +-- plugins/syntax/ascii_syntax.ml | 12 +- plugins/syntax/nat_syntax.ml | 10 +- plugins/syntax/numbers_syntax.ml | 68 ++--- plugins/syntax/r_syntax.ml | 12 +- plugins/syntax/string_syntax.ml | 8 +- plugins/syntax/z_syntax.ml | 42 +-- pretyping/cases.ml | 66 ++--- pretyping/cases.mli | 8 +- pretyping/coercion.ml | 65 ++--- pretyping/coercion.mli | 14 +- pretyping/detyping.ml | 2 - pretyping/evarconv.ml | 4 +- pretyping/glob_ops.ml | 2 +- pretyping/glob_ops.mli | 2 +- pretyping/pretyping.ml | 44 ++-- pretyping/pretyping.mli | 2 +- pretyping/typing.ml | 6 +- pretyping/typing.mli | 2 +- pretyping/unification.ml | 4 +- printing/ppconstr.ml | 10 +- printing/ppconstr.mli | 2 +- printing/pputils.ml | 2 +- printing/prettyp.ml | 6 +- printing/printer.ml | 6 +- proofs/clenv.ml | 2 +- proofs/goal.ml | 2 +- proofs/proof_global.ml | 4 +- proofs/proof_using.ml | 2 +- proofs/refine.ml | 2 +- stm/stm.ml | 83 +++--- stm/stm.mli | 2 +- stm/vio_checking.ml | 2 +- tactics/contradiction.ml | 2 +- tactics/equality.ml | 2 +- tactics/inv.ml | 2 +- tactics/tacticals.ml | 24 +- tactics/tacticals.mli | 2 +- tactics/tactics.ml | 100 ++++--- toplevel/coqtop.ml | 4 +- toplevel/vernac.ml | 10 +- vernac/auto_ind_decl.ml | 12 +- vernac/classes.ml | 2 +- vernac/command.ml | 11 +- vernac/explainErr.ml | 4 +- vernac/explainErr.mli | 2 +- vernac/indschemes.ml | 2 +- vernac/lemmas.ml | 2 +- vernac/metasyntax.ml | 2 +- vernac/obligations.ml | 5 +- vernac/vernacentries.ml | 22 +- vernac/vernacentries.mli | 2 +- 122 files changed, 1083 insertions(+), 1099 deletions(-) (limited to 'intf') diff --git a/.merlin b/.merlin index 5cae15f5f..5c8c73851 100644 --- a/.merlin +++ b/.merlin @@ -36,6 +36,8 @@ S toplevel B toplevel S vernac B vernac +S plugins/ltac +B plugins/ltac S tools B tools diff --git a/checker/modops.ml b/checker/modops.ml index aba9da2fe..07dda8a06 100644 --- a/checker/modops.ml +++ b/checker/modops.ml @@ -32,10 +32,10 @@ let error_no_such_label_sub l l1 = error ("The field "^ Label.to_string l^" is missing in "^l1^".") -let error_not_a_module_loc loc s = - user_err ~loc (str ("\""^Label.to_string s^"\" is not a module")) +let error_not_a_module_loc ?loc s = + user_err ?loc (str ("\""^Label.to_string s^"\" is not a module")) -let error_not_a_module s = error_not_a_module_loc Loc.ghost s +let error_not_a_module s = error_not_a_module_loc s let error_with_module () = error "Unsupported 'with' constraint in module implementation" diff --git a/dev/top_printers.ml b/dev/top_printers.ml index 52cf8cf97..96cb9443c 100644 --- a/dev/top_printers.ml +++ b/dev/top_printers.ml @@ -521,7 +521,7 @@ let _ = extend_vernac_command_grammar ("PrintConstr", 0) None [GramTerminal "PrintConstr"; GramNonTerminal - (Loc.ghost,rawwit wit_constr,Extend.Aentry Pcoq.Constr.constr)] + (Loc.internal_ghost, rawwit wit_constr,Extend.Aentry Pcoq.Constr.constr)] let _ = try @@ -537,46 +537,46 @@ let _ = extend_vernac_command_grammar ("PrintPureConstr", 0) None [GramTerminal "PrintPureConstr"; GramNonTerminal - (Loc.ghost,rawwit wit_constr,Extend.Aentry Pcoq.Constr.constr)] + (Loc.internal_ghost,rawwit wit_constr,Extend.Aentry Pcoq.Constr.constr)] (* Setting printer of unbound global reference *) open Names open Libnames -let encode_path loc prefix mpdir suffix id = +let encode_path ?loc prefix mpdir suffix id = let dir = match mpdir with | None -> [] | Some (mp,dir) -> (DirPath.repr (dirpath_of_string (string_of_mp mp))@ DirPath.repr dir) in - Qualid (loc, make_qualid + Qualid (Loc.tag ?loc @@ make_qualid (DirPath.make (List.rev (Id.of_string prefix::dir@suffix))) id) -let raw_string_of_ref loc _ = function +let raw_string_of_ref ?loc _ = function | ConstRef cst -> let (mp,dir,id) = repr_con cst in - encode_path loc "CST" (Some (mp,dir)) [] (Label.to_id id) + encode_path ?loc "CST" (Some (mp,dir)) [] (Label.to_id id) | IndRef (kn,i) -> let (mp,dir,id) = repr_mind kn in - encode_path loc "IND" (Some (mp,dir)) [Label.to_id id] + encode_path ?loc "IND" (Some (mp,dir)) [Label.to_id id] (Id.of_string ("_"^string_of_int i)) | ConstructRef ((kn,i),j) -> let (mp,dir,id) = repr_mind kn in - encode_path loc "CSTR" (Some (mp,dir)) + encode_path ?loc "CSTR" (Some (mp,dir)) [Label.to_id id;Id.of_string ("_"^string_of_int i)] (Id.of_string ("_"^string_of_int j)) | VarRef id -> - encode_path loc "SECVAR" None [] id + encode_path ?loc "SECVAR" None [] id -let short_string_of_ref loc _ = function - | VarRef id -> Ident (loc,id) - | ConstRef cst -> Ident (loc,Label.to_id (pi3 (repr_con cst))) - | IndRef (kn,0) -> Ident (loc,Label.to_id (pi3 (repr_mind kn))) +let short_string_of_ref ?loc _ = function + | VarRef id -> Ident (Loc.tag ?loc id) + | ConstRef cst -> Ident (Loc.tag ?loc @@ Label.to_id (pi3 (repr_con cst))) + | IndRef (kn,0) -> Ident (Loc.tag ?loc @@ Label.to_id (pi3 (repr_mind kn))) | IndRef (kn,i) -> - encode_path loc "IND" None [Label.to_id (pi3 (repr_mind kn))] + encode_path ?loc "IND" None [Label.to_id (pi3 (repr_mind kn))] (Id.of_string ("_"^string_of_int i)) | ConstructRef ((kn,i),j) -> - encode_path loc "CSTR" None + encode_path ?loc "CSTR" None [Label.to_id (pi3 (repr_mind kn));Id.of_string ("_"^string_of_int i)] (Id.of_string ("_"^string_of_int j)) diff --git a/engine/evarutil.ml b/engine/evarutil.ml index 1624dc93e..ac64ab834 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -366,7 +366,7 @@ let push_rel_context_to_named_context env sigma typ = * Entry points to define new evars * *------------------------------------*) -let default_source = (Loc.ghost,Evar_kinds.InternalHole) +let default_source = Loc.tag @@ Evar_kinds.InternalHole let restrict_evar evd evk filter candidates = let evd = Sigma.to_evar_map evd in diff --git a/engine/evd.ml b/engine/evd.ml index 5419a10a8..9e81ccd36 100644 --- a/engine/evd.ml +++ b/engine/evd.ml @@ -151,7 +151,7 @@ let make_evar hyps ccl = { evar_hyps = hyps; evar_body = Evar_empty; evar_filter = Filter.identity; - evar_source = (Loc.ghost,Evar_kinds.InternalHole); + evar_source = Loc.tag @@ Evar_kinds.InternalHole; evar_candidates = None; evar_extra = Store.empty } @@ -704,11 +704,11 @@ let extract_all_conv_pbs evd = let loc_of_conv_pb evd (pbty,env,t1,t2) = match kind_of_term (fst (decompose_app t1)) with - | Evar (evk1,_) -> fst (evar_source evk1 evd) + | Evar (evk1,_) -> Some (fst (evar_source evk1 evd)) | _ -> match kind_of_term (fst (decompose_app t2)) with - | Evar (evk2,_) -> fst (evar_source evk2 evd) - | _ -> Loc.ghost + | Evar (evk2,_) -> Some (fst (evar_source evk2 evd)) + | _ -> None (** The following functions return the set of evars immediately contained in the object *) @@ -1086,8 +1086,8 @@ let retract_coercible_metas evd = let evar_source_of_meta mv evd = match meta_name evd mv with - | Anonymous -> (Loc.ghost,Evar_kinds.GoalEvar) - | Name id -> (Loc.ghost,Evar_kinds.VarInstance id) + | Anonymous -> Loc.tag Evar_kinds.GoalEvar + | Name id -> Loc.tag @@ Evar_kinds.VarInstance id let dependent_evar_ident ev evd = let evi = find evd ev in diff --git a/engine/evd.mli b/engine/evd.mli index 9c40c8b71..005332470 100644 --- a/engine/evd.mli +++ b/engine/evd.mli @@ -414,7 +414,7 @@ val extract_changed_conv_pbs : evar_map -> (Evar.Set.t -> evar_constraint -> bool) -> evar_map * evar_constraint list val extract_all_conv_pbs : evar_map -> evar_map * evar_constraint list -val loc_of_conv_pb : evar_map -> evar_constraint -> Loc.t +val loc_of_conv_pb : evar_map -> evar_constraint -> Loc.t option (** The following functions return the set of evars immediately contained in the object; need the term to be evar-normal otherwise diff --git a/engine/proofview.ml b/engine/proofview.ml index f054038e9..84bcecc44 100644 --- a/engine/proofview.ml +++ b/engine/proofview.ml @@ -66,7 +66,7 @@ let dependent_init = for type classes. *) let store = Evd.Store.set Evd.Store.empty typeclass_resolvable () in (* Goals don't have a source location. *) - let src = (Loc.ghost,Evar_kinds.GoalEvar) in + let src = Loc.tag @@ Evar_kinds.GoalEvar in (* Main routine *) let rec aux = function | TNil sigma -> [], { solution = sigma; comb = []; shelf = [] } diff --git a/engine/uState.ml b/engine/uState.ml index c66af02bb..c9653b6cd 100644 --- a/engine/uState.ml +++ b/engine/uState.ml @@ -266,10 +266,10 @@ let universe_context ?names ctx = try let info = Univ.LMap.find (Univ.LSet.choose left) (snd ctx.uctx_names) in - Option.default Loc.ghost info.uloc - with Not_found -> Loc.ghost + info.uloc + with Not_found -> None in - user_err ~loc ~hdr:"universe_context" + user_err ?loc ~hdr:"universe_context" ((str(CString.plural n "Universe") ++ spc () ++ Univ.LSet.pr (pr_uctx_level ctx) left ++ spc () ++ str (CString.conjugate_verb_to_be n) ++ diff --git a/grammar/argextend.mlp b/grammar/argextend.mlp index d00ee4e5d..5ec8086d0 100644 --- a/grammar/argextend.mlp +++ b/grammar/argextend.mlp @@ -9,7 +9,7 @@ open Q_util let loc = Ploc.dummy -let default_loc = <:expr< Loc.ghost >> +let default_loc = <:expr< Loc.internal_ghost >> IFDEF STRICT THEN let ploc_vala x = Ploc.VaVal x diff --git a/grammar/tacextend.mlp b/grammar/tacextend.mlp index 1dd8da12a..aa463ada5 100644 --- a/grammar/tacextend.mlp +++ b/grammar/tacextend.mlp @@ -20,8 +20,6 @@ let make_fun loc cl = let l = cl @ [default_patt loc] in MLast.ExFun (loc, ploc_vala l) (* correspond to <:expr< fun [ $list:l$ ] >> *) -let dloc = <:expr< Loc.ghost >> - let plugin_name = <:expr< __coq_plugin_name >> let mlexpr_of_ident id = @@ -75,7 +73,7 @@ let rec mlexpr_of_symbol = function let make_prod_item = function | ExtTerminal s -> <:expr< Tacentries.TacTerm $str:s$ >> | ExtNonTerminal (g, id) -> - <:expr< Tacentries.TacNonTerm $default_loc$ $mlexpr_of_symbol g$ $mlexpr_of_ident id$ >> + <:expr< Tacentries.TacNonTerm (Loc.tag ( $mlexpr_of_symbol g$ , $mlexpr_of_ident id$ ) ) >> let mlexpr_of_clause cl = mlexpr_of_list (fun (a,_,_) -> mlexpr_of_list make_prod_item a) cl @@ -114,7 +112,7 @@ let declare_tactic loc tacname ~level classification clause = match clause with (** Arguments are not passed directly to the ML tactic in the TacML node, the ML tactic retrieves its arguments in the [ist] environment instead. This is the rôle of the [lift_constr_tac_to_ml_tac] function. *) - let body = <:expr< Tacexpr.TacFun ($vars$, Tacexpr.TacML ($dloc$, $ml$, [])) >> in + let body = <:expr< Tacexpr.TacFun ($vars$, Tacexpr.TacML (Loc.tag ( $ml$ , []))) >> in let name = <:expr< Names.Id.of_string $name$ >> in declare_str_items loc [ <:str_item< do { diff --git a/ide/coqOps.ml b/ide/coqOps.ml index 222b9eed9..7825fb45e 100644 --- a/ide/coqOps.ml +++ b/ide/coqOps.ml @@ -465,20 +465,22 @@ object(self) self#attach_tooltip ~loc sentence (Printf.sprintf "%s %s %s" filepath ident ty) | Message(Error, loc, msg), Some (id,sentence) -> - let uloc = Option.default Loc.ghost loc in log_pp ?id Pp.(str "ErrorMsg" ++ msg); remove_flag sentence `PROCESSING; let rmsg = Pp.string_of_ppcmds msg in - add_flag sentence (`ERROR (uloc, rmsg)); + Option.iter (fun loc -> + add_flag sentence (`ERROR (loc, rmsg)); + ) loc; self#mark_as_needed sentence; - self#attach_tooltip sentence ?loc rmsg; + self#attach_tooltip ?loc sentence rmsg; self#position_tag_at_sentence ?loc Tags.Script.error sentence | Message(Warning, loc, msg), Some (id,sentence) -> - let uloc = Option.default Loc.ghost loc in log_pp ?id Pp.(str "WarningMsg" ++ msg); let rmsg = Pp.string_of_ppcmds msg in - add_flag sentence (`WARNING (uloc, rmsg)); - self#attach_tooltip sentence ?loc rmsg; + Option.iter (fun loc -> + add_flag sentence (`WARNING (loc, rmsg)); + ) loc; + self#attach_tooltip ?loc sentence rmsg; self#position_tag_at_sentence ?loc Tags.Script.warning sentence; messages#push Warning msg | Message(lvl, loc, msg), Some (id,sentence) -> @@ -528,14 +530,14 @@ object(self) let start, stop, phrase = self#get_sentence sentence in self#position_tag_at_iter ?loc start stop tag phrase - method private process_interp_error queue sentence loc msg tip id = + method private process_interp_error ?loc queue sentence msg tip id = Coq.bind (Coq.return ()) (function () -> let start, stop, phrase = self#get_sentence sentence in buffer#remove_tag Tags.Script.to_process ~start ~stop; self#discard_command_queue queue; pop_info (); if Stateid.equal id tip || Stateid.equal id Stateid.dummy then begin - self#position_tag_at_iter ~loc start stop Tags.Script.error phrase; + self#position_tag_at_iter ?loc start stop Tags.Script.error phrase; buffer#place_cursor ~where:stop; messages#clear; messages#push Feedback.Error msg; @@ -649,9 +651,9 @@ object(self) if Queue.is_empty queue then loop tip [] else loop tip (List.rev topstack) | Fail (id, loc, msg) -> - let loc = Option.cata Loc.make_loc Loc.ghost loc in + let loc = Option.map Loc.make_loc loc in let sentence = Doc.pop document in - self#process_interp_error queue sentence loc msg tip id in + self#process_interp_error ?loc queue sentence msg tip id in Coq.bind coq_query handle_answer in let tip = diff --git a/ide/ide_slave.ml b/ide/ide_slave.ml index bf86db08f..7f30a4acc 100644 --- a/ide/ide_slave.ml +++ b/ide/ide_slave.ml @@ -390,8 +390,8 @@ let quit = ref false (** Serializes the output of Stm.get_ast *) let print_ast id = match Stm.get_ast id with - | Some (expr, loc) -> begin - try Texmacspp.tmpp expr loc + | Some (loc, expr) -> begin + try Texmacspp.tmpp ~loc expr with e -> Xml_datatype.PCData ("ERROR " ^ Printexc.to_string e) end | None -> Xml_datatype.PCData "ERROR" diff --git a/ide/texmacspp.ml b/ide/texmacspp.ml index f86b260df..0a07a830b 100644 --- a/ide/texmacspp.ml +++ b/ide/texmacspp.ml @@ -17,12 +17,12 @@ open Extend open Libnames open Constrexpr_ops -let unlock loc = - let start, stop = Loc.unloc loc in +let unlock ?loc = + let start, stop = Option.cata Loc.unloc (0,0) loc in (string_of_int start, string_of_int stop) -let xmlWithLoc loc ename attr xml = - let start, stop = unlock loc in +let xmlWithLoc ?loc ename attr xml = + let start, stop = unlock ?loc in Element(ename, [ "begin", start; "end", stop ] @ attr, xml) let get_fst_attr_in_xml_list attr xml_list = @@ -49,32 +49,32 @@ let compare_begin_att xml1 xml2 = | (_, s1), (_, s2) when int_of_string s1 < int_of_string s2 -> -1 | _ -> 0 -let xmlBeginSection loc name = xmlWithLoc loc "beginsection" ["name", name] [] +let xmlBeginSection ?loc name = xmlWithLoc ?loc "beginsection" ["name", name] [] -let xmlEndSegment loc name = xmlWithLoc loc "endsegment" ["name", name] [] +let xmlEndSegment ?loc name = xmlWithLoc ?loc "endsegment" ["name", name] [] -let xmlThm typ name loc xml = - xmlWithLoc loc "theorem" ["type", typ; "name", name] xml +let xmlThm ?loc typ name xml = + xmlWithLoc ?loc "theorem" ["type", typ; "name", name] xml -let xmlDef typ name loc xml = - xmlWithLoc loc "definition" ["type", typ; "name", name] xml +let xmlDef ?loc typ name xml = + xmlWithLoc ?loc "definition" ["type", typ; "name", name] xml -let xmlNotation attr name loc xml = - xmlWithLoc loc "notation" (("name", name) :: attr) xml +let xmlNotation ?loc attr name xml = + xmlWithLoc ?loc "notation" (("name", name) :: attr) xml -let xmlReservedNotation attr name loc = - xmlWithLoc loc "reservednotation" (("name", name) :: attr) [] +let xmlReservedNotation ?loc attr name = + xmlWithLoc ?loc "reservednotation" (("name", name) :: attr) [] -let xmlCst name ?(attr=[]) loc = - xmlWithLoc loc "constant" (("name", name) :: attr) [] +let xmlCst ?loc ?(attr=[]) name = + xmlWithLoc ?loc "constant" (("name", name) :: attr) [] -let xmlOperator name ?(attr=[]) ?(pprules=[]) loc = - xmlWithLoc loc "operator" +let xmlOperator ?loc ?(attr=[]) ?(pprules=[]) name = + xmlWithLoc ?loc "operator" (("name", name) :: List.map (fun (a,b) -> "format"^a,b) pprules @ attr) [] -let xmlApply loc ?(attr=[]) xml = xmlWithLoc loc "apply" attr xml +let xmlApply ?loc ?(attr=[]) xml = xmlWithLoc ?loc "apply" attr xml -let xmlToken loc ?(attr=[]) xml = xmlWithLoc loc "token" attr xml +let xmlToken ?loc ?(attr=[]) xml = xmlWithLoc ?loc "token" attr xml let xmlTyped xml = Element("typed", (backstep_loc xml), xml) @@ -88,23 +88,23 @@ let xmlWith xml = Element("with", [], xml) let xmlAssign id xml = Element("assign", ["target",string_of_id id], [xml]) -let xmlInductive kind loc xml = xmlWithLoc loc "inductive" ["kind",kind] xml +let xmlInductive ?loc kind xml = xmlWithLoc ?loc "inductive" ["kind",kind] xml let xmlCoFixpoint xml = Element("cofixpoint", [], xml) let xmlFixpoint xml = Element("fixpoint", [], xml) -let xmlCheck loc xml = xmlWithLoc loc "check" [] xml +let xmlCheck ?loc xml = xmlWithLoc ?loc "check" [] xml -let xmlAssumption kind loc xml = xmlWithLoc loc "assumption" ["kind",kind] xml +let xmlAssumption ?loc kind xml = xmlWithLoc ?loc "assumption" ["kind",kind] xml -let xmlComment loc xml = xmlWithLoc loc "comment" [] xml +let xmlComment ?loc xml = xmlWithLoc ?loc "comment" [] xml -let xmlCanonicalStructure attr loc = xmlWithLoc loc "canonicalstructure" attr [] +let xmlCanonicalStructure ?loc attr = xmlWithLoc ?loc "canonicalstructure" attr [] -let xmlQed ?(attr=[]) loc = xmlWithLoc loc "qed" attr [] +let xmlQed ?loc ?(attr=[]) = xmlWithLoc ?loc "qed" attr [] -let xmlPatvar id loc = xmlWithLoc loc "patvar" ["id", id] [] +let xmlPatvar ?loc id = xmlWithLoc ?loc "patvar" ["id", id] [] let xmlReference ref = let name = Libnames.string_of_reference ref in @@ -112,38 +112,38 @@ let xmlReference ref = let b, e = string_of_int i, string_of_int j in Element("reference",["name", name; "begin", b; "end", e] ,[]) -let xmlRequire loc ?(attr=[]) xml = xmlWithLoc loc "require" attr xml -let xmlImport loc ?(attr=[]) xml = xmlWithLoc loc "import" attr xml +let xmlRequire ?loc ?(attr=[]) xml = xmlWithLoc ?loc "require" attr xml +let xmlImport ?loc ?(attr=[]) xml = xmlWithLoc ?loc "import" attr xml -let xmlAddLoadPath loc ?(attr=[]) xml = xmlWithLoc loc "addloadpath" attr xml -let xmlRemoveLoadPath loc ?(attr=[]) = xmlWithLoc loc "removeloadpath" attr -let xmlAddMLPath loc ?(attr=[]) = xmlWithLoc loc "addmlpath" attr +let xmlAddLoadPath ?loc ?(attr=[]) xml = xmlWithLoc ?loc "addloadpath" attr xml +let xmlRemoveLoadPath ?loc ?(attr=[]) = xmlWithLoc ?loc "removeloadpath" attr +let xmlAddMLPath ?loc ?(attr=[]) = xmlWithLoc ?loc "addmlpath" attr -let xmlExtend loc xml = xmlWithLoc loc "extend" [] xml +let xmlExtend ?loc xml = xmlWithLoc ?loc "extend" [] xml -let xmlScope loc action ?(attr=[]) name xml = - xmlWithLoc loc "scope" (["name",name;"action",action] @ attr) xml +let xmlScope ?loc ?(attr=[]) action name xml = + xmlWithLoc ?loc "scope" (["name",name;"action",action] @ attr) xml -let xmlProofMode loc name = xmlWithLoc loc "proofmode" ["name",name] [] +let xmlProofMode ?loc name = xmlWithLoc ?loc "proofmode" ["name",name] [] -let xmlProof loc xml = xmlWithLoc loc "proof" [] xml +let xmlProof ?loc xml = xmlWithLoc ?loc "proof" [] xml let xmlSectionSubsetDescr name ssd = Element("sectionsubsetdescr",["name",name], [PCData (Proof_using.to_string ssd)]) -let xmlDeclareMLModule loc s = - xmlWithLoc loc "declarexmlmodule" [] +let xmlDeclareMLModule ?loc s = + xmlWithLoc ?loc "declarexmlmodule" [] (List.map (fun x -> Element("path",["value",x],[])) s) (* tactics *) -let xmlLtac loc xml = xmlWithLoc loc "ltac" [] xml +let xmlLtac ?loc xml = xmlWithLoc ?loc "ltac" [] xml (* toplevel commands *) -let xmlGallina loc xml = xmlWithLoc loc "gallina" [] xml +let xmlGallina ?loc xml = xmlWithLoc ?loc "gallina" [] xml -let xmlTODO loc x = - xmlWithLoc loc "todo" [] [PCData (Pp.string_of_ppcmds (Ppvernac.pr_vernac x))] +let xmlTODO ?loc x = + xmlWithLoc ?loc "todo" [] [PCData (Pp.string_of_ppcmds (Ppvernac.pr_vernac x))] let string_of_name n = match n with @@ -189,7 +189,7 @@ match sm with | SetOnlyParsing -> ["onlyparsing", ""] | SetCompatVersion v -> ["compat", Flags.pr_version v] | SetFormat (system, (loc, s)) -> - let start, stop = unlock loc in + let start, stop = unlock ~loc in ["format-"^system, s; "begin", start; "end", stop] let string_of_assumption_kind l a many = @@ -217,7 +217,7 @@ let rec pp_bindlist bl = let names = (List.map (fun (loc, name) -> - xmlCst (string_of_name name) loc) loc_names) in + xmlCst ~loc (string_of_name name)) loc_names) in match e with | _, CHole _ -> names | _ -> names @ [pp_expr e]) @@ -237,13 +237,13 @@ and pp_local_binder lb = (* don't know what it is for now *) pp_expr ~attr:attrs value | CLocalAssum (namll, _, ce) -> let ppl = - List.map (fun (loc, nam) -> (xmlCst (string_of_name nam) loc)) namll in + List.map (fun (loc, nam) -> (xmlCst ~loc (string_of_name nam))) namll in xmlTyped (ppl @ [pp_expr ce]) | CLocalPattern _ -> assert false and pp_local_decl_expr lde = (* don't know what it is for now *) match lde with - | AssumExpr (_, ce) -> pp_expr ce + | AssumExpr (_, ce) -> pp_expr ce | DefExpr (_, ce, _) -> pp_expr ce and pp_inductive_expr ((_, ((l, id),_)), lbl, ceo, _, cl_or_rdexpr) = (* inductive_expr *) @@ -265,7 +265,7 @@ and pp_recursion_order_expr optid roe = (* don't know what it is for now *) match optid with | None -> [] | Some (loc, id) -> - let start, stop = unlock loc in + let start, stop = unlock ~loc in ["begin", start; "end", stop ; "name", Id.to_string id] in let kind, expr = match roe with @@ -276,7 +276,7 @@ and pp_recursion_order_expr optid roe = (* don't know what it is for now *) Element ("recursion_order", ["kind", kind] @ attrs, expr) and pp_fixpoint_expr (((loc, id), pl), (optid, roe), lbl, ce, ceo) = (* fixpoint_expr *) - let start, stop = unlock loc in + let start, stop = unlock ~loc in let id = Id.to_string id in [Element ("lident", ["begin", start; "end", stop ; "name", id], [])] @ (* fixpoint name *) @@ -290,7 +290,7 @@ and pp_fixpoint_expr (((loc, id), pl), (optid, roe), lbl, ce, ceo) = and pp_cofixpoint_expr (((loc, id), pl), lbl, ce, ceo) = (* cofixpoint_expr *) (* Nota: it is like fixpoint_expr without (optid, roe) * so could be merged if there is no more differences *) - let start, stop = unlock loc in + let start, stop = unlock ~loc in let id = Id.to_string id in [Element ("lident", ["begin", start; "end", stop ; "name", id], [])] @ (* cofixpoint name *) @@ -300,37 +300,37 @@ and pp_cofixpoint_expr (((loc, id), pl), lbl, ce, ceo) = (* cofixpoint_expr *) | Some ce -> [pp_expr ce] | None -> [] end -and pp_lident (loc, id) = xmlCst (Id.to_string id) loc +and pp_lident (loc, id) = xmlCst ~loc (Id.to_string id) and pp_simple_binder (idl, ce) = List.map pp_lident idl @ [pp_expr ce] and pp_cases_pattern_expr (loc, cpe) = match cpe with | CPatAlias (cpe, id) -> - xmlApply loc - (xmlOperator "alias" ~attr:["name", string_of_id id] loc :: + xmlApply ~loc + (xmlOperator ~loc ~attr:["name", string_of_id id] "alias" :: [pp_cases_pattern_expr cpe]) | CPatCstr (ref, None, cpel2) -> - xmlApply loc - (xmlOperator "reference" - ~attr:["name", Libnames.string_of_reference ref] loc :: + xmlApply ~loc + (xmlOperator ~loc "reference" + ~attr:["name", Libnames.string_of_reference ref] :: [Element ("impargs", [], []); Element ("args", [], (List.map pp_cases_pattern_expr cpel2))]) | CPatCstr (ref, Some cpel1, cpel2) -> - xmlApply loc - (xmlOperator "reference" - ~attr:["name", Libnames.string_of_reference ref] loc :: + xmlApply ~loc + (xmlOperator ~loc "reference" + ~attr:["name", Libnames.string_of_reference ref] :: [Element ("impargs", [], (List.map pp_cases_pattern_expr cpel1)); Element ("args", [], (List.map pp_cases_pattern_expr cpel2))]) | CPatAtom optr -> let attrs = match optr with | None -> [] | Some r -> ["name", Libnames.string_of_reference r] in - xmlApply loc (xmlOperator "atom" ~attr:attrs loc :: []) + xmlApply ~loc (xmlOperator ~loc "atom" ~attr:attrs :: []) | CPatOr cpel -> - xmlApply loc (xmlOperator "or" loc :: List.map pp_cases_pattern_expr cpel) + xmlApply ~loc (xmlOperator ~loc "or" :: List.map pp_cases_pattern_expr cpel) | CPatNotation (n, (subst_constr, subst_rec), cpel) -> - xmlApply loc - (xmlOperator "notation" loc :: - [xmlOperator n loc; + xmlApply ~loc + (xmlOperator ~loc "notation" :: + [xmlOperator ~loc n; Element ("subst", [], [Element ("subterms", [], List.map pp_cases_pattern_expr subst_constr); @@ -341,29 +341,29 @@ and pp_cases_pattern_expr (loc, cpe) = List.map pp_cases_pattern_expr cpel)) subst_rec)]); Element ("args", [], (List.map pp_cases_pattern_expr cpel))]) - | CPatPrim tok -> pp_token loc tok + | CPatPrim tok -> pp_token ~loc tok | CPatRecord rcl -> - xmlApply loc - (xmlOperator "record" loc :: + xmlApply ~loc + (xmlOperator ~loc "record" :: List.map (fun (r, cpe) -> Element ("field", ["reference", Libnames.string_of_reference r], [pp_cases_pattern_expr cpe])) rcl) | CPatDelimiters (delim, cpe) -> - xmlApply loc - (xmlOperator "delimiter" ~attr:["name", delim] loc :: + xmlApply ~loc + (xmlOperator ~loc "delimiter" ~attr:["name", delim] :: [pp_cases_pattern_expr cpe]) | CPatCast _ -> assert false and pp_case_expr (e, name, pat) = match name, pat with | None, None -> xmlScrutinee [pp_expr e] | Some (loc, name), None -> - let start, stop= unlock loc in + let start, stop= unlock ~loc in xmlScrutinee ~attr:["name", string_of_name name; "begin", start; "end", stop] [pp_expr e] | Some (loc, name), Some p -> - let start, stop= unlock loc in + let start, stop= unlock ~loc in xmlScrutinee ~attr:["name", string_of_name name; "begin", start; "end", stop] [Element ("in", [], [pp_cases_pattern_expr p]) ; pp_expr e] @@ -378,12 +378,12 @@ and pp_branch_expr_list bel = let ppe = [pp_expr e] in xmlCase (ppcepl @ ppe)) bel) -and pp_token loc tok = +and pp_token ?loc tok = let tokstr = match tok with | String s -> PCData s | Numeral n -> PCData (to_string n) in - xmlToken loc [tokstr] + xmlToken ?loc [tokstr] and pp_local_binder_list lbl = let l = (List.map pp_local_binder lbl) in Element ("recurse", (backstep_loc l), l) @@ -393,142 +393,140 @@ and pp_const_expr_list cel = and pp_expr ?(attr=[]) (loc, e) = match e with | CRef (r, _) -> - xmlCst ~attr - (Libnames.string_of_reference r) (Libnames.loc_of_reference r) + xmlCst ~loc:(Libnames.loc_of_reference r) ~attr (Libnames.string_of_reference r) | CProdN (bl, e) -> - xmlApply loc - (xmlOperator "forall" loc :: [pp_bindlist bl] @ [pp_expr e]) + xmlApply ~loc + (xmlOperator ~loc "forall" :: [pp_bindlist bl] @ [pp_expr e]) | CApp ((_, hd), args) -> - xmlApply ~attr loc (pp_expr hd :: List.map (fun (e,_) -> pp_expr e) args) + xmlApply ~loc ~attr (pp_expr hd :: List.map (fun (e,_) -> pp_expr e) args) | CAppExpl ((_, r, _), args) -> - xmlApply ~attr loc - (xmlCst (Libnames.string_of_reference r) - (Libnames.loc_of_reference r) :: List.map pp_expr args) + xmlApply ~loc ~attr + (xmlCst ~loc:(Libnames.loc_of_reference r) (Libnames.string_of_reference r) + :: List.map pp_expr args) | CNotation (notation, ([],[],[])) -> - xmlOperator notation loc + xmlOperator ~loc notation | CNotation (notation, (args, cell, lbll)) -> let fmts = Notation.find_notation_extra_printing_rules notation in - let oper = xmlOperator notation loc ~pprules:fmts in + let oper = xmlOperator ~loc notation ~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))) + xmlApply ~loc (oper :: (List.sort compare_begin_att (args @ cels @ lbls))) | CSort(s) -> - xmlOperator (string_of_glob_sort s) loc + xmlOperator ~loc (string_of_glob_sort s) | CDelimiters (scope, ce) -> - xmlApply loc (xmlOperator "delimiter" ~attr:["name", scope] loc :: + xmlApply ~loc (xmlOperator ~loc "delimiter" ~attr:["name", scope] :: [pp_expr ce]) - | CPrim tok -> pp_token loc tok + | 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]) + xmlApply ~loc + (xmlOperator ~loc ~attr:["kind", kind] "generalization" :: [pp_expr e]) | CCast (e, tc) -> begin match tc with | CastConv t | CastVM t |CastNative t -> - xmlApply loc - (xmlOperator ":" loc ~attr:["kind", (string_of_cast_sort tc)] :: + xmlApply ~loc + (xmlOperator ~loc ":" ~attr:["kind", (string_of_cast_sort tc)] :: [pp_expr e; pp_expr t]) | CastCoerce -> - xmlApply loc - (xmlOperator ":" loc ~attr:["kind", "CastCoerce"] :: + xmlApply ~loc + (xmlOperator ~loc ":" ~attr:["kind", "CastCoerce"] :: [pp_expr e]) end | 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] :: + xmlApply ~loc + (xmlOperator ~loc "evar" ~attr:["id", string_of_id ek] :: ppcel) - | CPatVar id -> xmlPatvar (string_of_id id) loc - | CHole (_, _, _) -> xmlCst ~attr "_" loc + | CPatVar id -> xmlPatvar ~loc (string_of_id id) + | CHole (_, _, _) -> xmlCst ~loc ~attr "_" | CIf (test, (_, ret), th, el) -> let return = match ret with | None -> [] | Some r -> [xmlReturn [pp_expr r]] in - xmlApply loc - (xmlOperator "if" loc :: + xmlApply ~loc + (xmlOperator ~loc "if" :: return @ [pp_expr th] @ [pp_expr el]) | CLetTuple (names, (_, ret), value, body) -> let return = match ret with | None -> [] | Some r -> [xmlReturn [pp_expr r]] in - xmlApply loc - (xmlOperator "lettuple" loc :: + xmlApply ~loc + (xmlOperator ~loc "lettuple" :: return @ - (List.map (fun (loc, var) -> xmlCst (string_of_name var) loc) names) @ + (List.map (fun (loc, var) -> xmlCst ~loc (string_of_name var)) names) @ [pp_expr value; pp_expr body]) | CCases (sty, ret, cel, bel) -> let return = match ret with | None -> [] | Some r -> [xmlReturn [pp_expr r]] in - xmlApply loc - (xmlOperator "match" loc ~attr:["style", (string_of_case_style sty)] :: + xmlApply ~loc + (xmlOperator ~loc ~attr:["style", (string_of_case_style sty)] "match" :: (return @ [Element ("scrutinees", [], List.map pp_case_expr cel)] @ [pp_branch_expr_list bel])) | 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 (Loc.tag ~loc value); pp_expr body]) + xmlApply ~loc + (xmlOperator ~loc "let" :: + [xmlCst ~loc:varloc (string_of_name var) ; pp_expr (Loc.tag ~loc value); pp_expr body]) | CLambdaN (bl, e) -> - xmlApply loc - (xmlOperator "lambda" loc :: [pp_bindlist bl] @ [pp_expr e]) + xmlApply ~loc + (xmlOperator ~loc "lambda" :: [pp_bindlist bl] @ [pp_expr e]) | CCoFix (_, _) -> assert false | CFix (lid, fel) -> - xmlApply loc - (xmlOperator "fix" loc :: + xmlApply ~loc + (xmlOperator ~loc "fix" :: List.flatten (List.map (fun (a,b,cl,c,d) -> pp_fixpoint_expr ((a,None),b,cl,c,Some d)) fel)) -let pp_comment (c) = +let pp_comment c = match c with | CommentConstr e -> [pp_expr e] | CommentString s -> [Element ("string", [], [PCData s])] - | CommentInt i -> [PCData (string_of_int i)] + | CommentInt i -> [PCData (string_of_int i)] -let rec tmpp v loc = +let rec tmpp ?loc v = match v with (* Control *) | VernacLoad (verbose,f) -> - xmlWithLoc loc "load" ["verbose",string_of_bool verbose;"file",f] [] + xmlWithLoc ?loc "load" ["verbose",string_of_bool verbose;"file",f] [] | VernacTime (loc,e) -> - xmlApply loc (Element("time",[],[]) :: - [tmpp e loc]) + xmlApply ~loc (Element("time",[],[]) :: + [tmpp ~loc e]) | VernacRedirect (s, (loc,e)) -> - xmlApply loc (Element("redirect",["path", s],[]) :: - [tmpp e loc]) + xmlApply ~loc (Element("redirect",["path", s],[]) :: + [tmpp ~loc e]) | VernacTimeout (s,e) -> - xmlApply loc (Element("timeout",["val",string_of_int s],[]) :: - [tmpp e loc]) - | VernacFail e -> xmlApply loc (Element("fail",[],[]) :: [tmpp e loc]) + xmlApply ?loc (Element("timeout",["val",string_of_int s],[]) :: + [tmpp ?loc e]) + | VernacFail e -> xmlApply ?loc (Element("fail",[],[]) :: [tmpp ?loc e]) (* Syntax *) | VernacSyntaxExtension (_, ((_, name), sml)) -> let attrs = List.flatten (List.map attribute_of_syntax_modifier sml) in - xmlReservedNotation attrs name loc + xmlReservedNotation ?loc attrs name - | VernacOpenCloseScope (_,(true,name)) -> xmlScope loc "open" name [] - | VernacOpenCloseScope (_,(false,name)) -> xmlScope loc "close" name [] + | VernacOpenCloseScope (_,(true,name)) -> xmlScope ?loc "open" name [] + | VernacOpenCloseScope (_,(false,name)) -> xmlScope ?loc "close" name [] | VernacDelimiters (name,Some tag) -> - xmlScope loc "delimit" name ~attr:["delimiter",tag] [] + xmlScope ?loc "delimit" name ~attr:["delimiter",tag] [] | VernacDelimiters (name,None) -> - xmlScope loc "undelimit" name ~attr:[] [] + xmlScope ?loc "undelimit" name ~attr:[] [] | VernacInfix (_,((_,name),sml),ce,sn) -> let attrs = List.flatten (List.map attribute_of_syntax_modifier sml) in let sc_attr = match sn with | Some scope -> ["scope", scope] | None -> [] in - xmlNotation (sc_attr @ attrs) name loc [pp_expr ce] + xmlNotation ?loc (sc_attr @ attrs) name [pp_expr ce] | VernacNotation (_, ce, (lstr, sml), sn) -> let name = snd lstr in let attrs = List.flatten (List.map attribute_of_syntax_modifier sml) in @@ -536,12 +534,12 @@ let rec tmpp v loc = match sn with | Some scope -> ["scope", scope] | None -> [] in - xmlNotation (sc_attr @ attrs) name loc [pp_expr ce] - | VernacBindScope _ as x -> xmlTODO loc x - | VernacNotationAddFormat _ as x -> xmlTODO loc x + xmlNotation ?loc (sc_attr @ attrs) name [pp_expr ce] + | VernacBindScope _ as x -> xmlTODO ?loc x + | VernacNotationAddFormat _ as x -> xmlTODO ?loc x | VernacUniverse _ | VernacConstraint _ - | VernacPolymorphic (_, _) as x -> xmlTODO loc x + | VernacPolymorphic (_, _) as x -> xmlTODO ?loc x (* Gallina *) | VernacDefinition (ldk, ((_,id),_), de) -> let l, dk = @@ -557,26 +555,26 @@ let rec tmpp v loc = | DefineBody (_, None , ce, Some _) -> ce in let str_dk = Kindops.string_of_definition_kind (l, false, dk) in let str_id = Id.to_string id in - (xmlDef str_dk str_id loc [pp_expr e]) + (xmlDef ?loc str_dk str_id [pp_expr e]) | VernacStartTheoremProof (tk, [ Some ((_,id),_), ([], statement, None) ], b) -> let str_tk = Kindops.string_of_theorem_kind tk in let str_id = Id.to_string id in - (xmlThm str_tk str_id loc [pp_expr statement]) - | VernacStartTheoremProof _ as x -> xmlTODO loc x + (xmlThm ?loc str_tk str_id [pp_expr statement]) + | VernacStartTheoremProof _ as x -> xmlTODO ?loc x | VernacEndProof pe -> begin match pe with - | Admitted -> xmlQed loc + | Admitted -> xmlQed ?loc ?attr:None | Proved (_, Some ((_, id), Some tk)) -> let nam = Id.to_string id in let typ = Kindops.string_of_theorem_kind tk in - xmlQed ~attr:["name", nam; "type", typ] loc + xmlQed ?loc ~attr:["name", nam; "type", typ] | Proved (_, Some ((_, id), None)) -> let nam = Id.to_string id in - xmlQed ~attr:["name", nam] loc - | Proved _ -> xmlQed loc + xmlQed ?loc ~attr:["name", nam] + | Proved _ -> xmlQed ?loc ?attr:None end - | VernacExactProof _ as x -> xmlTODO loc x + | VernacExactProof _ as x -> xmlTODO ?loc x | VernacAssumption ((l, a), _, sbwcl) -> let binders = List.map (fun (_, (id, c)) -> (List.map fst id, c)) sbwcl in let many = @@ -585,7 +583,7 @@ let rec tmpp v loc = List.flatten (List.map pp_simple_binder binders) in let l = match l with Some x -> x | None -> Decl_kinds.Global in let kind = string_of_assumption_kind l a many in - xmlAssumption kind loc exprs + xmlAssumption ?loc kind exprs | VernacInductive (_, _, iednll) -> let kind = let (_, _, _, k, _),_ = List.hd iednll in @@ -603,7 +601,7 @@ let rec tmpp v loc = (List.map (fun (ie, dnl) -> (pp_inductive_expr ie) @ (List.map pp_decl_notation dnl)) iednll) in - xmlInductive kind loc exprs + xmlInductive ?loc kind exprs | VernacFixpoint (_, fednll) -> let exprs = List.flatten (* should probably not be flattened *) @@ -619,13 +617,13 @@ let rec tmpp v loc = (fun (cfe, dnl) -> (pp_cofixpoint_expr cfe) @ (List.map pp_decl_notation dnl)) cfednll) in xmlCoFixpoint exprs - | VernacScheme _ as x -> xmlTODO loc x - | VernacCombinedScheme _ as x -> xmlTODO loc x + | VernacScheme _ as x -> xmlTODO ?loc x + | VernacCombinedScheme _ as x -> xmlTODO ?loc x (* Gallina extensions *) - | VernacBeginSection (_, id) -> xmlBeginSection loc (Id.to_string id) - | VernacEndSegment (_, id) -> xmlEndSegment loc (Id.to_string id) - | VernacNameSectionHypSet _ as x -> xmlTODO loc x + | VernacBeginSection (_, id) -> xmlBeginSection ?loc (Id.to_string id) + | VernacEndSegment (_, id) -> xmlEndSegment ?loc (Id.to_string id) + | VernacNameSectionHypSet _ as x -> xmlTODO ?loc x | VernacRequire (from, import, l) -> let import = match import with | None -> [] @@ -636,137 +634,136 @@ let rec tmpp v loc = | None -> [] | Some r -> ["from", Libnames.string_of_reference r] in - xmlRequire loc ~attr:(from @ import) (List.map (fun ref -> + xmlRequire ?loc ~attr:(from @ import) (List.map (fun ref -> xmlReference ref) l) | VernacImport (true,l) -> - xmlImport loc ~attr:["export","true"] (List.map (fun ref -> + xmlImport ?loc ~attr:["export","true"] (List.map (fun ref -> xmlReference ref) l) | VernacImport (false,l) -> - xmlImport loc (List.map (fun ref -> - xmlReference ref) l) + xmlImport ?loc (List.map (fun ref -> xmlReference ref) l) | VernacCanonical r -> let attr = match r with | AN (Qualid (_, q)) -> ["qualid", string_of_qualid q] | AN (Ident (_, id)) -> ["id", Id.to_string id] | ByNotation (_, (s, _)) -> ["notation", s] in - xmlCanonicalStructure attr loc - | VernacCoercion _ as x -> xmlTODO loc x - | VernacIdentityCoercion _ as x -> xmlTODO loc x + xmlCanonicalStructure ?loc attr + | VernacCoercion _ as x -> xmlTODO ?loc x + | VernacIdentityCoercion _ as x -> xmlTODO ?loc x (* Type classes *) - | VernacInstance _ as x -> xmlTODO loc x + | VernacInstance _ as x -> xmlTODO ?loc x - | VernacContext _ as x -> xmlTODO loc x + | VernacContext _ as x -> xmlTODO ?loc x - | VernacDeclareInstances _ as x -> xmlTODO loc x + | VernacDeclareInstances _ as x -> xmlTODO ?loc x - | VernacDeclareClass _ as x -> xmlTODO loc x + | VernacDeclareClass _ as x -> xmlTODO ?loc x (* Modules and Module Types *) - | VernacDeclareModule _ as x -> xmlTODO loc x - | VernacDefineModule _ as x -> xmlTODO loc x - | VernacDeclareModuleType _ as x -> xmlTODO loc x - | VernacInclude _ as x -> xmlTODO loc x + | VernacDeclareModule _ as x -> xmlTODO ?loc x + | VernacDefineModule _ as x -> xmlTODO ?loc x + | VernacDeclareModuleType _ as x -> xmlTODO ?loc x + | VernacInclude _ as x -> xmlTODO ?loc x (* Solving *) | (VernacSolveExistential _) as x -> - xmlLtac loc [PCData (Pp.string_of_ppcmds (Ppvernac.pr_vernac x))] + xmlLtac ?loc [PCData (Pp.string_of_ppcmds (Ppvernac.pr_vernac x))] (* Auxiliary file and library management *) | VernacAddLoadPath (recf,name,None) -> - xmlAddLoadPath loc ~attr:["rec",string_of_bool recf;"path",name] [] + xmlAddLoadPath ?loc ~attr:["rec",string_of_bool recf;"path",name] [] | VernacAddLoadPath (recf,name,Some dp) -> - xmlAddLoadPath loc ~attr:["rec",string_of_bool recf;"path",name] + xmlAddLoadPath ?loc ~attr:["rec",string_of_bool recf;"path",name] [PCData (Names.DirPath.to_string dp)] - | VernacRemoveLoadPath name -> xmlRemoveLoadPath loc ~attr:["path",name] [] + | VernacRemoveLoadPath name -> xmlRemoveLoadPath ?loc ~attr:["path",name] [] | VernacAddMLPath (recf,name) -> - xmlAddMLPath loc ~attr:["rec",string_of_bool recf;"path",name] [] - | VernacDeclareMLModule sl -> xmlDeclareMLModule loc sl - | VernacChdir _ as x -> xmlTODO loc x + xmlAddMLPath ?loc ~attr:["rec",string_of_bool recf;"path",name] [] + | VernacDeclareMLModule sl -> xmlDeclareMLModule ?loc sl + | VernacChdir _ as x -> xmlTODO ?loc x (* State management *) - | VernacWriteState _ as x -> xmlTODO loc x - | VernacRestoreState _ as x -> xmlTODO loc x + | VernacWriteState _ as x -> xmlTODO ?loc x + | VernacRestoreState _ as x -> xmlTODO ?loc x (* Resetting *) - | VernacResetName _ as x -> xmlTODO loc x - | VernacResetInitial as x -> xmlTODO loc x - | VernacBack _ as x -> xmlTODO loc x + | VernacResetName _ as x -> xmlTODO ?loc x + | VernacResetInitial as x -> xmlTODO ?loc x + | VernacBack _ as x -> xmlTODO ?loc x | VernacBackTo _ -> PCData "VernacBackTo" (* Commands *) - | VernacCreateHintDb _ as x -> xmlTODO loc x - | VernacRemoveHints _ as x -> xmlTODO loc x - | VernacHints _ as x -> xmlTODO loc x + | VernacCreateHintDb _ as x -> xmlTODO ?loc x + | VernacRemoveHints _ as x -> xmlTODO ?loc x + | VernacHints _ as x -> xmlTODO ?loc x | VernacSyntacticDefinition ((_, name), (idl, ce), _, _) -> let name = Id.to_string name in let attrs = List.map (fun id -> ("id", Id.to_string id)) idl in - xmlNotation attrs name loc [pp_expr ce] - | VernacDeclareImplicits _ as x -> xmlTODO loc x - | VernacArguments _ as x -> xmlTODO loc x - | VernacArgumentsScope _ as x -> xmlTODO loc x - | VernacReserve _ as x -> xmlTODO loc x - | VernacGeneralizable _ as x -> xmlTODO loc x - | VernacSetOpacity _ as x -> xmlTODO loc x - | VernacSetStrategy _ as x -> xmlTODO loc x - | VernacUnsetOption _ as x -> xmlTODO loc x - | VernacSetOption _ as x -> xmlTODO loc x - | VernacSetAppendOption _ as x -> xmlTODO loc x - | VernacAddOption _ as x -> xmlTODO loc x - | VernacRemoveOption _ as x -> xmlTODO loc x - | VernacMemOption _ as x -> xmlTODO loc x - | VernacPrintOption _ as x -> xmlTODO loc x - | VernacCheckMayEval (_,_,e) -> xmlCheck loc [pp_expr e] - | VernacGlobalCheck _ as x -> xmlTODO loc x - | VernacDeclareReduction _ as x -> xmlTODO loc x - | VernacPrint _ as x -> xmlTODO loc x - | VernacSearch _ as x -> xmlTODO loc x - | VernacLocate _ as x -> xmlTODO loc x - | VernacRegister _ as x -> xmlTODO loc x + xmlNotation ?loc attrs name [pp_expr ce] + | VernacDeclareImplicits _ as x -> xmlTODO ?loc x + | VernacArguments _ as x -> xmlTODO ?loc x + | VernacArgumentsScope _ as x -> xmlTODO ?loc x + | VernacReserve _ as x -> xmlTODO ?loc x + | VernacGeneralizable _ as x -> xmlTODO ?loc x + | VernacSetOpacity _ as x -> xmlTODO ?loc x + | VernacSetStrategy _ as x -> xmlTODO ?loc x + | VernacUnsetOption _ as x -> xmlTODO ?loc x + | VernacSetOption _ as x -> xmlTODO ?loc x + | VernacSetAppendOption _ as x -> xmlTODO ?loc x + | VernacAddOption _ as x -> xmlTODO ?loc x + | VernacRemoveOption _ as x -> xmlTODO ?loc x + | VernacMemOption _ as x -> xmlTODO ?loc x + | VernacPrintOption _ as x -> xmlTODO ?loc x + | VernacCheckMayEval (_,_,e) -> xmlCheck ?loc [pp_expr e] + | VernacGlobalCheck _ as x -> xmlTODO ?loc x + | VernacDeclareReduction _ as x -> xmlTODO ?loc x + | VernacPrint _ as x -> xmlTODO ?loc x + | VernacSearch _ as x -> xmlTODO ?loc x + | VernacLocate _ as x -> xmlTODO ?loc x + | VernacRegister _ as x -> xmlTODO ?loc x | VernacComments (cl) -> - xmlComment loc (List.flatten (List.map pp_comment cl)) + xmlComment ?loc (List.flatten (List.map pp_comment cl)) (* Stm backdoor *) - | VernacStm _ as x -> xmlTODO loc x + | VernacStm _ as x -> xmlTODO ?loc x (* Proof management *) - | VernacGoal _ as x -> xmlTODO loc x - | VernacAbort _ as x -> xmlTODO loc x + | VernacGoal _ as x -> xmlTODO ?loc x + | VernacAbort _ as x -> xmlTODO ?loc x | VernacAbortAll -> PCData "VernacAbortAll" - | VernacRestart as x -> xmlTODO loc x - | VernacUndo _ as x -> xmlTODO loc x - | VernacUndoTo _ as x -> xmlTODO loc x - | VernacBacktrack _ as x -> xmlTODO loc x - | VernacFocus _ as x -> xmlTODO loc x - | VernacUnfocus as x -> xmlTODO loc x - | VernacUnfocused as x -> xmlTODO loc x - | VernacBullet _ as x -> xmlTODO loc x - | VernacSubproof _ as x -> xmlTODO loc x - | VernacEndSubproof as x -> xmlTODO loc x - | VernacShow _ as x -> xmlTODO loc x - | VernacCheckGuard as x -> xmlTODO loc x + | VernacRestart as x -> xmlTODO ?loc x + | VernacUndo _ as x -> xmlTODO ?loc x + | VernacUndoTo _ as x -> xmlTODO ?loc x + | VernacBacktrack _ as x -> xmlTODO ?loc x + | VernacFocus _ as x -> xmlTODO ?loc x + | VernacUnfocus as x -> xmlTODO ?loc x + | VernacUnfocused as x -> xmlTODO ?loc x + | VernacBullet _ as x -> xmlTODO ?loc x + | VernacSubproof _ as x -> xmlTODO ?loc x + | VernacEndSubproof as x -> xmlTODO ?loc x + | VernacShow _ as x -> xmlTODO ?loc x + | VernacCheckGuard as x -> xmlTODO ?loc x | VernacProof (tac,using) -> let tac = None (** FIXME *) in let using = Option.map (xmlSectionSubsetDescr "using") using in - xmlProof loc (Option.List.(cons tac (cons using []))) - | VernacProofMode name -> xmlProofMode loc name + xmlProof ?loc (Option.List.(cons tac (cons using []))) + | VernacProofMode name -> xmlProofMode ?loc name (* Toplevel control *) - | VernacToplevelControl _ as x -> xmlTODO loc x + | VernacToplevelControl _ as x -> xmlTODO ?loc x (* For extension *) | VernacExtend _ as x -> - xmlExtend loc [PCData (Pp.string_of_ppcmds (Ppvernac.pr_vernac x))] + xmlExtend ?loc [PCData (Pp.string_of_ppcmds (Ppvernac.pr_vernac x))] (* Flags *) - | VernacProgram e -> xmlApply loc (Element("program",[],[]) :: [tmpp e loc]) + | VernacProgram e -> xmlApply ?loc (Element("program",[],[]) :: [tmpp ?loc e]) | VernacLocal (b,e) -> - xmlApply loc (Element("local",["flag",string_of_bool b],[]) :: - [tmpp e loc]) + xmlApply ?loc (Element("local",["flag",string_of_bool b],[]) :: + [tmpp ?loc e]) -let tmpp v loc = - match tmpp v loc with +let tmpp ?loc v = + match tmpp ?loc v with | Element("ltac",_,_) as x -> x - | xml -> xmlGallina loc [xml] + | xml -> xmlGallina ?loc [xml] diff --git a/ide/texmacspp.mli b/ide/texmacspp.mli index 858847fb6..c1086a633 100644 --- a/ide/texmacspp.mli +++ b/ide/texmacspp.mli @@ -9,4 +9,4 @@ open Xml_datatype open Vernacexpr -val tmpp : vernac_expr -> Loc.t -> xml +val tmpp : ?loc:Loc.t -> vernac_expr -> xml diff --git a/ide/xmlprotocol.ml b/ide/xmlprotocol.ml index bf52b0b52..53eb1a95f 100644 --- a/ide/xmlprotocol.ml +++ b/ide/xmlprotocol.ml @@ -863,7 +863,7 @@ let to_feedback_content = do_match "feedback_content" (fun s a -> match s,a with | "workerstatus", [ns] -> let n, s = to_pair to_string to_string ns in WorkerStatus(n,s) - | "custom", [loc;name;x]-> Custom (to_loc loc, to_string name, x) + | "custom", [loc;name;x]-> Custom (to_option to_loc loc, to_string name, x) | "filedependency", [from; dep] -> FileDependency (to_option to_string from, to_string dep) | "fileloaded", [dirpath; filename] -> @@ -896,7 +896,7 @@ let of_feedback_content = function constructor "feedback_content" "workerstatus" [of_pair of_string of_string (n,s)] | Custom (loc, name, x) -> - constructor "feedback_content" "custom" [of_loc loc; of_string name; x] + constructor "feedback_content" "custom" [of_option of_loc loc; of_string name; x] | FileDependency (from, depends_on) -> constructor "feedback_content" "filedependency" [ of_option of_string from; diff --git a/interp/constrexpr_ops.ml b/interp/constrexpr_ops.ml index 61115c00b..4b61ab494 100644 --- a/interp/constrexpr_ops.ml +++ b/interp/constrexpr_ops.ml @@ -242,13 +242,12 @@ let local_binder_loc = function | CLocalPattern (loc,_) -> loc let local_binders_loc bll = match bll with - | [] -> Loc.ghost - | h :: l -> - Loc.merge (local_binder_loc h) (local_binder_loc (List.last bll)) + | [] -> None + | h :: l -> Some (Loc.merge (local_binder_loc h) (local_binder_loc (List.last bll))) (** Pseudo-constructors *) -let mkIdentC id = Loc.tag @@ CRef (Ident (Loc.ghost, id),None) +let mkIdentC id = Loc.tag @@ CRef (Ident (Loc.tag 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) @@ -268,23 +267,23 @@ let add_name_in_env env n = let (fresh_var, fresh_var_hook) = Hook.make ~default:(fun _ _ -> assert false) () -let expand_binders ~loc mkC 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:(Loc.merge loc1 loc) bl c in + let env, c = loop ~loc:(Loc.opt_merge loc1 loc) bl c in let env = add_name_in_env env n in - (env, Loc.tag ~loc @@ CLetIn (n,oty,b,c)) + (env, Loc.tag ?loc @@ CLetIn (n,oty,b,c)) | CLocalAssum ((loc1,_)::_ as nl, bk, t) -> - let env, c = loop ~loc:(Loc.merge loc1 loc) bl c in + let env, c = loop ~loc:(Loc.opt_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) - | CLocalAssum ([],_,_) -> loop loc bl c + (env, mkC ?loc (nl,bk,t) c) + | CLocalAssum ([],_,_) -> loop ?loc bl c | CLocalPattern (loc1, (p, ty)) -> - let env, c = loop ~loc:(Loc.merge loc1 loc) bl c in + let env, c = loop ~loc:(Loc.opt_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 @@ -292,27 +291,27 @@ let expand_binders ~loc mkC bl c = | None -> Loc.tag ~loc:loc1 @@ CHole (None, IntroAnonymous, None) in let e = Loc.tag @@ CRef (Libnames.Ident (loc1, ni), None) in - let c = Loc.tag ~loc @@ + let c = Loc.tag ?loc @@ CCases (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 + let (_, c) = loop ?loc bl c in 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 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 = Loc.tag ~loc @@ CLambdaN ([b],c) in - expand_binders ~loc mk 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 -let prod_constr_expr c bl = mkCProdN (local_binders_loc bl) bl c +let abstract_constr_expr c bl = mkCLambdaN ?loc:(local_binders_loc bl) bl c +let prod_constr_expr c bl = mkCProdN ?loc:(local_binders_loc bl) bl c let coerce_reference_to_id = function | Ident (_,id) -> id diff --git a/interp/constrexpr_ops.mli b/interp/constrexpr_ops.mli index ae5ec2be5..82e4f54b0 100644 --- a/interp/constrexpr_ops.mli +++ b/interp/constrexpr_ops.mli @@ -37,7 +37,7 @@ val binder_kind_eq : binder_kind -> binder_kind -> bool val constr_loc : constr_expr -> Loc.t val cases_pattern_expr_loc : cases_pattern_expr -> Loc.t val raw_cases_pattern_expr_loc : raw_cases_pattern_expr -> Loc.t -val local_binders_loc : local_binder_expr list -> Loc.t +val local_binders_loc : local_binder_expr list -> Loc.t option (** {6 Constructors}*) @@ -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: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: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 8d9f8552d..5960a6baa 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -146,13 +146,13 @@ let insert_delimiters e = function | None -> e | Some sc -> Loc.tag @@ CDelimiters (sc,e) -let insert_pat_delimiters loc p = function +let insert_pat_delimiters ?loc p = function | None -> p - | Some sc -> Loc.tag ~loc @@ CPatDelimiters (sc,p) + | Some sc -> Loc.tag ?loc @@ CPatDelimiters (sc,p) -let insert_pat_alias loc p = function +let insert_pat_alias ?loc p = function | Anonymous -> p - | Name id -> Loc.tag ~loc @@ CPatAlias (p,id) + | Name id -> Loc.tag ?loc @@ CPatAlias (p,id) (**********************************************************************) (* conversion of references *) @@ -163,15 +163,15 @@ let extern_evar n l = CEvar (n,l) For instance, in the debugger the tables of global references may be inaccurate *) -let default_extern_reference loc vars r = - Qualid (loc,shortest_qualid_of_global vars r) +let default_extern_reference ?loc vars r = + Qualid (Loc.tag ?loc @@ shortest_qualid_of_global vars r) let my_extern_reference = ref default_extern_reference let set_extern_reference f = my_extern_reference := f let get_extern_reference () = !my_extern_reference -let extern_reference loc vars l = !my_extern_reference loc vars l +let extern_reference ?loc vars l = !my_extern_reference ?loc vars l (**********************************************************************) (* mapping patterns to cases_pattern_expr *) @@ -266,16 +266,16 @@ let make_notation loc ntn (terms,termlists,binders as subst) = (fun (loc,p) -> Loc.tag ~loc @@ CPrim p) destPrim terms -let make_pat_notation loc ntn (terms,termlists as subst) args = - if not (List.is_empty termlists) then (loc, CPatNotation (ntn,subst,args)) else +let make_pat_notation ?loc ntn (terms,termlists as subst) args = + if not (List.is_empty termlists) then (Loc.tag ?loc @@ CPatNotation (ntn,subst,args)) else make_notation_gen loc ntn - (fun (loc,ntn,l) -> Loc.tag ~loc @@ CPatNotation (ntn,(l,[]),args)) - (fun (loc,p) -> Loc.tag ~loc @@ CPatPrim p) + (fun (loc,ntn,l) -> Loc.tag ?loc @@ CPatNotation (ntn,(l,[]),args)) + (fun (loc,p) -> Loc.tag ?loc @@ CPatPrim p) destPatPrim terms -let mkPat loc qid l = +let mkPat ?loc qid l = Loc.tag ?loc @@ (* Normally irrelevant test with v8 syntax, but let's do it anyway *) - if List.is_empty l then Loc.tag ~loc @@ CPatAtom (Some qid) else Loc.tag ~loc @@ CPatCstr (qid,None,l) + if List.is_empty l then CPatAtom (Some qid) else CPatCstr (qid,None,l) let pattern_printable_in_both_syntax (ind,_ as c) = let impl_st = extract_impargs_data (implicits_of_global (ConstructRef c)) in @@ -293,7 +293,7 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = match pat with | loc, PatCstr(cstrsp,args,na) when !Flags.in_debugger||Inductiveops.constructor_has_local_defs cstrsp -> - let c = extern_reference loc Id.Set.empty (ConstructRef cstrsp) in + let c = extern_reference ~loc Id.Set.empty (ConstructRef cstrsp) in let args = List.map (extern_cases_pattern_in_scope scopes vars) args in Loc.tag ~loc @@ CPatCstr (c, Some (add_patt_for_params (fst cstrsp) args), []) | _ -> @@ -304,7 +304,7 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = | None -> raise No_match | Some key -> let loc = cases_pattern_loc pat in - insert_pat_alias loc (insert_pat_delimiters loc (Loc.tag ~loc @@ CPatPrim p) key) na + insert_pat_alias ~loc (insert_pat_delimiters ~loc (Loc.tag ~loc @@ CPatPrim p) key) na with No_match -> try if !Flags.raw_print || !print_no_symbol then raise No_match; @@ -330,12 +330,12 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = | (_loc, CPatAtom(None)) :: tail -> ip q tail acc (* we don't want to have 'x = _' in our patterns *) | head :: tail -> ip q tail - ((extern_reference loc Id.Set.empty (ConstRef c), head) :: acc) + ((extern_reference ~loc Id.Set.empty (ConstRef c), head) :: acc) in Loc.tag ~loc @@ CPatRecord(List.rev (ip projs args [])) with Not_found | No_match | Exit -> - let c = extern_reference loc Id.Set.empty (ConstructRef cstrsp) in + let c = extern_reference ~loc Id.Set.empty (ConstructRef cstrsp) in if !Topconstr.asymmetric_patterns then if pattern_printable_in_both_syntax cstrsp then Loc.tag ~loc @@ CPatCstr (c, None, args) @@ -345,8 +345,8 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = match drop_implicits_in_patt (ConstructRef cstrsp) 0 full_args with | Some true_args -> Loc.tag ~loc @@ CPatCstr (c, None, true_args) | None -> Loc.tag ~loc @@ CPatCstr (c, Some full_args, []) - in insert_pat_alias loc p na -and apply_notation_to_pattern loc gr ((subst,substlist),(nb_to_drop,more_args)) + in insert_pat_alias ~loc p na +and apply_notation_to_pattern ?loc gr ((subst,substlist),(nb_to_drop,more_args)) (tmp_scope, scopes as allscopes) vars = function | NotationRule (sc,ntn) -> @@ -373,11 +373,11 @@ and apply_notation_to_pattern loc gr ((subst,substlist),(nb_to_drop,more_args)) |Some true_args -> true_args |None -> raise No_match in - insert_pat_delimiters loc - (make_pat_notation loc ntn (l,ll) l2') key + insert_pat_delimiters ?loc + (make_pat_notation ?loc ntn (l,ll) l2') key end | SynDefRule kn -> - let qid = Qualid (loc, shortest_qualid_of_syndef vars kn) in + let qid = Qualid (Loc.tag ?loc @@ shortest_qualid_of_syndef vars kn) in let l1 = List.rev_map (fun (c,(scopt,scl)) -> extern_cases_pattern_in_scope (scopt,scl@scopes) vars c) @@ -390,7 +390,7 @@ and apply_notation_to_pattern loc gr ((subst,substlist),(nb_to_drop,more_args)) |None -> raise No_match in assert (List.is_empty substlist); - mkPat loc qid (List.rev_append l1 l2') + mkPat ?loc qid (List.rev_append l1 l2') and extern_notation_pattern (tmp_scope,scopes as allscopes) vars (loc, t) = function | [] -> raise No_match | (keyrule,pat,n as _rule)::rules -> @@ -398,9 +398,9 @@ and extern_notation_pattern (tmp_scope,scopes as allscopes) vars (loc, t) = func if List.mem keyrule !print_non_active_notations then raise No_match; match t with | PatCstr (cstr,_,na) -> - let p = apply_notation_to_pattern loc (ConstructRef cstr) + let p = apply_notation_to_pattern ~loc (ConstructRef cstr) (match_notation_constr_cases_pattern (loc, t) pat) allscopes vars keyrule in - insert_pat_alias loc p na + insert_pat_alias ~loc p na | PatVar Anonymous -> Loc.tag ~loc @@ CPatAtom None | PatVar (Name id) -> Loc.tag ~loc @@ CPatAtom (Some (Ident (loc,id))) with @@ -411,7 +411,7 @@ let rec extern_notation_ind_pattern allscopes vars ind args = function | (keyrule,pat,n as _rule)::rules -> try if List.mem keyrule !print_non_active_notations then raise No_match; - apply_notation_to_pattern Loc.ghost (IndRef ind) + apply_notation_to_pattern (IndRef ind) (match_notation_constr_ind_pattern ind args pat) allscopes vars keyrule with No_match -> extern_notation_ind_pattern allscopes vars ind args rules @@ -420,7 +420,7 @@ let extern_ind_pattern_in_scope (scopes:local_scopes) vars ind args = (* pboutill: There are letins in pat which is incompatible with notations and not explicit application. *) if !Flags.in_debugger||Inductiveops.inductive_has_local_defs ind then - let c = extern_reference Loc.ghost vars (IndRef ind) in + let c = extern_reference vars (IndRef ind) in let args = List.map (extern_cases_pattern_in_scope scopes vars) args in Loc.tag @@ CPatCstr (c, Some (add_patt_for_params ind args), []) else @@ -430,14 +430,14 @@ let extern_ind_pattern_in_scope (scopes:local_scopes) vars ind args = match availability_of_prim_token p sc scopes with | None -> raise No_match | Some key -> - insert_pat_delimiters Loc.ghost (Loc.tag @@ CPatPrim p) key + insert_pat_delimiters (Loc.tag @@ CPatPrim p) key with No_match -> try if !Flags.raw_print || !print_no_symbol then raise No_match; extern_notation_ind_pattern scopes vars ind args (uninterp_ind_pattern_notations ind) with No_match -> - let c = extern_reference Loc.ghost vars (IndRef ind) in + let c = extern_reference vars (IndRef ind) in let args = List.map (extern_cases_pattern_in_scope scopes vars) args in match drop_implicits_in_patt (IndRef ind) 0 args with |Some true_args -> Loc.tag @@ CPatCstr (c, None, true_args) @@ -490,7 +490,7 @@ let explicitize inctx impl (cf,f) args = is_significant_implicit (Lazy.force a)) in if visible then - (Lazy.force a,Some (Loc.ghost, ExplByName (name_of_implicit imp))) :: tail + (Lazy.force a,Some (Loc.tag @@ ExplByName (name_of_implicit imp))) :: tail else tail | a::args, _::impl -> (Lazy.force a,None) :: exprec (q+1) (args,impl) @@ -615,9 +615,11 @@ let extern_optimal_prim_token scopes r r' = (* mapping decl *) let extended_glob_local_binder_of_decl loc = function - | (p,bk,None,t) -> GLocalAssum (loc,p,bk,t) - | (p,bk,Some x,(_,GHole ( _, Misctypes.IntroAnonymous, None))) -> GLocalDef (loc,p,bk,x,None) - | (p,bk,Some x,t) -> GLocalDef (loc,p,bk,x,Some t) + | (p,bk,None,t) -> GLocalAssum (p,bk,t) + | (p,bk,Some x,(_,GHole ( _, Misctypes.IntroAnonymous, None))) -> GLocalDef (p,bk,x,None) + | (p,bk,Some x,t) -> GLocalDef (p,bk,x,Some t) + +let extended_glob_local_binder_of_decl ?loc u = Loc.tag ?loc (extended_glob_local_binder_of_decl loc u) (**********************************************************************) (* mapping glob_constr to constr_expr *) @@ -645,7 +647,7 @@ let rec extern inctx scopes vars r = with No_match -> Loc.map_with_loc (fun ~loc -> function | GRef (ref,us) -> extern_global (select_stronger_impargs (implicits_of_global ref)) - (extern_reference loc vars ref) (extern_universes us) + (extern_reference ~loc vars ref) (extern_universes us) | GVar id -> CRef (Ident (loc,id),None) @@ -699,7 +701,7 @@ let rec extern inctx scopes vars r = (* we give up since the constructor is not complete *) | (arg, scopes) :: tail -> let head = extern true scopes vars arg in - ip q locs' tail ((extern_reference loc Id.Set.empty (ConstRef c), head) :: acc) + ip q locs' tail ((extern_reference ~loc Id.Set.empty (ConstRef c), head) :: acc) in CRecord (List.rev (ip projs locals args [])) with @@ -707,7 +709,7 @@ let rec extern inctx scopes vars r = let args = extern_args (extern true) vars args in extern_app inctx (select_stronger_impargs (implicits_of_global ref)) - (Some ref,extern_reference rloc vars ref) (extern_universes us) args + (Some ref,extern_reference ~loc:rloc vars ref) (extern_universes us) args end | _ -> @@ -722,12 +724,12 @@ let rec extern inctx scopes vars r = | GProd (na,bk,t,c) -> let t = extern_typ scopes vars t in let (idl,c) = factorize_prod scopes (add_vname vars na) na bk t c in - CProdN ([(Loc.ghost,na)::idl,Default bk,t],c) + CProdN ([(Loc.tag na)::idl,Default bk,t],c) | GLambda (na,bk,t,c) -> let t = extern_typ scopes vars t in let (idl,c) = factorize_lambda inctx scopes (add_vname vars na) na bk t c in - CLambdaN ([(Loc.ghost,na)::idl,Default bk,t],c) + CLambdaN ([(Loc.tag na)::idl,Default bk,t],c) | GCases (sty,rtntypopt,tml,eqns) -> let vars' = @@ -741,12 +743,12 @@ let rec extern inctx scopes vars r = | None -> None | Some ntn -> if occur_glob_constr id ntn then - Some (Loc.ghost, Anonymous) + Some (Loc.tag Anonymous) else None end | Anonymous, _ -> None | Name id, (_, GVar id') when Id.equal id id' -> None - | Name _, _ -> Some (Loc.ghost,na) in + | Name _, _ -> Some (Loc.tag na) in (sub_extern false scopes vars tm, na', Option.map (fun (loc,(ind,nal)) -> @@ -760,15 +762,15 @@ let rec extern inctx scopes vars r = CCases (sty,rtntypopt',tml,eqns) | GLetTuple (nal,(na,typopt),tm,b) -> - CLetTuple (List.map (fun na -> (Loc.ghost,na)) nal, - (Option.map (fun _ -> (Loc.ghost,na)) typopt, + CLetTuple (List.map (fun na -> (Loc.tag na)) nal, + (Option.map (fun _ -> (Loc.tag 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 (c,(na,typopt),b1,b2) -> CIf (sub_extern false scopes vars c, - (Option.map (fun _ -> (Loc.ghost,na)) typopt, + (Option.map (fun _ -> (Loc.tag na)) typopt, Option.map (extern_typ scopes (add_vname vars na)) typopt), sub_extern inctx scopes vars b1, sub_extern inctx scopes vars b2) @@ -779,28 +781,28 @@ let rec extern inctx scopes vars r = let listdecl = Array.mapi (fun i fi -> let (bl,ty,def) = blv.(i), tyv.(i), bv.(i) in - let bl = List.map (extended_glob_local_binder_of_decl loc) bl in + let bl = List.map (extended_glob_local_binder_of_decl ~loc) bl in let (assums,ids,bl) = extern_local_binder scopes vars bl in let vars0 = List.fold_right (name_fold Id.Set.add) ids vars in let vars1 = List.fold_right (name_fold Id.Set.add) ids vars' in let n = match fst nv.(i) with | None -> None - | Some x -> Some (Loc.ghost, out_name (List.nth assums x)) + | Some x -> Some (Loc.tag @@ out_name (List.nth assums x)) in let ro = extern_recursion_order scopes vars (snd nv.(i)) in - ((Loc.ghost, fi), (n, ro), bl, extern_typ scopes vars0 ty, + ((Loc.tag fi), (n, ro), bl, extern_typ scopes vars0 ty, extern false scopes vars1 def)) idv in CFix ((loc,idv.(n)),Array.to_list listdecl) | GCoFix n -> let listdecl = Array.mapi (fun i fi -> - let bl = List.map (extended_glob_local_binder_of_decl loc) blv.(i) in + let bl = List.map (extended_glob_local_binder_of_decl ~loc) blv.(i) in let (_,ids,bl) = extern_local_binder scopes vars bl in let vars0 = List.fold_right (name_fold Id.Set.add) ids vars in let vars1 = List.fold_right (name_fold Id.Set.add) ids vars' in - ((Loc.ghost, fi),bl,extern_typ scopes vars0 tyv.(i), + ((Loc.tag fi),bl,extern_typ scopes vars0 tyv.(i), sub_extern false scopes vars1 bv.(i))) idv in CCoFix ((loc,idv.(n)),Array.to_list listdecl)) @@ -841,14 +843,14 @@ and factorize_lambda inctx scopes vars na bk aty c = and extern_local_binder scopes vars = function [] -> ([],[],[]) - | GLocalDef (_,na,bk,bd,ty)::l -> + | (_, GLocalDef (na,bk,bd,ty))::l -> let (assums,ids,l) = extern_local_binder scopes (name_fold Id.Set.add na vars) l in (assums,na::ids, - CLocalDef((Loc.ghost,na), extern false scopes vars bd, + CLocalDef((Loc.tag na), extern false scopes vars bd, Option.map (extern false scopes vars) ty) :: l) - | GLocalAssum (_,na,bk,ty)::l -> + | (_, GLocalAssum (na,bk,ty))::l -> let ty = extern_typ scopes vars ty in (match extern_local_binder scopes (name_fold Id.Set.add na vars) l with (assums,ids,CLocalAssum(nal,k,ty')::l) @@ -856,12 +858,12 @@ and extern_local_binder scopes vars = function match na with Name id -> not (occur_var_constr_expr id ty') | _ -> true -> (na::assums,na::ids, - CLocalAssum((Loc.ghost,na)::nal,k,ty')::l) + CLocalAssum((Loc.tag na)::nal,k,ty')::l) | (assums,ids,l) -> (na::assums,na::ids, - CLocalAssum([(Loc.ghost,na)],Default bk,ty) :: l)) + CLocalAssum([(Loc.tag na)],Default bk,ty) :: l)) - | GLocalPattern (_,(p,_),_,bk,ty)::l -> + | (_, GLocalPattern ((p,_),_,bk,ty))::l -> let ty = if !Flags.raw_print then Some (extern_typ scopes vars ty) else None in let p = extern_cases_pattern vars p in @@ -1078,5 +1080,5 @@ let extern_rel_context where env sigma sign = let where = Option.map EConstr.of_constr where in let a = detype_rel_context where [] (names_of_rel_context env,env) sigma sign in let vars = vars_of_env env in - let a = List.map (extended_glob_local_binder_of_decl Loc.ghost) a in + let a = List.map (extended_glob_local_binder_of_decl) a in pi3 (extern_local_binder (None,[]) vars a) diff --git a/interp/constrextern.mli b/interp/constrextern.mli index b39339450..ea627cff1 100644 --- a/interp/constrextern.mli +++ b/interp/constrextern.mli @@ -37,7 +37,7 @@ val extern_closed_glob : ?lax:bool -> bool -> env -> Evd.evar_map -> closed_glob val extern_constr : ?lax:bool -> bool -> env -> Evd.evar_map -> constr -> constr_expr val extern_constr_in_scope : bool -> scope_name -> env -> Evd.evar_map -> constr -> constr_expr -val extern_reference : Loc.t -> Id.Set.t -> global_reference -> reference +val extern_reference : ?loc:Loc.t -> Id.Set.t -> global_reference -> reference val extern_type : bool -> env -> Evd.evar_map -> types -> constr_expr val extern_sort : Evd.evar_map -> sorts -> glob_sort val extern_rel_context : constr option -> env -> Evd.evar_map -> @@ -55,9 +55,9 @@ val print_projections : bool ref (** Customization of the global_reference printer *) val set_extern_reference : - (Loc.t -> Id.Set.t -> global_reference -> reference) -> unit + (?loc:Loc.t -> Id.Set.t -> global_reference -> reference) -> unit val get_extern_reference : - unit -> (Loc.t -> Id.Set.t -> global_reference -> reference) + unit -> (?loc:Loc.t -> Id.Set.t -> global_reference -> reference) (** This governs printing of implicit arguments. If [with_implicits] is on and not [with_arguments] then implicit args are printed prefixed diff --git a/interp/constrintern.ml b/interp/constrintern.ml index d1b931a22..585f03808 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -345,13 +345,13 @@ let rec check_capture ty = function | [] -> () -let locate_if_hole loc na = function +let locate_if_hole ?loc na = function | _, GHole (_,naming,arg) -> (try match na with - | Name id -> glob_constr_of_notation_constr loc + | Name id -> glob_constr_of_notation_constr ?loc (Reserve.find_reserved_type id) | Anonymous -> raise Not_found - with Not_found -> Loc.tag ~loc @@ GHole (Evar_kinds.BinderType na, naming, arg)) + with Not_found -> Loc.tag ?loc @@ GHole (Evar_kinds.BinderType na, naming, arg)) | x -> x let reset_hidden_inductive_implicit_test env = @@ -424,7 +424,7 @@ let intern_assumption intern lvar env nal bk ty = List.fold_left (fun (env, bl) (loc, na as locna) -> (push_name_env lvar impls env locna, - (loc,(na,k,locate_if_hole loc na ty))::bl)) + (loc,(na,k,locate_if_hole ~loc na ty))::bl)) (env, []) nal | Generalized (b,b',t) -> let env, b = intern_generalized_binder intern_type lvar env (List.hd nal) b b' t ty in @@ -454,27 +454,28 @@ let intern_local_pattern intern lvar env p = env) env (free_vars_of_pat [] p) -let glob_local_binder_of_extended = function - | GLocalAssum (loc,na,bk,t) -> (na,bk,None,t) - | GLocalDef (loc,na,bk,c,Some t) -> (na,bk,Some c,t) - | GLocalDef (loc,na,bk,c,None) -> +let glob_local_binder_of_extended = Loc.with_loc (fun ~loc -> function + | GLocalAssum (na,bk,t) -> (na,bk,None,t) + | GLocalDef (na,bk,c,Some t) -> (na,bk,Some c,t) + | GLocalDef (na,bk,c,None) -> let t = Loc.tag ~loc @@ GHole(Evar_kinds.BinderType na,Misctypes.IntroAnonymous,None) in (na,bk,Some c,t) - | GLocalPattern (loc,_,_,_,_) -> + | GLocalPattern (_,_,_,_) -> Loc.raise ~loc (Stream.Error "pattern with quote not allowed here.") + ) let intern_cases_pattern_fwd = ref (fun _ -> failwith "intern_cases_pattern_fwd") let intern_local_binder_aux ?(global_level=false) intern lvar (env,bl) = function | CLocalAssum(nal,bk,ty) -> let env, bl' = intern_assumption intern lvar env nal bk ty in - let bl' = List.map (fun (loc,(na,c,t)) -> GLocalAssum (loc,na,c,t)) bl' in + let bl' = List.map (fun (loc,(na,c,t)) -> Loc.tag ~loc @@ GLocalAssum (na,c,t)) bl' in env, bl' @ bl | CLocalDef((loc,na as locna),def,ty) -> let term = intern env def in let ty = Option.map (intern env) ty in (push_name_env lvar (impls_term_list term) env locna, - GLocalDef (loc,na,Explicit,term,ty) :: bl) + (Loc.tag ~loc @@ GLocalDef (na,Explicit,term,ty)) :: bl) | CLocalPattern (loc,(p,ty)) -> let tyc = match ty with @@ -494,7 +495,7 @@ let intern_local_binder_aux ?(global_level=false) intern lvar (env,bl) = functio let bk = Default Explicit in let _, bl' = intern_assumption intern lvar env [na] bk tyc in let _,(_,bk,t) = List.hd bl' in - (env, GLocalPattern(loc,(cp,il),id,bk,t) :: bl) + (env, (Loc.tag ~loc @@ GLocalPattern((cp,il),id,bk,t)) :: bl) let intern_generalization intern env lvar loc bk ak c = let c = intern {env with unb = true} c in @@ -582,13 +583,13 @@ let make_letins = let rec subordinate_letins letins = function (* binders come in reverse order; the non-let are returned in reverse order together *) (* with the subordinated let-in in writing order *) - | GLocalDef (loc,na,_,b,t)::l -> + | (loc, GLocalDef (na,_,b,t))::l -> subordinate_letins (LPLetIn (loc,(na,b,t))::letins) l - | GLocalAssum (loc,na,bk,t)::l -> + | (loc, GLocalAssum (na,bk,t))::l -> let letins',rest = subordinate_letins [] l in letins',((loc,(na,bk,t)),letins)::rest - | GLocalPattern (loc,u,id,bk,t) :: l -> - subordinate_letins (LPCases (loc,u,id)::letins) ([GLocalAssum (loc,Name id,bk,t)] @ l) + | (loc, GLocalPattern (u,id,bk,t)) :: l -> + subordinate_letins (LPCases (loc,u,id)::letins) ([Loc.tag ~loc @@ GLocalAssum (Name id,bk,t)] @ l) | [] -> letins,[] @@ -602,11 +603,11 @@ let terms_of_binders bl = let params = List.make (Inductiveops.inductive_nparams (fst c)) hole in CAppExpl ((None,r,None),params @ List.map term_of_pat l)) pt in let rec extract_variables = function - | 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." - | GLocalPattern (loc,(u,_),_,_,_) :: l -> term_of_pat u :: extract_variables l + | (loc, GLocalAssum (Name id,_,_))::l -> (Loc.tag ~loc @@ CRef (Ident (loc,id), None)) :: extract_variables l + | (loc, GLocalDef (Name id,_,_,_))::l -> extract_variables l + | (loc, GLocalDef (Anonymous,_,_,_))::l + | (loc, GLocalAssum (Anonymous,_,_))::l -> error "Cannot turn \"_\" into a term." + | (loc, GLocalPattern ((u,_),_,_,_)) :: l -> term_of_pat u :: extract_variables l | [] -> [] in extract_variables bl @@ -697,7 +698,7 @@ let instantiate_notation_constr loc intern ntnvars subst infos c = let ty = Loc.tag ~loc @@ GHole (Evar_kinds.BinderType na,naming,arg) in Loc.tag ~loc @@ GLambda (na,Explicit,ty,aux subst' subinfos c') | t -> - glob_constr_of_notation_constr_with_binders loc + glob_constr_of_notation_constr_with_binders ~loc (traverse_binder subst avoid) (aux subst') subinfos t and subst_var (terms, _binderopt, _terminopt) (renaming, env) id = (* subst remembers the delimiters stack in the interpretation *) @@ -728,7 +729,7 @@ let make_subst ids l = let intern_notation intern env lvar loc ntn fullargs = let ntn,(args,argslist,bll as fullargs) = contract_notation ntn fullargs in - let ((ids,c),df) = interp_notation loc ntn (env.tmp_scope,env.scopes) in + let ((ids,c),df) = interp_notation ~loc ntn (env.tmp_scope,env.scopes) in Dumpglob.dump_notation_location (ntn_loc loc fullargs ntn) ntn df; let ids,idsl,idsbl = split_by_type ids in let terms = make_subst ids args in @@ -809,8 +810,8 @@ let find_appl_head_data c = List.skipn_at_least n scopes,[] | _ -> c,[],[],[] -let error_not_enough_arguments loc = - user_err ~loc (str "Abbreviation is not applied enough.") +let error_not_enough_arguments ?loc = + user_err ?loc (str "Abbreviation is not applied enough.") let check_no_explicitation l = let is_unset (a, b) = match b with None -> false | Some _ -> true in @@ -843,7 +844,7 @@ let intern_qualid loc qid intern env lvar us args = | SynDef sp -> let (ids,c) = Syntax_def.search_syntactic_definition sp in let nids = List.length ids in - if List.length args < nids then error_not_enough_arguments loc; + if List.length args < nids then error_not_enough_arguments ~loc; let args1,args2 = List.chop nids args in check_no_explicitation args1; let terms = make_subst ids (List.map fst args1) in @@ -893,7 +894,7 @@ let intern_applied_reference intern env namedctx (_, ntnvars as lvar) us args = let interp_reference vars r = let (r,_,_,_),_ = - intern_applied_reference (fun _ -> error_not_enough_arguments Loc.ghost) + intern_applied_reference (fun _ -> error_not_enough_arguments ?loc:None) {ids = Id.Set.empty; unb = false ; tmp_scope = None; scopes = []; impls = empty_internalization_env} [] (vars, Id.Map.empty) None [] r @@ -990,10 +991,10 @@ let add_implicits_check_length fail nargs nargs_with_letin impls_st len_pl1 pl2 else Int.equal args_len nargs_with_letin || (fst (fail (nargs - List.length impl_list + i)))) ,l) |imp::q as il,[] -> if is_status_implicit imp && maximal_insertion_of imp - then let (b,out) = aux i (q,[]) in (b,(Loc.ghost,RCPatAtom(None))::out) + then let (b,out) = aux i (q,[]) in (b,(Loc.tag @@ RCPatAtom(None))::out) else fail (remaining_args (len_pl1+i) il) |imp::q,(hh::tt as l) -> if is_status_implicit imp - then let (b,out) = aux i (q,l) in (b,(Loc.ghost, RCPatAtom(None))::out) + then let (b,out) = aux i (q,l) in (b,(Loc.tag @@ RCPatAtom(None))::out) else let (b,out) = aux (succ i) (q,tt) in (b,hh::out) in aux 0 (impl_list,pl2) @@ -1239,7 +1240,7 @@ let drop_notations_pattern looked_for = (* Convention: do not deactivate implicit arguments and scopes for further arguments *) test_kind top g; let nvars = List.length vars in - if List.length pats < nvars then error_not_enough_arguments loc; + if List.length pats < nvars then error_not_enough_arguments ~loc; let pats1,pats2 = List.chop nvars pats in let subst = make_subst vars pats1 in let idspl1 = List.map (in_not false loc scopes (subst, Id.Map.empty) []) args in @@ -1288,20 +1289,20 @@ let drop_notations_pattern looked_for = Loc.tag ~loc @@ RCPatCstr (g, List.map2 (in_pat_sc scopes) argscs1 expl_pl @ List.map (in_pat false scopes) pl, []) | CPatNotation ("- _",([_loc,CPatPrim(Numeral p)],[]),[]) when Bigint.is_strictly_pos p -> - fst (Notation.interp_prim_token_cases_pattern_expr loc (ensure_kind false loc) (Numeral (Bigint.neg p)) scopes) + fst (Notation.interp_prim_token_cases_pattern_expr ~loc (ensure_kind false loc) (Numeral (Bigint.neg p)) scopes) | CPatNotation ("( _ )",([a],[]),[]) -> in_pat top scopes a | CPatNotation (ntn, fullargs,extrargs) -> let ntn,(args,argsl as fullargs) = contract_pat_notation ntn fullargs in - let ((ids',c),df) = Notation.interp_notation loc ntn scopes in + let ((ids',c),df) = Notation.interp_notation ~loc ntn scopes in let (ids',idsl',_) = split_by_type ids' in Dumpglob.dump_notation_location (patntn_loc loc fullargs ntn) ntn df; let substlist = make_subst idsl' argsl in let subst = make_subst ids' args in in_not top loc scopes (subst,substlist) extrargs c | CPatDelimiters (key, e) -> - in_pat top (None,find_delimiters_scope loc key::snd scopes) e - | CPatPrim p -> fst (Notation.interp_prim_token_cases_pattern_expr loc (test_kind false) p scopes) + in_pat top (None,find_delimiters_scope ~loc key::snd scopes) e + | CPatPrim p -> fst (Notation.interp_prim_token_cases_pattern_expr ~loc (test_kind false) p scopes) | CPatAtom Some id -> begin match drop_syndef top scopes id [] with @@ -1540,7 +1541,9 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let before, after = split_at_annot bl n in let (env',rbefore) = List.fold_left intern_local_binder (env,[]) before in let ro = f (intern env') in - let n' = Option.map (fun _ -> List.count (function GLocalAssum _ -> true | _ -> false (* remove let-ins *)) rbefore) n in + let n' = Option.map (fun _ -> List.count (function | _, GLocalAssum _ -> true + | _ -> false (* remove let-ins *)) + rbefore) n in n', ro, List.fold_left intern_local_binder (env',rbefore) after in let n, ro, (env',rbl) = @@ -1559,7 +1562,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let (_,bli,tyi,_) = idl_temp.(i) in let fix_args = (List.map (fun (na, bk, _, _) -> (build_impls bk na)) bli) in push_name_env ntnvars (impls_type_list ~args:fix_args tyi) - en (Loc.ghost, Name name)) 0 env' lf in + en (Loc.tag @@ Name name)) 0 env' lf in (a,b,c,intern {env'' with tmp_scope = None} bd)) dl idl_temp in Loc.tag ~loc @@ GRec (GFix @@ -1586,7 +1589,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let (bli,tyi,_) = idl_tmp.(i) in let cofix_args = List.map (fun (na, bk, _, _) -> (build_impls bk na)) bli in push_name_env ntnvars (impls_type_list ~args:cofix_args tyi) - en (Loc.ghost, Name name)) 0 env' lf in + en (Loc.tag @@ Name name)) 0 env' lf in (b,c,intern {env'' with tmp_scope = None} bd)) dl idl_tmp in Loc.tag ~loc @@ GRec (GCoFix n, @@ -1617,10 +1620,10 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = | CGeneralization (b,a,c) -> intern_generalization intern env ntnvars loc b a c | CPrim p -> - fst (Notation.interp_prim_token loc p (env.tmp_scope,env.scopes)) + fst (Notation.interp_prim_token ~loc p (env.tmp_scope,env.scopes)) | CDelimiters (key, e) -> intern {env with tmp_scope = None; - scopes = find_delimiters_scope loc key :: env.scopes} e + scopes = find_delimiters_scope ~loc key :: env.scopes} e | CAppExpl ((isproj,ref,us), args) -> let (f,_,args_scopes,_),args = let args = List.map (fun a -> (a,None)) args in @@ -1679,7 +1682,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = (tm,ind)::inds, Option.fold_right Id.Set.add extra_id ex_ids, List.rev_append match_td matchs) tms ([],Id.Set.empty,[]) in let env' = Id.Set.fold - (fun var bli -> push_name_env ntnvars (Variable,[],[],[]) bli (Loc.ghost,Name var)) + (fun var bli -> push_name_env ntnvars (Variable,[],[],[]) bli (Loc.tag @@ Name var)) (Id.Set.union ex_ids as_in_vars) (reset_hidden_inductive_implicit_test env) in (* PatVars before a real pattern do not need to be matched *) let stripped_match_from_in = @@ -1715,7 +1718,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let ((b',(na',_)),_,_) = intern_case_item env' Id.Set.empty (b,na,None) in let p' = Option.map (fun u -> let env'' = push_name_env ntnvars (Variable,[],[],[]) (reset_hidden_inductive_implicit_test env') - (Loc.ghost,na') in + (Loc.tag na') in intern_type env'' u) po in Loc.tag ~loc @@ GLetTuple (List.map snd nal, (na', p'), b', @@ -1725,7 +1728,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = 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 -> let env'' = push_name_env ntnvars (Variable,[],[],[]) (reset_hidden_inductive_implicit_test env) - (Loc.ghost,na') in + (Loc.tag na') in intern_type env'' p) po in Loc.tag ~loc @@ GIf (c', (na', p'), intern env b1, intern env b2) @@ -1779,7 +1782,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = ) and intern_type env = intern (set_type_scope env) - and intern_local_binder env bind = + and intern_local_binder env bind : intern_env * Glob_term.extended_glob_local_binder list = intern_local_binder_aux intern ntnvars env bind (* Expands a multiple pattern into a disjunction of multiple patterns *) @@ -1815,7 +1818,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let extra_id,na = match tm', na with | (loc , GVar id), None when not (Id.Map.mem id (snd lvar)) -> Some id,(loc,Name id) | (loc, GRef (VarRef id, _)), None -> Some id,(loc,Name id) - | _, None -> None,(Loc.ghost,Anonymous) + | _, None -> None,(Loc.tag Anonymous) | _, Some (loc,na) -> None,(loc,na) in (* the "in" part *) let match_td,typ = match t with @@ -1837,7 +1840,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = match case_rel_ctxt,arg_pats with (* LetIn in the rel_context *) | LocalDef _ :: t, l when not with_letin -> - canonize_args t l forbidden_names match_acc ((Loc.ghost,Anonymous)::var_acc) + canonize_args t l forbidden_names match_acc ((Loc.tag Anonymous)::var_acc) | [],[] -> (add_name match_acc na, var_acc) | _::t, (loc, PatVar x)::tt -> @@ -2052,12 +2055,12 @@ let interp_notation_constr ?(impls=empty_internalization_env) nenv a = let interp_binder env sigma na t = let t = intern_gen IsType env t in - let t' = locate_if_hole (loc_of_glob_constr t) na t in + let t' = locate_if_hole ~loc:(loc_of_glob_constr t) na t in understand ~expected_type:IsType env sigma t' let interp_binder_evars env evdref na t = let t = intern_gen IsType env t in - let t' = locate_if_hole (loc_of_glob_constr t) na t in + let t' = locate_if_hole ~loc:(loc_of_glob_constr t) na t in understand_tcc_evars env evdref ~expected_type:IsType t' open Environ @@ -2084,7 +2087,7 @@ let interp_rawcontext_evars env evdref k bl = List.fold_left (fun (env,params,n,impls) (na, k, b, t) -> let t' = - if Option.is_empty b then locate_if_hole (loc_of_glob_constr t) na t + if Option.is_empty b then locate_if_hole ~loc:(loc_of_glob_constr t) na t else t in let t = understand_tcc_evars env evdref ~expected_type:IsType t' in diff --git a/interp/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml index 51152bb24..fa7712bdc 100644 --- a/interp/implicit_quantifiers.ml +++ b/interp/implicit_quantifiers.ml @@ -252,7 +252,7 @@ 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 - (Loc.tag @@ CRef (Ident (Loc.ghost, id'),None), Id.Set.add id' avoid) + (Loc.tag @@ CRef (Ident (Loc.tag id'),None), Id.Set.add id' avoid) let destClassApp (loc, cl) = match cl with diff --git a/interp/notation.ml b/interp/notation.ml index 3bcec3001..150be040f 100644 --- a/interp/notation.ml +++ b/interp/notation.ml @@ -220,10 +220,10 @@ let remove_delimiters scope = with Not_found -> assert false (* A delimiter for scope [scope] should exist *) -let find_delimiters_scope loc key = +let find_delimiters_scope ?loc key = try String.Map.find key !delimiters_map with Not_found -> - user_err ~loc ~hdr:"find_delimiters" + user_err ?loc ~hdr:"find_delimiters" (str "Unknown scope delimiting key " ++ str key ++ str ".") (* Uninterpretation tables *) @@ -291,7 +291,7 @@ let notation_constr_key = function (* Rem: NApp(NRef ref,[]) stands for @ref *) type required_module = full_path * string list type 'a prim_token_interpreter = - Loc.t -> 'a -> glob_constr + ?loc:Loc.t -> 'a -> glob_constr type cases_pattern_status = bool (* true = use prim token in patterns *) @@ -299,7 +299,7 @@ type 'a prim_token_uninterpreter = glob_constr list * (glob_constr -> 'a option) * cases_pattern_status type internal_prim_token_interpreter = - Loc.t -> prim_token -> required_module * (unit -> glob_constr) + ?loc:Loc.t -> prim_token -> required_module * (unit -> glob_constr) let prim_token_interpreter_tab = (Hashtbl.create 7 : (scope_name, internal_prim_token_interpreter) Hashtbl.t) @@ -309,7 +309,7 @@ let add_prim_token_interpreter sc interp = let cont = Hashtbl.find prim_token_interpreter_tab sc in Hashtbl.replace prim_token_interpreter_tab sc (interp cont) with Not_found -> - let cont = (fun _loc _p -> raise Not_found) in + let cont = (fun ?loc _p -> raise Not_found) in Hashtbl.add prim_token_interpreter_tab sc (interp cont) let declare_prim_token_interpreter sc interp (patl,uninterp,b) = @@ -325,22 +325,22 @@ let mkString = function | None -> None | Some s -> if Unicode.is_utf8 s then Some (String s) else None -let delay dir int loc x = (dir, (fun () -> int loc x)) +let delay dir int ?loc x = (dir, (fun () -> int ?loc x)) let declare_numeral_interpreter sc dir interp (patl,uninterp,inpat) = declare_prim_token_interpreter sc - (fun cont loc -> function Numeral n-> delay dir interp loc n | p -> cont loc p) + (fun cont ?loc -> function Numeral n-> delay dir interp ?loc n | p -> cont ?loc p) (patl, (fun r -> Option.map mkNumeral (uninterp r)), inpat) let declare_string_interpreter sc dir interp (patl,uninterp,inpat) = declare_prim_token_interpreter sc - (fun cont loc -> function String s -> delay dir interp loc s | p -> cont loc p) + (fun cont ?loc -> function String s -> delay dir interp ?loc s | p -> cont ?loc p) (patl, (fun r -> mkString (uninterp r)), inpat) -let check_required_module loc sc (sp,d) = +let check_required_module ?loc sc (sp,d) = try let _ = Nametab.global_of_path sp in () with Not_found -> - user_err ~loc ~hdr:"prim_token_interpreter" + user_err ?loc ~hdr:"prim_token_interpreter" (str "Cannot interpret in " ++ str sc ++ str " without requiring first module " ++ str (List.last d) ++ str ".") (* Look if some notation or numeral printer in [scope] can be used in @@ -445,23 +445,23 @@ let notation_of_prim_token = function | Numeral n -> "- "^(to_string (neg n)) | String _ -> raise Not_found -let find_prim_token g loc p sc = +let find_prim_token ?loc g p sc = (* Try for a user-defined numerical notation *) try let (_,c),df = find_notation (notation_of_prim_token p) sc in - g (Notation_ops.glob_constr_of_notation_constr loc c),df + g (Notation_ops.glob_constr_of_notation_constr ?loc c),df with Not_found -> (* Try for a primitive numerical notation *) - let (spdir,interp) = Hashtbl.find prim_token_interpreter_tab sc loc p in - check_required_module loc sc spdir; + let (spdir,interp) = (Hashtbl.find prim_token_interpreter_tab sc) ?loc p in + check_required_module ?loc sc spdir; g (interp ()), ((dirpath (fst spdir),DirPath.empty),"") -let interp_prim_token_gen g loc p local_scopes = +let interp_prim_token_gen g ?loc p local_scopes = let scopes = make_current_scopes local_scopes in let p_as_ntn = try notation_of_prim_token p with Not_found -> "" in - try find_interpretation p_as_ntn (find_prim_token g loc p) scopes + try find_interpretation p_as_ntn (find_prim_token ?loc g p) scopes with Not_found -> - user_err ~loc ~hdr:"interp_prim_token" + user_err ?loc ~hdr:"interp_prim_token" ((match p with | Numeral n -> str "No interpretation for numeral " ++ str (to_string n) | String s -> str "No interpretation for string " ++ qs s) ++ str ".") @@ -480,14 +480,14 @@ let rec rcp_of_glob looked_for gt = Loc.map (function | _ -> raise Not_found ) gt -let interp_prim_token_cases_pattern_expr loc looked_for p = - interp_prim_token_gen (rcp_of_glob looked_for) loc p +let interp_prim_token_cases_pattern_expr ?loc looked_for p = + interp_prim_token_gen (rcp_of_glob looked_for) ?loc p -let interp_notation loc ntn local_scopes = +let interp_notation ?loc ntn local_scopes = let scopes = make_current_scopes local_scopes in try find_interpretation ntn (find_notation ntn) scopes with Not_found -> - user_err ~loc + user_err ?loc (str "Unknown interpretation for notation \"" ++ str ntn ++ str "\".") let uninterp_notations c = @@ -541,7 +541,7 @@ let uninterp_prim_token_cases_pattern c = let availability_of_prim_token n printer_scope local_scopes = let f scope = - try ignore (Hashtbl.find prim_token_interpreter_tab scope Loc.ghost n); true + try ignore ((Hashtbl.find prim_token_interpreter_tab scope) n); true with Not_found -> false in let scopes = make_current_scopes local_scopes in Option.map snd (find_without_delimiters f (Some printer_scope,None) scopes) @@ -823,7 +823,7 @@ let pr_scope_classes sc = let pr_notation_info prglob ntn c = str "\"" ++ str ntn ++ str "\" := " ++ - prglob (Notation_ops.glob_constr_of_notation_constr Loc.ghost c) + prglob (Notation_ops.glob_constr_of_notation_constr c) let pr_named_scope prglob scope sc = (if String.equal scope default_scope then @@ -891,25 +891,25 @@ let global_reference_of_notation test (ntn,(sc,c,_)) = Some (ntn,sc,ref) | _ -> None -let error_ambiguous_notation loc _ntn = - user_err ~loc (str "Ambiguous notation.") +let error_ambiguous_notation ?loc _ntn = + user_err ?loc (str "Ambiguous notation.") -let error_notation_not_reference loc ntn = - user_err ~loc +let error_notation_not_reference ?loc ntn = + user_err ?loc (str "Unable to interpret " ++ quote (str ntn) ++ str " as a reference.") -let interp_notation_as_global_reference loc test ntn sc = +let interp_notation_as_global_reference ?loc test ntn sc = let scopes = match sc with | Some sc -> - let scope = find_scope (find_delimiters_scope Loc.ghost sc) in + let scope = find_scope (find_delimiters_scope sc) in String.Map.add sc scope String.Map.empty | None -> !scope_map in let ntns = browse_notation true ntn scopes in let refs = List.map (global_reference_of_notation test) ntns in match Option.List.flatten refs with | [_,_,ref] -> ref - | [] -> error_notation_not_reference loc ntn + | [] -> error_notation_not_reference ?loc ntn | refs -> let f (ntn,sc,ref) = let def = find_default ntn !scope_stack in @@ -919,8 +919,8 @@ let interp_notation_as_global_reference loc test ntn sc = in match List.filter f refs with | [_,_,ref] -> ref - | [] -> error_notation_not_reference loc ntn - | _ -> error_ambiguous_notation loc ntn + | [] -> error_notation_not_reference ?loc ntn + | _ -> error_ambiguous_notation ?loc ntn let locate_notation prglob ntn scope = let ntns = factorize_entries (browse_notation false ntn !scope_map) in diff --git a/interp/notation.mli b/interp/notation.mli index 2e92a00a8..10c7b85e4 100644 --- a/interp/notation.mli +++ b/interp/notation.mli @@ -55,7 +55,7 @@ val find_scope : scope_name -> scope val declare_delimiters : scope_name -> delimiters -> unit val remove_delimiters : scope_name -> unit -val find_delimiters_scope : Loc.t -> delimiters -> scope_name +val find_delimiters_scope : ?loc:Loc.t -> delimiters -> scope_name (** {6 Declare and uses back and forth an interpretation of primitive token } *) @@ -69,7 +69,7 @@ type required_module = full_path * string list type cases_pattern_status = bool (** true = use prim token in patterns *) type 'a prim_token_interpreter = - Loc.t -> 'a -> glob_constr + ?loc:Loc.t -> 'a -> glob_constr type 'a prim_token_uninterpreter = glob_constr list * (glob_constr -> 'a option) * cases_pattern_status @@ -83,9 +83,9 @@ val declare_string_interpreter : scope_name -> required_module -> (** Return the [term]/[cases_pattern] bound to a primitive token in a given scope context*) -val interp_prim_token : Loc.t -> prim_token -> local_scopes -> +val interp_prim_token : ?loc:Loc.t -> prim_token -> local_scopes -> glob_constr * (notation_location * scope_name option) -val interp_prim_token_cases_pattern_expr : Loc.t -> (global_reference -> unit) -> prim_token -> +val interp_prim_token_cases_pattern_expr : ?loc:Loc.t -> (global_reference -> unit) -> prim_token -> local_scopes -> raw_cases_pattern_expr * (notation_location * scope_name option) (** Return the primitive token associated to a [term]/[cases_pattern]; @@ -114,7 +114,7 @@ val declare_notation_interpretation : notation -> scope_name option -> val declare_uninterpretation : interp_rule -> interpretation -> unit (** Return the interpretation bound to a notation *) -val interp_notation : Loc.t -> notation -> local_scopes -> +val interp_notation : ?loc:Loc.t -> notation -> local_scopes -> interpretation * (notation_location * scope_name option) type notation_rule = interp_rule * interpretation * int option @@ -137,7 +137,7 @@ val level_of_notation : notation -> level (** raise [Not_found] if no level *) (** {6 Miscellaneous} *) -val interp_notation_as_global_reference : Loc.t -> (global_reference -> bool) -> +val interp_notation_as_global_reference : ?loc:Loc.t -> (global_reference -> bool) -> notation -> delimiters option -> global_reference (** Checks for already existing notations *) diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml index 32c900504..32c564156 100644 --- a/interp/notation_ops.ml +++ b/interp/notation_ops.ml @@ -117,13 +117,13 @@ let name_to_ident = function let to_id g e id = let e,na = g e (Name id) in e,name_to_ident na -let rec cases_pattern_fold_map loc g e = Loc.with_unloc (function +let rec cases_pattern_fold_map ?loc g e = Loc.with_unloc (function | PatVar na -> - let e',na' = g e na in e', Loc.tag ~loc @@ PatVar na' + let e',na' = g e na in e', Loc.tag ?loc @@ PatVar na' | PatCstr (cstr,patl,na) -> let e',na' = g e na in - let e',patl' = List.fold_map (cases_pattern_fold_map loc g) e patl in - e', Loc.tag ~loc @@ PatCstr (cstr,patl',na') + let e',patl' = List.fold_map (cases_pattern_fold_map ?loc g) e patl in + e', Loc.tag ?loc @@ PatCstr (cstr,patl',na') ) let subst_binder_type_vars l = function @@ -152,8 +152,8 @@ let rec subst_glob_vars l gc = Loc.map (function let ldots_var = Id.of_string ".." -let glob_constr_of_notation_constr_with_binders loc g f e nc = - let lt x = Loc.tag ~loc x in lt @@ match nc with +let glob_constr_of_notation_constr_with_binders ?loc g f e nc = + let lt x = Loc.tag ?loc x in lt @@ match nc with | NVar id -> GVar id | NApp (a,args) -> GApp (f e a, List.map (f e) args) | NList (x,y,iter,tail,swap) -> @@ -181,13 +181,13 @@ let glob_constr_of_notation_constr_with_binders loc g f e nc = | Some (ind,nal) -> let e',nal' = List.fold_right (fun na (e',nal) -> let e',na' = g e' na in e',na'::nal) nal (e',[]) in - e',Some (loc,(ind,nal')) in + e',Some (Loc.tag ?loc (ind,nal')) in let e',na' = g e' na in (e',(f e tm,(na',t'))::tml')) tml (e,[]) in let fold (idl,e) na = let (e,na) = g e na in ((name_cons na idl,e),na) in let eqnl' = List.map (fun (patl,rhs) -> let ((idl,e),patl) = - List.fold_map (cases_pattern_fold_map loc fold) ([],e) patl in + List.fold_map (cases_pattern_fold_map ?loc fold) ([],e) patl in lt (idl,patl,f e rhs)) eqnl in GCases (sty,Option.map (f e') rtntypopt,tml',eqnl') | NLetTuple (nal,(na,po),b,c) -> @@ -208,9 +208,9 @@ let glob_constr_of_notation_constr_with_binders loc g f e nc = | NHole (x, naming, arg) -> GHole (x, naming, arg) | NRef x -> GRef (x,None) -let glob_constr_of_notation_constr loc x = +let glob_constr_of_notation_constr ?loc x = let rec aux () x = - glob_constr_of_notation_constr_with_binders loc (fun () id -> ((),id)) aux () x + glob_constr_of_notation_constr_with_binders ?loc (fun () id -> ((),id)) aux () x in aux () x (******************************************************************************) @@ -795,17 +795,17 @@ let bind_bindinglist_env alp (terms,onlybinders,termlists,binderlists as sigma) | (Some _ as x), None | None, (Some _ as x) -> x | None, None -> None in let unify_binding_kind bk bk' = if bk == bk' then bk' else raise No_match in - let unify_binder alp b b' = + let unify_binder alp (loc, b) (loc', b') = match b, b' with - | GLocalAssum (loc,na,bk,t), GLocalAssum (_,na',bk',t') -> + | GLocalAssum (na,bk,t), GLocalAssum (na',bk',t') -> let alp, na = unify_name alp na na' in - alp, GLocalAssum (loc, na, unify_binding_kind bk bk', unify_term alp t t') - | GLocalDef (loc,na,bk,c,t), GLocalDef (_,na',bk',c',t') -> + alp, Loc.tag ~loc @@ GLocalAssum (na, unify_binding_kind bk bk', unify_term alp t t') + | GLocalDef (na,bk,c,t), GLocalDef (na',bk',c',t') -> let alp, na = unify_name alp na na' in - alp, GLocalDef (loc, na, unify_binding_kind bk bk', unify_term alp c c', unify_opt_term alp t t') - | GLocalPattern (loc,(p,ids),id,bk,t), GLocalPattern (_,(p',_),_,bk',t') -> + alp, Loc.tag ~loc @@ GLocalDef (na, unify_binding_kind bk bk', unify_term alp c c', unify_opt_term alp t t') + | GLocalPattern ((p,ids),id,bk,t), GLocalPattern ((p',_),_,bk',t') -> let alp, p = unify_pat alp p p' in - alp, GLocalPattern (loc, (p,ids), id, unify_binding_kind bk bk', unify_term alp t t') + alp, Loc.tag ~loc @@ GLocalPattern ((p,ids), id, unify_binding_kind bk bk', unify_term alp t t') | _ -> raise No_match in let rec unify alp bl bl' = match bl, bl' with @@ -832,18 +832,18 @@ let bind_bindinglist_as_term_env alp (terms,onlybinders,termlists,binderlists) v let unify_pat p p' = if cases_pattern_eq (map_cases_pattern_name_left (name_app (rename_var (snd alp))) p) p' then p' else raise No_match in - let unify_term_binder c b' = + let unify_term_binder c (loc, b') = Loc.tag ~loc @@ match c, b' with - | (_, GVar id), GLocalAssum (loc, na', bk', t') -> - GLocalAssum (loc, unify_id id na', bk', t') - | c, GLocalPattern (loc, (p',ids), id, bk', t') -> + | (_, GVar id), GLocalAssum (na', bk', t') -> + GLocalAssum (unify_id id na', bk', t') + | c, GLocalPattern ((p',ids), id, bk', t') -> let p = pat_binder_of_term c in - GLocalPattern (loc, (unify_pat p p',ids), id, bk', t') + GLocalPattern ((unify_pat p p',ids), id, bk', t') | _ -> raise No_match in let rec unify cl bl' = match cl, bl' with | [], [] -> [] - | c :: cl, GLocalDef (_, _, _, _, t) :: bl' -> unify cl bl' + | c :: cl, (_loc, GLocalDef ( _, _, _, t)) :: bl' -> unify cl bl' | c :: cl, b' :: bl' -> unify_term_binder c b' :: unify cl bl' | _ -> raise No_match in let bl = unify cl bl' in @@ -898,17 +898,17 @@ let glue_letin_with_decls = true let rec match_iterated_binders islambda decls bi = Loc.with_loc (fun ~loc -> function | GLambda (Name p,bk,t,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b))]))) when islambda && Id.equal p e -> - match_iterated_binders islambda (GLocalPattern(loc,(cp,ids),p,bk,t)::decls) b + match_iterated_binders islambda ((Loc.tag ~loc @@ GLocalPattern((cp,ids),p,bk,t))::decls) b | GLambda (na,bk,t,b) when islambda -> - match_iterated_binders islambda (GLocalAssum(loc,na,bk,t)::decls) b + match_iterated_binders islambda ((Loc.tag ~loc @@ GLocalAssum(na,bk,t))::decls) b | GProd (Name p,bk,t,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b))]))) when not islambda && Id.equal p e -> - match_iterated_binders islambda (GLocalPattern(loc,(cp,ids),p,bk,t)::decls) b + match_iterated_binders islambda ((Loc.tag ~loc @@ GLocalPattern((cp,ids),p,bk,t))::decls) b | GProd ((Name _ as na),bk,t,b) when not islambda -> - match_iterated_binders islambda (GLocalAssum(loc,na,bk,t)::decls) b + match_iterated_binders islambda ((Loc.tag ~loc @@ GLocalAssum(na,bk,t))::decls) b | GLetIn (na,c,t,b) when glue_letin_with_decls -> match_iterated_binders islambda - (GLocalDef (loc,na,Explicit (*?*), c,t)::decls) b + ((Loc.tag ~loc @@ GLocalDef (na,Explicit (*?*), c,t))::decls) b | b -> (decls, Loc.tag ~loc b) ) bi @@ -989,13 +989,13 @@ let rec match_ inner u alp metas sigma a1 a2 = (* "λ p, let 'cp = p in t" -> "λ 'cp, t" *) | GLambda (Name p,bk,t1,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b1))]))), NBinderList (x,_,NLambda (Name _id2,_,b2),termin) when Id.equal p e -> - let (decls,b) = match_iterated_binders true [GLocalPattern(loc,(cp,ids),p,bk,t1)] b1 in + let (decls,b) = match_iterated_binders true [Loc.tag ~loc @@ GLocalPattern((cp,ids),p,bk,t1)] b1 in let alp,sigma = bind_bindinglist_env alp sigma x decls in match_in u alp metas sigma b termin (* Matching recursive notations for binders: ad hoc cases supporting let-in *) | GLambda (na1,bk,t1,b1), NBinderList (x,_,NLambda (Name _id2,_,b2),termin)-> - let (decls,b) = match_iterated_binders true [GLocalAssum (loc,na1,bk,t1)] b1 in + let (decls,b) = match_iterated_binders true [Loc.tag ~loc @@ GLocalAssum (na1,bk,t1)] b1 in (* TODO: address the possibility that termin is a Lambda itself *) let alp,sigma = bind_bindinglist_env alp sigma x decls in match_in u alp metas sigma b termin @@ -1003,13 +1003,13 @@ let rec match_ inner u alp metas sigma a1 a2 = (* "∀ p, let 'cp = p in t" -> "∀ 'cp, t" *) | GProd (Name p,bk,t1,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b1))]))), NBinderList (x,_,NProd (Name _id2,_,b2),(NVar v as termin)) when Id.equal p e -> - let (decls,b) = match_iterated_binders true [GLocalPattern (loc,(cp,ids),p,bk,t1)] b1 in + let (decls,b) = match_iterated_binders true [Loc.tag ~loc @@ GLocalPattern ((cp,ids),p,bk,t1)] b1 in let alp,sigma = bind_bindinglist_env alp sigma x decls in match_in u alp metas sigma b termin | GProd (na1,bk,t1,b1), NBinderList (x,_,NProd (Name _id2,_,b2),termin) when na1 != Anonymous -> - let (decls,b) = match_iterated_binders false [GLocalAssum (loc,na1,bk,t1)] b1 in + let (decls,b) = match_iterated_binders false [Loc.tag ~loc @@ GLocalAssum (na1,bk,t1)] b1 in (* TODO: address the possibility that termin is a Prod itself *) let alp,sigma = bind_bindinglist_env alp sigma x decls in match_in u alp metas sigma b termin @@ -1021,15 +1021,15 @@ let rec match_ inner u alp metas sigma a1 a2 = | GLambda (Name p,bk,t,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b1))]))), NLambda (Name id,_,b2) when is_bindinglist_meta id metas -> - let alp,sigma = bind_bindinglist_env alp sigma id [GLocalPattern (loc,(cp,ids),p,bk,t)] in + let alp,sigma = bind_bindinglist_env alp sigma id [Loc.tag ~loc @@ GLocalPattern ((cp,ids),p,bk,t)] in match_in u alp metas sigma b1 b2 | GLambda (na,bk,t,b1), NLambda (Name id,_,b2) when is_bindinglist_meta id metas -> - let alp,sigma = bind_bindinglist_env alp sigma id [GLocalAssum (loc,na,bk,t)] in + let alp,sigma = bind_bindinglist_env alp sigma id [Loc.tag ~loc @@ GLocalAssum (na,bk,t)] in match_in u alp metas sigma b1 b2 | GProd (na,bk,t,b1), NProd (Name id,_,b2) when is_bindinglist_meta id metas && na != Anonymous -> - let alp,sigma = bind_bindinglist_env alp sigma id [GLocalAssum (loc,na,bk,t)] in + let alp,sigma = bind_bindinglist_env alp sigma id [Loc.tag ~loc @@ GLocalAssum (na,bk,t)] in match_in u alp metas sigma b1 b2 (* Matching compositionally *) @@ -1121,10 +1121,10 @@ let rec match_ inner u alp metas sigma a1 a2 = | _ -> assert false in let (alp,sigma) = if is_bindinglist_meta id metas then - bind_bindinglist_env alp sigma id [GLocalAssum (Loc.ghost,Name id',Explicit,t1)] + bind_bindinglist_env alp sigma id [Loc.tag @@ GLocalAssum (Name id',Explicit,t1)] else match_names metas (alp,sigma) (Name id') na in - match_in u alp metas sigma (mkGApp Loc.ghost a1 (Loc.tag @@ GVar id')) b2 + match_in u alp metas sigma (mkGApp a1 (Loc.tag @@ GVar id')) b2 | (GRec _ | GEvar _), _ | _,_ -> raise No_match diff --git a/interp/notation_ops.mli b/interp/notation_ops.mli index a61ba172e..64f811dc2 100644 --- a/interp/notation_ops.mli +++ b/interp/notation_ops.mli @@ -33,12 +33,12 @@ val notation_constr_of_glob_constr : notation_interp_env -> (** Re-interpret a notation as a [glob_constr], taking care of binders *) -val glob_constr_of_notation_constr_with_binders : Loc.t -> +val glob_constr_of_notation_constr_with_binders : ?loc:Loc.t -> ('a -> Name.t -> 'a * Name.t) -> ('a -> notation_constr -> glob_constr) -> 'a -> notation_constr -> glob_constr -val glob_constr_of_notation_constr : Loc.t -> notation_constr -> glob_constr +val glob_constr_of_notation_constr : ?loc:Loc.t -> notation_constr -> glob_constr (** {5 Matching a notation pattern against a [glob_constr]} *) diff --git a/interp/smartlocate.ml b/interp/smartlocate.ml index 64d260cc1..fd9599ec0 100644 --- a/interp/smartlocate.ml +++ b/interp/smartlocate.ml @@ -67,14 +67,14 @@ let smart_global ?head = function | AN r -> global_with_alias ?head r | ByNotation (loc,(ntn,sc)) -> - Notation.interp_notation_as_global_reference loc (fun _ -> true) ntn sc + Notation.interp_notation_as_global_reference ~loc (fun _ -> true) ntn sc let smart_global_inductive = function | AN r -> global_inductive_with_alias r | ByNotation (loc,(ntn,sc)) -> destIndRef - (Notation.interp_notation_as_global_reference loc isIndRef ntn sc) + (Notation.interp_notation_as_global_reference ~loc isIndRef ntn sc) let loc_of_smart_reference = function | AN r -> loc_of_reference r diff --git a/interp/topconstr.ml b/interp/topconstr.ml index c8fbdaf28..2ffeb1f83 100644 --- a/interp/topconstr.ml +++ b/interp/topconstr.ml @@ -282,7 +282,7 @@ let locs_of_notation loc locs ntn = let ntn_loc loc (args,argslist,binderslist) = locs_of_notation loc (List.map constr_loc (args@List.flatten argslist)@ - List.map local_binders_loc binderslist) + List.map_filter local_binders_loc binderslist) let patntn_loc loc (args,argslist) = locs_of_notation loc diff --git a/intf/glob_term.mli b/intf/glob_term.mli index a8311e093..5aa5bdeeb 100644 --- a/intf/glob_term.mli +++ b/intf/glob_term.mli @@ -79,10 +79,11 @@ and cases_clause = (Id.t list * cases_pattern list * glob_constr) Loc.located of [t] are members of [il]. *) and cases_clauses = cases_clause list -type extended_glob_local_binder = - | GLocalAssum of Loc.t * Name.t * binding_kind * glob_constr - | GLocalDef of Loc.t * Name.t * binding_kind * glob_constr * glob_constr option - | GLocalPattern of Loc.t * (cases_pattern * Id.t list) * Id.t * binding_kind * glob_constr +type extended_glob_local_binder_r = + | GLocalAssum of Name.t * binding_kind * glob_constr + | GLocalDef of Name.t * binding_kind * glob_constr * glob_constr option + | GLocalPattern of (cases_pattern * Id.t list) * Id.t * binding_kind * glob_constr +and extended_glob_local_binder = extended_glob_local_binder_r Loc.located (** A globalised term together with a closure representing the value of its free variables. Intended for use when these variables are taken diff --git a/lib/aux_file.ml b/lib/aux_file.ml index 1b6651a55..fe6e87388 100644 --- a/lib/aux_file.ml +++ b/lib/aux_file.ml @@ -46,19 +46,21 @@ let contents x = x let empty_aux_file = H.empty -let get aux loc key = M.find key (H.find (Loc.unloc loc) aux) +let get ?loc aux key = M.find key (H.find (Option.cata Loc.unloc (0,0) loc) aux) -let record_in_aux_at loc key v = +let record_in_aux_at ?loc key v = Option.iter (fun oc -> - let i, j = Loc.unloc loc in - Printf.fprintf oc "%d %d %s %S\n" i j key v) - !oc + match loc with + | Some loc -> let i, j = Loc.unloc loc in + Printf.fprintf oc "%d %d %s %S\n" i j key v + | None -> Printf.fprintf oc "--- %s %S\n" key v + ) !oc -let current_loc = ref Loc.ghost +let current_loc : Loc.t option ref = ref None -let record_in_aux_set_at loc = current_loc := loc +let record_in_aux_set_at ?loc = current_loc := loc -let record_in_aux key v = record_in_aux_at !current_loc key v +let record_in_aux key v = record_in_aux_at ?loc:!current_loc key v let set h loc k v = let m = try H.find loc h with Not_found -> M.empty in @@ -91,4 +93,4 @@ let load_aux_file_for vfile = Flags.if_verbose Feedback.msg_info Pp.(str"Loading file "++str aux_fname++str": "++str s); empty_aux_file -let set h loc k v = set h (Loc.unloc loc) k v +let set ?loc h k v = set h (Option.cata Loc.unloc (0,0) loc) k v diff --git a/lib/aux_file.mli b/lib/aux_file.mli index 86e322b71..41b7b0855 100644 --- a/lib/aux_file.mli +++ b/lib/aux_file.mli @@ -9,9 +9,9 @@ type aux_file val load_aux_file_for : string -> aux_file -val get : aux_file -> Loc.t -> string -> string val empty_aux_file : aux_file -val set : aux_file -> Loc.t -> string -> string -> aux_file +val get : ?loc:Loc.t -> aux_file -> string -> string +val set : ?loc:Loc.t -> aux_file -> string -> string -> aux_file module H : Map.S with type key = int * int module M : Map.S with type key = string @@ -22,6 +22,6 @@ val start_aux_file : aux_file:string -> v_file:string -> unit val stop_aux_file : unit -> unit val recording : unit -> bool -val record_in_aux_at : Loc.t -> string -> string -> unit +val record_in_aux_at : ?loc:Loc.t -> string -> string -> unit val record_in_aux : string -> string -> unit -val record_in_aux_set_at : Loc.t -> unit +val record_in_aux_set_at : ?loc:Loc.t -> unit diff --git a/lib/cWarnings.ml b/lib/cWarnings.ml index 2f569d284..4e692af36 100644 --- a/lib/cWarnings.ml +++ b/lib/cWarnings.ml @@ -20,10 +20,10 @@ type t = { let warnings : (string, t) Hashtbl.t = Hashtbl.create 97 let categories : (string, string list) Hashtbl.t = Hashtbl.create 97 -let current_loc = ref Loc.ghost +let current_loc = ref None let flags = ref "" -let set_current_loc = (:=) current_loc +let set_current_loc loc = current_loc := Some loc let get_flags () = !flags @@ -35,29 +35,22 @@ let add_warning_in_category ~name ~category = in Hashtbl.replace categories category (name::ws) -let refine_loc = function - | None when not (Loc.is_ghost !current_loc) -> Some !current_loc - | loc -> loc - let create ~name ~category ?(default=Enabled) pp = Hashtbl.add warnings name { default; category; status = default }; add_warning_in_category ~name ~category; if default <> Disabled then add_warning_in_category ~name ~category:"default"; - fun ?loc x -> let w = Hashtbl.find warnings name in + fun ?loc x -> + let w = Hashtbl.find warnings name in + let loc = Option.append loc !current_loc in match w.status with | Disabled -> () - | AsError -> - begin match refine_loc loc with - | Some loc -> CErrors.user_err ~loc (pp x) - | None -> CErrors.user_err (pp x) - end + | AsError -> CErrors.user_err ?loc (pp x) | Enabled -> let msg = pp x ++ spc () ++ str "[" ++ str name ++ str "," ++ str category ++ str "]" in - let loc = refine_loc loc in Feedback.msg_warning ?loc msg let warn_unknown_warning = diff --git a/lib/feedback.ml b/lib/feedback.ml index df6fe3a62..3552fa4a0 100644 --- a/lib/feedback.ml +++ b/lib/feedback.ml @@ -30,7 +30,7 @@ type feedback_content = | FileDependency of string option * string | FileLoaded of string * string (* Extra metadata *) - | Custom of Loc.t * string * xml + | Custom of Loc.t option * string * xml (* Generic messages *) | Message of level * Loc.t option * Pp.std_ppcmds diff --git a/lib/feedback.mli b/lib/feedback.mli index bdd236ac7..dc104132a 100644 --- a/lib/feedback.mli +++ b/lib/feedback.mli @@ -38,7 +38,7 @@ type feedback_content = | FileDependency of string option * string | FileLoaded of string * string (* Extra metadata *) - | Custom of Loc.t * string * xml + | Custom of Loc.t option * string * xml (* Generic messages *) | Message of level * Loc.t option * Pp.std_ppcmds diff --git a/lib/loc.ml b/lib/loc.ml index 2a785fac4..3051ca7b9 100644 --- a/lib/loc.ml +++ b/lib/loc.ml @@ -26,12 +26,6 @@ let make_loc (bp, ep) = { fname = ""; line_nb = -1; bol_pos = 0; line_nb_last = -1; bol_pos_last = 0; bp = bp; ep = ep; } -let ghost = { - fname = ""; line_nb = -1; bol_pos = 0; line_nb_last = -1; bol_pos_last = 0; - bp = 0; ep = 0; } - -let is_ghost loc = loc.ep = 0 - let merge loc1 loc2 = if loc1.bp < loc2.bp then if loc1.ep < loc2.ep then { @@ -51,15 +45,24 @@ let merge loc1 loc2 = bp = loc2.bp; ep = loc1.ep; } else loc2 +let merge_opt l1 l2 = Option.cata (fun l1 -> merge l1 l2) l2 l1 +let opt_merge l1 l2 = Option.cata (fun l2 -> merge l1 l2) l1 l2 + let unloc loc = (loc.bp, loc.ep) -let dummy_loc = ghost let join_loc = merge (** Located type *) type 'a located = t * 'a +let is_ghost loc = loc.ep = 0 + +let ghost = { + fname = ""; line_nb = -1; bol_pos = 0; line_nb_last = -1; bol_pos_last = 0; + bp = 0; ep = 0; } + +let internal_ghost = ghost let to_pair x = x let tag ?loc x = Option.default ghost loc, x let obj (_,x) = x diff --git a/lib/loc.mli b/lib/loc.mli index 10f63a8dd..82792613c 100644 --- a/lib/loc.mli +++ b/lib/loc.mli @@ -32,14 +32,15 @@ val unloc : t -> int * int val make_loc : int * int -> t (** Make a location out of its start and end position *) -val ghost : t -(** Dummy location *) - +val internal_ghost : t val is_ghost : t -> bool (** Test whether the location is meaningful *) val merge : t -> t -> t +val merge_opt : t option -> t -> t +val opt_merge : t -> t option -> t + (** {5 Located exceptions} *) val add_loc : Exninfo.info -> t -> Exninfo.info @@ -71,15 +72,7 @@ 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 -(* Current not used *) +(* Currently not used *) val located_iter2 : ('a -> 'b -> unit) -> 'a located -> 'b located -> unit (** Projects out a located object *) - -(** {5 Backward compatibility} *) - -val dummy_loc : t -(** Same as [ghost] *) - -val join_loc : t -> t -> t -(** Same as [merge] *) diff --git a/lib/stateid.ml b/lib/stateid.ml index ae25735c5..34d49685b 100644 --- a/lib/stateid.ml +++ b/lib/stateid.ml @@ -41,7 +41,7 @@ type ('a,'b) request = { exn_info : t * t; stop : t; document : 'b; - loc : Loc.t; + loc : Loc.t option; uuid : 'a; name : string } diff --git a/lib/stateid.mli b/lib/stateid.mli index 1d87a343b..d9e75f584 100644 --- a/lib/stateid.mli +++ b/lib/stateid.mli @@ -34,7 +34,7 @@ type ('a,'b) request = { exn_info : t * t; stop : t; document : 'b; - loc : Loc.t; + loc : Loc.t option; uuid : 'a; name : string } diff --git a/library/declare.ml b/library/declare.ml index 31c9c24bc..6b505ac09 100644 --- a/library/declare.ml +++ b/library/declare.ml @@ -512,8 +512,8 @@ let do_constraint poly l = let open Misctypes in let u_of_id x = match x with - | GProp -> Loc.dummy_loc, (false, Univ.Level.prop) - | GSet -> Loc.dummy_loc, (false, Univ.Level.set) + | GProp -> Loc.tag (false, Univ.Level.prop) + | GSet -> Loc.tag (false, Univ.Level.set) | GType None -> user_err ~hdr:"Constraint" (str "Cannot declare constraints on anonymous universes") diff --git a/parsing/g_constr.ml4 b/parsing/g_constr.ml4 index 33e7236f5..e6c82c894 100644 --- a/parsing/g_constr.ml4 +++ b/parsing/g_constr.ml4 @@ -226,19 +226,19 @@ GEXTEND Gram record_field_declaration: [ [ id = global; bl = binders; ":="; c = lconstr -> - (id, mkCLambdaN (!@loc) bl c) ] ] + (id, mkCLambdaN ~loc:!@loc bl c) ] ] ; binder_constr: [ [ "forall"; bl = open_binders; ","; c = operconstr LEVEL "200" -> - mkCProdN (!@loc) bl c + mkCProdN ~loc:!@loc bl c | "fun"; bl = open_binders; "=>"; c = operconstr LEVEL "200" -> - mkCLambdaN (!@loc) bl c + mkCLambdaN ~loc:!@loc bl c | "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(c, CastConv t)) -> (Loc.tag ~loc:(constr_loc t) @@ Some t), c (* Tolerance, see G_vernac.def_body *) | _, _ -> ty, c1 in - Loc.tag ~loc:!@loc @@ CLetIn(id,mkCLambdaN (constr_loc c1) bl c1, + Loc.tag ~loc:!@loc @@ CLetIn(id,mkCLambdaN ~loc:(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 @@ -411,7 +411,7 @@ GEXTEND Gram (fun na -> CLocalAssum (na::nal,Default Implicit,c)) | nal=LIST1 name; "}" -> (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))) + Loc.tag ~loc:(Loc.merge (fst na) !@loc) @@ CHole (Some (Evar_kinds.BinderType (snd na)), IntroAnonymous, None))) | ":"; c=lconstr; "}" -> (fun na -> CLocalAssum ([na],Default Implicit,c)) ] ] diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4 index 3d0322b10..7a3a2ace1 100644 --- a/parsing/g_vernac.ml4 +++ b/parsing/g_vernac.ml4 @@ -349,14 +349,14 @@ GEXTEND Gram ; record_binder_body: [ [ l = binders; oc = of_type_with_opt_coercion; - t = lconstr -> fun id -> (oc,AssumExpr (id,mkCProdN (!@loc) l t)) + t = lconstr -> fun id -> (oc,AssumExpr (id,mkCProdN ~loc:!@loc l t)) | l = binders; oc = of_type_with_opt_coercion; t = lconstr; ":="; b = lconstr -> fun id -> - (oc,DefExpr (id,mkCLambdaN ~loc:!@loc l b,Some (mkCProdN (!@loc) l t))) + (oc,DefExpr (id,mkCLambdaN ~loc:!@loc l b,Some (mkCProdN ~loc:!@loc l t))) | l = binders; ":="; b = lconstr -> fun id -> 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:!@loc l b',Some (mkCProdN ~loc:!@loc l t))) | _ -> (None,DefExpr(id,mkCLambdaN ~loc:!@loc l b,None)) ] ] ; @@ -378,9 +378,9 @@ GEXTEND Gram constructor_type: [[ l = binders; 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 -> (not (Option.is_empty coe),(id,mkCProdN ~loc:!@loc l c)) | -> - fun l id -> (false,(id,mkCProdN (!@loc) l (Loc.tag ~loc:!@loc @@ CHole (None, Misctypes.IntroAnonymous, None)))) ] + fun l id -> (false,(id,mkCProdN ~loc:!@loc l (Loc.tag ~loc:!@loc @@ CHole (None, Misctypes.IntroAnonymous, None)))) ] -> t l ]] ; @@ -592,15 +592,15 @@ GEXTEND Gram d = def_body -> let s = coerce_reference_to_id qid in VernacDefinition - ((Some Global,CanonicalStructure),((Loc.ghost,s),None),d) + ((Some Global,CanonicalStructure),((Loc.tag s),None),d) (* Coercions *) | IDENT "Coercion"; qid = global; d = def_body -> let s = coerce_reference_to_id qid in - VernacDefinition ((None,Coercion),((Loc.ghost,s),None),d) + VernacDefinition ((None,Coercion),((Loc.tag s),None),d) | IDENT "Coercion"; IDENT "Local"; qid = global; d = def_body -> let s = coerce_reference_to_id qid in - VernacDefinition ((Some Decl_kinds.Local,Coercion),((Loc.ghost,s),None),d) + VernacDefinition ((Some Decl_kinds.Local,Coercion),((Loc.tag s),None),d) | IDENT "Identity"; IDENT "Coercion"; IDENT "Local"; f = identref; ":"; s = class_rawexpr; ">->"; t = class_rawexpr -> VernacIdentityCoercion (true, f, s, t) diff --git a/plugins/firstorder/g_ground.ml4 b/plugins/firstorder/g_ground.ml4 index 3c0319319..70748f056 100644 --- a/plugins/firstorder/g_ground.ml4 +++ b/plugins/firstorder/g_ground.ml4 @@ -61,7 +61,7 @@ let default_intuition_tac = let name = { Tacexpr.mltac_plugin = "ground_plugin"; mltac_tactic = "auto_with"; } in let entry = { Tacexpr.mltac_name = name; mltac_index = 0 } in Tacenv.register_ml_tactic name [| tac |]; - Tacexpr.TacML (Loc.ghost, entry, []) + Tacexpr.TacML (Loc.tag (entry, [])) let (set_default_solver, default_solver, print_default_solver) = Tactic_option.declare_tactic_option ~default:default_intuition_tac "Firstorder default solver" diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml index 07a0d5a50..946ee55d4 100644 --- a/plugins/funind/glob_term_to_relation.ml +++ b/plugins/funind/glob_term_to_relation.ml @@ -672,9 +672,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = in let case_pats = build_constructors_of_type (fst ind) nal_as_glob_constr in assert (Int.equal (Array.length case_pats) 1); - let br = - (Loc.ghost,([],[case_pats.(0)],e)) - in + let br = Loc.tag ([],[case_pats.(0)],e) in let match_expr = mkGCases(None,[b,(Anonymous,None)],[br]) in build_entry_lc env funnames avoid match_expr @@ -1254,7 +1252,7 @@ let rec rebuild_return_type (loc, rt) = 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], + | _ -> Loc.tag ~loc @@ Constrexpr.CProdN([[Loc.tag Anonymous], Constrexpr.Default Decl_kinds.Explicit,Loc.tag ~loc rt], Loc.tag @@ Constrexpr.CSort(GType [])) @@ -1307,12 +1305,12 @@ let do_build_inductive (fun (n,t,typ) acc -> match typ with | Some typ -> - Loc.tag @@ Constrexpr.CLetIn((Loc.ghost, n),with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t, + Loc.tag @@ Constrexpr.CLetIn((Loc.tag 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 -> Loc.tag @@ Constrexpr.CProdN - ([[(Loc.ghost,n)],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t], + ([[(Loc.tag n)],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t], acc ) ) @@ -1374,12 +1372,12 @@ let do_build_inductive (fun (n,t,typ) acc -> match typ with | Some typ -> - Loc.tag @@ Constrexpr.CLetIn((Loc.ghost, n),with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t, + Loc.tag @@ Constrexpr.CLetIn((Loc.tag 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 -> Loc.tag @@ Constrexpr.CProdN - ([[(Loc.ghost,n)],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t], + ([[(Loc.tag n)],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t], acc ) ) @@ -1406,18 +1404,18 @@ let do_build_inductive (fun (n,t,typ) -> match typ with | Some typ -> - Constrexpr.CLocalDef((Loc.ghost,n), Constrextern.extern_glob_constr Id.Set.empty t, + Constrexpr.CLocalDef((Loc.tag n), Constrextern.extern_glob_constr Id.Set.empty t, Some (with_full_print (Constrextern.extern_glob_constr Id.Set.empty) typ)) | None -> Constrexpr.CLocalAssum - ([(Loc.ghost,n)], Constrexpr_ops.default_binder_kind, Constrextern.extern_glob_constr Id.Set.empty t) + ([(Loc.tag n)], Constrexpr_ops.default_binder_kind, Constrextern.extern_glob_constr Id.Set.empty t) ) rels_params in let ext_rels_constructors = Array.map (List.map (fun (id,t) -> - false,((Loc.ghost,id), + false,((Loc.tag id), with_full_print (Constrextern.extern_glob_type Id.Set.empty) ((* zeta_normalize *) (alpha_rt rel_params_ids t)) ) @@ -1425,7 +1423,7 @@ let do_build_inductive (rel_constructors) in let rel_ind i ext_rel_constructors = - (((Loc.ghost,relnames.(i)), None), + (((Loc.tag @@ relnames.(i)), None), rel_params, Some rel_arities.(i), ext_rel_constructors),[] diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index cad96946c..1da85c467 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -156,7 +156,7 @@ let build_newrecursive let (rec_sign,rec_impls) = List.fold_left (fun (env,impls) (((_,recname),_),bl,arityc,_) -> - let arityc = Constrexpr_ops.mkCProdN Loc.ghost bl arityc in + let arityc = Constrexpr_ops.mkCProdN bl arityc in let arity,ctx = Constrintern.interp_type env0 sigma arityc in let evdref = ref (Evd.from_env env0) in let _, (_, impls') = Constrintern.interp_context_evars env evdref bl in @@ -355,7 +355,7 @@ let generate_principle (evd:Evd.evar_map ref) pconstants on_error (*i The next call to mk_rel_id is valid since we have just construct the graph Ensures by : do_built i*) - let f_R_mut = Ident (Loc.ghost,mk_rel_id (List.nth names 0)) in + let f_R_mut = Ident (Loc.tag @@ mk_rel_id (List.nth names 0)) in let ind_kn = fst (locate_with_msg (pr_reference f_R_mut++str ": Not an inductive type!") @@ -453,7 +453,7 @@ let generate_correction_proof_wf f_ref tcc_lemma_ref let register_wf ?(is_mes=false) fname rec_impls wf_rel_expr wf_arg using_lemmas args ret_type body pre_hook = - let type_of_f = Constrexpr_ops.mkCProdN Loc.ghost args ret_type in + let type_of_f = Constrexpr_ops.mkCProdN args ret_type in let rec_arg_num = let names = List.map @@ -470,7 +470,7 @@ let register_wf ?(is_mes=false) fname rec_impls wf_rel_expr wf_arg using_lemmas let unbounded_eq = let f_app_args = Loc.tag @@ Constrexpr.CAppExpl( - (None,(Ident (Loc.ghost,fname)),None) , + (None,(Ident (Loc.tag fname)),None) , (List.map (function | _,Anonymous -> assert false @@ -480,10 +480,10 @@ let register_wf ?(is_mes=false) fname rec_impls wf_rel_expr wf_arg using_lemmas ) ) in - Loc.tag @@ Constrexpr.CApp ((None,Constrexpr_ops.mkRefC (Qualid (Loc.ghost,(qualid_of_string "Logic.eq")))), + Loc.tag @@ Constrexpr.CApp ((None,Constrexpr_ops.mkRefC (Qualid (Loc.tag (qualid_of_string "Logic.eq")))), [(f_app_args,None);(body,None)]) in - let eq = Constrexpr_ops.mkCProdN Loc.ghost args unbounded_eq in + let eq = Constrexpr_ops.mkCProdN args unbounded_eq in let hook ((f_ref,_) as fconst) tcc_lemma_ref (functional_ref,_) (eq_ref,_) rec_arg_num rec_arg_type nb_args relation = try @@ -537,13 +537,13 @@ let register_mes fname rec_impls wf_mes_expr wf_rel_expr_opt wf_arg using_lemmas | None -> let ltof = let make_dir l = DirPath.make (List.rev_map Id.of_string l) in - Libnames.Qualid (Loc.ghost,Libnames.qualid_of_path + Libnames.Qualid (Loc.tag @@ Libnames.qualid_of_path (Libnames.make_path (make_dir ["Arith";"Wf_nat"]) (Id.of_string "ltof"))) in let fun_from_mes = let applied_mes = Constrexpr_ops.mkAppC(wf_mes_expr,[Constrexpr_ops.mkIdentC wf_arg]) in - Constrexpr_ops.mkLambdaC ([(Loc.ghost,Name wf_arg)],Constrexpr_ops.default_binder_kind,wf_arg_type,applied_mes) + Constrexpr_ops.mkLambdaC ([(Loc.tag @@ Name wf_arg)],Constrexpr_ops.default_binder_kind,wf_arg_type,applied_mes) in let wf_rel_from_mes = Constrexpr_ops.mkAppC(Constrexpr_ops.mkRefC ltof,[wf_arg_type;fun_from_mes]) @@ -554,7 +554,7 @@ let register_mes fname rec_impls wf_mes_expr wf_rel_expr_opt wf_arg using_lemmas let a = Names.Id.of_string "___a" in let b = Names.Id.of_string "___b" in Constrexpr_ops.mkLambdaC( - [Loc.ghost,Name a;Loc.ghost,Name b], + [Loc.tag @@ Name a;Loc.tag @@ Name b], Constrexpr.Default Explicit, wf_arg_type, Constrexpr_ops.mkAppC(wf_rel_expr, @@ -891,14 +891,14 @@ let make_graph (f_ref:global_reference) = ) in let b' = add_args (snd id) new_args b in - ((((id,None), ( Some (Loc.ghost,rec_id),CStructRec),nal_tas@bl,t,Some b'),[]):(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list)) + ((((id,None), ( Some (Loc.tag rec_id),CStructRec),nal_tas@bl,t,Some b'),[]):(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list)) ) fixexprl in l | _ -> let id = Label.to_id (con_label c) in - [(((Loc.ghost,id),None),(None,Constrexpr.CStructRec),nal_tas,t,Some b),[]] + [(((Loc.tag id),None),(None,Constrexpr.CStructRec),nal_tas,t,Some b),[]] in let mp,dp,_ = repr_con c in do_generate_principle [c,Univ.Instance.empty] error_error false false expr_list; diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml index 8f0c98624..de8dc53f1 100644 --- a/plugins/funind/indfun_common.ml +++ b/plugins/funind/indfun_common.ml @@ -106,7 +106,7 @@ let list_add_set_eq eq_fun x l = let const_of_id id = let _,princ_ref = - qualid_of_reference (Libnames.Ident (Loc.ghost,id)) + qualid_of_reference (Libnames.Ident (Loc.tag id)) in try Constrintern.locate_reference princ_ref with Not_found -> diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index 8da4f9233..e4536278c 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -274,7 +274,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes List.map (fun decl -> List.map - (fun id -> Loc.ghost, IntroNaming (IntroIdentifier id)) + (fun id -> Loc.tag @@ IntroNaming (IntroIdentifier id)) (generate_fresh_id (Id.of_string "y") ids (List.length (fst (decompose_prod_assum evd (RelDecl.get_type decl))))) ) branches diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml index 9dc1d48df..f2654b5de 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 = Loc.tag @@ CRef (Libnames.Ident (Loc.ghost,id), None) in + let ans = Loc.tag @@ CRef (Libnames.Ident (Loc.tag id), None) in let _ = ignore (Constrintern.intern_constr (Global.env()) ans) in true with e when CErrors.noncritical e -> false @@ -825,7 +825,7 @@ let merge_rec_params_and_arity prms1 prms2 shift (concl:constr) = let _ = prNamedRConstr (string_of_name nme) tp in let _ = prstr " ; " in let typ = glob_constr_to_constr_expr tp in - CLocalAssum ([(Loc.ghost,nme)], Constrexpr_ops.default_binder_kind, typ) :: acc) + CLocalAssum ([(Loc.tag nme)], Constrexpr_ops.default_binder_kind, typ) :: acc) [] params in let concl = Constrextern.extern_constr false (Global.env()) Evd.empty concl in let arity,_ = @@ -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 - Loc.tag @@ CProdN ([[(Loc.ghost,nm)],Constrexpr_ops.default_binder_kind,typ] , acc) , newenv) + Loc.tag @@ CProdN ([[(Loc.tag 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 @@ -849,12 +849,12 @@ let merge_rec_params_and_arity prms1 prms2 shift (concl:constr) = FIXME: params et cstr_expr (arity) *) let glob_constr_list_to_inductive_expr prms1 prms2 mib1 mib2 shift (rawlist:(Id.t * glob_constr) list) = - let lident = (Loc.ghost, shift.ident), None in + let lident = (Loc.tag shift.ident), None in let bindlist , cstr_expr = (* params , arities *) merge_rec_params_and_arity prms1 prms2 shift mkSet in let lcstor_expr : (bool * (lident * constr_expr)) list = List.map (* zeta_normalize t ? *) - (fun (id,t) -> false, ((Loc.ghost,id),glob_constr_to_constr_expr t)) + (fun (id,t) -> false, ((Loc.tag id),glob_constr_to_constr_expr t)) rawlist in lident , bindlist , Some cstr_expr , lcstor_expr @@ -902,7 +902,7 @@ let merge_inductive (ind1: inductive) (ind2: inductive) (* Find infos on identifier id. *) let find_Function_infos_safe (id:Id.t): Indfun_common.function_info = let kn_of_id x = - let f_ref = Libnames.Ident (Loc.ghost,x) in + let f_ref = Libnames.Ident (Loc.tag x) in locate_with_msg (str "Don't know what to do with " ++ Libnames.pr_reference f_ref) locate_constant f_ref in try find_Function_infos (kn_of_id id) diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 25daedd0e..9e469c756 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -1306,7 +1306,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp CErrors.error "\"abstract\" cannot handle existentials"; let hook _ _ = let opacity = - let na_ref = Libnames.Ident (Loc.ghost,na) in + let na_ref = Libnames.Ident (Loc.tag na) in let na_global = Smartlocate.global_with_alias na_ref in match na_global with ConstRef c -> is_opaque_constant c @@ -1556,7 +1556,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num let hook _ _ = let term_ref = Nametab.locate (qualid_of_ident term_id) in let f_ref = declare_f function_name (IsProof Lemma) arg_types term_ref in - let _ = Extraction_plugin.Table.extraction_inline true [Ident (Loc.ghost,term_id)] in + let _ = Extraction_plugin.Table.extraction_inline true [Ident (Loc.tag term_id)] in (* message "start second proof"; *) let stop = try com_eqn (List.length res_vars) equation_id functional_ref f_ref term_ref (subst_var function_name equation_lemma_type); diff --git a/plugins/ltac/coretactics.ml4 b/plugins/ltac/coretactics.ml4 index 28ff6df83..6890b3198 100644 --- a/plugins/ltac/coretactics.ml4 +++ b/plugins/ltac/coretactics.ml4 @@ -305,10 +305,9 @@ END open Tacexpr let initial_atomic () = - let dloc = Loc.ghost in let nocl = {onhyps=Some[];concl_occs=AllOccurrences} in let iter (s, t) = - let body = TacAtom (dloc, t) in + let body = TacAtom (Loc.tag t) in Tacenv.register_ltac false false (Id.of_string s) body in let () = List.iter iter @@ -323,7 +322,7 @@ let initial_atomic () = List.iter iter [ "idtac",TacId []; "fail", TacFail(TacLocal,ArgArg 0,[]); - "fresh", TacArg(dloc,TacFreshId []) + "fresh", TacArg(Loc.tag @@ TacFreshId []) ] let () = Mltop.declare_cache_obj initial_atomic "coretactics" diff --git a/plugins/ltac/evar_tactics.ml b/plugins/ltac/evar_tactics.ml index 5d3f6df03..a5d9697ae 100644 --- a/plugins/ltac/evar_tactics.ml +++ b/plugins/ltac/evar_tactics.ml @@ -82,7 +82,7 @@ let instantiate_tac_by_name id c = end let let_evar name typ = - let src = (Loc.ghost,Evar_kinds.GoalEvar) in + let src = (Loc.tag Evar_kinds.GoalEvar) in Proofview.Goal.s_enter { s_enter = begin fun gl -> let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in diff --git a/plugins/ltac/extraargs.ml4 b/plugins/ltac/extraargs.ml4 index 53b726432..aa81f148e 100644 --- a/plugins/ltac/extraargs.ml4 +++ b/plugins/ltac/extraargs.ml4 @@ -228,11 +228,11 @@ ARGUMENT EXTEND hloc | [ "in" "|-" "*" ] -> [ ConclLocation () ] | [ "in" ident(id) ] -> - [ HypLocation ((Loc.ghost,id),InHyp) ] + [ HypLocation ((Loc.tag id),InHyp) ] | [ "in" "(" "Type" "of" ident(id) ")" ] -> - [ HypLocation ((Loc.ghost,id),InHypTypeOnly) ] + [ HypLocation ((Loc.tag id),InHypTypeOnly) ] | [ "in" "(" "Value" "of" ident(id) ")" ] -> - [ HypLocation ((Loc.ghost,id),InHypValueOnly) ] + [ HypLocation ((Loc.tag id),InHypValueOnly) ] END diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4 index 232bd851f..5123d6c40 100644 --- a/plugins/ltac/extratactics.ml4 +++ b/plugins/ltac/extratactics.ml4 @@ -76,7 +76,7 @@ END let induction_arg_of_quantified_hyp = function | AnonHyp n -> None,ElimOnAnonHyp n - | NamedHyp id -> None,ElimOnIdent (Loc.ghost,id) + | NamedHyp id -> None,ElimOnIdent (Loc.tag id) (* Versions *_main must come first!! so that "1" is interpreted as a ElimOnAnonHyp and not as a "constr", and "id" is interpreted as a @@ -679,8 +679,8 @@ let hResolve id c occ t = with | Pretype_errors.PretypeError (_,_,Pretype_errors.UnsolvableImplicit _) as e -> let (e, info) = CErrors.push e in - let loc = match Loc.get_loc info with None -> Loc.ghost | Some loc -> loc in - resolve_hole (subst_hole_with_term (fst (Loc.unloc loc)) c_raw t_hole) + let loc = match Loc.get_loc info with None -> 0,0 | Some loc -> Loc.unloc loc in + resolve_hole (subst_hole_with_term (fst loc) c_raw t_hole) in let t_constr,ctx = resolve_hole (subst_var_with_hole occ id t_raw) in let t_constr = EConstr.of_constr t_constr in @@ -784,7 +784,7 @@ let case_eq_intros_rewrite x = let rec find_a_destructable_match sigma t = let cl = induction_arg_of_quantified_hyp (NamedHyp (Id.of_string "x")) in let cl = [cl, (None, None), None], None in - let dest = TacAtom (Loc.ghost, TacInductionDestruct(false, false, cl)) in + let dest = TacAtom (Loc.tag @@ TacInductionDestruct(false, false, cl)) in match EConstr.kind sigma t with | Case (_,_,x,_) when closed0 sigma x -> if isVar sigma x then diff --git a/plugins/ltac/g_ltac.ml4 b/plugins/ltac/g_ltac.ml4 index d6cbd8db1..23643c5d3 100644 --- a/plugins/ltac/g_ltac.ml4 +++ b/plugins/ltac/g_ltac.ml4 @@ -460,7 +460,7 @@ END let pr_ltac_production_item = function | Tacentries.TacTerm s -> quote (str s) -| Tacentries.TacNonTerm (_, (arg, sep), id) -> +| Tacentries.TacNonTerm (_, ((arg, sep), id)) -> let sep = match sep with | None -> mt () | Some sep -> str "," ++ spc () ++ quote (str sep) @@ -470,7 +470,7 @@ let pr_ltac_production_item = function VERNAC ARGUMENT EXTEND ltac_production_item PRINTED BY pr_ltac_production_item | [ string(s) ] -> [ Tacentries.TacTerm s ] | [ ident(nt) "(" ident(p) ltac_production_sep_opt(sep) ")" ] -> - [ Tacentries.TacNonTerm (loc, (Names.Id.to_string nt, sep), p) ] + [ Tacentries.TacNonTerm (loc, ((Names.Id.to_string nt, sep), p)) ] END VERNAC COMMAND EXTEND VernacTacticNotation diff --git a/plugins/ltac/g_obligations.ml4 b/plugins/ltac/g_obligations.ml4 index 3e6e2db60..7d4075836 100644 --- a/plugins/ltac/g_obligations.ml4 +++ b/plugins/ltac/g_obligations.ml4 @@ -50,7 +50,7 @@ module Tactic = Pltac open Pcoq -let sigref = mkRefC (Qualid (Loc.ghost, Libnames.qualid_of_string "Coq.Init.Specif.sig")) +let sigref = mkRefC (Qualid (Loc.tag @@ Libnames.qualid_of_string "Coq.Init.Specif.sig")) type 'a withtac_argtype = (Tacexpr.raw_tactic_expr option, 'a) Genarg.abstract_argument_type diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml index ae596c411..26ac3c53e 100644 --- a/plugins/ltac/pptactic.ml +++ b/plugins/ltac/pptactic.ml @@ -51,7 +51,7 @@ let pr_global x = Nametab.pr_global_env Id.Set.empty x type 'a grammar_tactic_prod_item_expr = | TacTerm of string -| TacNonTerm of Loc.t * 'a * Names.Id.t +| TacNonTerm of ('a * Names.Id.t) Loc.located type grammar_terminals = Genarg.ArgT.any Extend.user_symbol grammar_tactic_prod_item_expr list @@ -212,7 +212,7 @@ type 'a extra_genarg_printer = let rec tacarg_using_rule_token pr_gen = function | [] -> [] | TacTerm s :: l -> keyword s :: tacarg_using_rule_token pr_gen l - | TacNonTerm (_, (symb, arg), _) :: l -> + | TacNonTerm (_, ((symb, arg), _)) :: l -> pr_gen symb arg :: tacarg_using_rule_token pr_gen l let pr_tacarg_using_rule pr_gen l = @@ -252,7 +252,7 @@ type 'a extra_genarg_printer = let prods = (KNmap.find key !prnotation_tab).pptac_prods in let rec pr = function | TacTerm s -> primitive s - | TacNonTerm (_, symb, _) -> str (Printf.sprintf "(%s)" (pr_user_symbol symb)) + | TacNonTerm (_, (symb, _)) -> str (Printf.sprintf "(%s)" (pr_user_symbol symb)) in pr_sequence pr prods with Not_found -> @@ -264,8 +264,8 @@ type 'a extra_genarg_printer = let rec pack prods args = match prods, args with | [], [] -> [] | TacTerm s :: prods, args -> TacTerm s :: pack prods args - | TacNonTerm (loc, symb, id) :: prods, arg :: args -> - TacNonTerm (loc, (symb, arg), id) :: pack prods args + | TacNonTerm (loc, (symb, id)) :: prods, arg :: args -> + TacNonTerm (loc, ((symb, arg), id)) :: pack prods args | _ -> raise Not_found in let prods = pack pp.pptac_prods l in @@ -275,7 +275,7 @@ type 'a extra_genarg_printer = let pr arg = str "_" in KerName.print key ++ spc() ++ pr_sequence pr l ++ str" (* Generic printer *)" - let pr_farg prtac arg = prtac (1, Any) (TacArg (Loc.ghost, arg)) + let pr_farg prtac arg = prtac (1, Any) (TacArg (Loc.tag arg)) let is_genarg tag wit = let ArgT.Any tag = tag in @@ -331,9 +331,9 @@ type 'a extra_genarg_printer = pr_extend_gen (pr_farg prtac) let pr_raw_alias prc prlc prtac prpat lev key args = - pr_alias_gen (pr_targ (fun l a -> prtac l (TacArg (Loc.ghost, a)))) lev key args + pr_alias_gen (pr_targ (fun l a -> prtac l (TacArg (Loc.tag a)))) lev key args let pr_glob_alias prc prlc prtac prpat lev key args = - pr_alias_gen (pr_targ (fun l a -> prtac l (TacArg (Loc.ghost, a)))) lev key args + pr_alias_gen (pr_targ (fun l a -> prtac l (TacArg (Loc.tag a)))) lev key args (**********************************************************************) (* The tactic printer *) @@ -352,7 +352,7 @@ type 'a extra_genarg_printer = let pr_ltac_or_var pr = function | ArgArg x -> pr x - | ArgVar (loc,id) -> pr_with_comments loc (pr_id id) + | ArgVar (loc,id) -> pr_with_comments ~loc (pr_id id) let pr_ltac_constant kn = if !Flags.in_debugger then pr_kn kn @@ -416,7 +416,7 @@ type 'a extra_genarg_printer = let pr_as_name = function | Anonymous -> mt () - | Name id -> spc () ++ keyword "as" ++ spc () ++ pr_lident (Loc.ghost,id) + | Name id -> spc () ++ keyword "as" ++ spc () ++ pr_lident (Loc.tag id) let pr_pose_as_style prc na c = spc() ++ prc c ++ pr_as_name na @@ -507,7 +507,7 @@ type 'a extra_genarg_printer = let pr_core_destruction_arg prc prlc = function | ElimOnConstr c -> pr_with_bindings prc prlc c - | ElimOnIdent (loc,id) -> pr_with_comments loc (pr_id id) + | ElimOnIdent (loc,id) -> pr_with_comments ~loc (pr_id id) | ElimOnAnonHyp n -> int n let pr_destruction_arg prc prlc (clear_flag,h) = @@ -574,7 +574,7 @@ type 'a extra_genarg_printer = let pr_let_clause k pr (id,(bl,t)) = hov 0 (keyword k ++ spc () ++ pr_lident id ++ prlist pr_funvar bl ++ - str " :=" ++ brk (1,1) ++ pr (TacArg (Loc.ghost,t))) + str " :=" ++ brk (1,1) ++ pr (TacArg (Loc.tag t))) let pr_let_clauses recflag pr = function | hd::tl -> @@ -1037,7 +1037,7 @@ type 'a extra_genarg_printer = | TacId l -> keyword "idtac" ++ prlist (pr_arg (pr_message_token pr.pr_name)) l, latom | TacAtom (loc,t) -> - pr_with_comments loc (hov 1 (pr_atom pr strip_prod_binders tag_atom t)), ltatom + pr_with_comments ~loc (hov 1 (pr_atom pr strip_prod_binders tag_atom t)), ltatom | TacArg(_,Tacexp e) -> pr.pr_tactic (latom,E) e, latom | TacArg(_,ConstrMayEval (ConstrTerm c)) -> @@ -1051,16 +1051,16 @@ type 'a extra_genarg_printer = | TacArg(_,TacCall(loc,(f,[]))) -> pr.pr_reference f, latom | TacArg(_,TacCall(loc,(f,l))) -> - pr_with_comments loc (hov 1 ( + pr_with_comments ~loc (hov 1 ( pr.pr_reference f ++ spc () ++ prlist_with_sep spc pr_tacarg l)), lcall | TacArg (_,a) -> pr_tacarg a, latom - | TacML (loc,s,l) -> - pr_with_comments loc (pr.pr_extend 1 s l), lcall + | TacML (loc,(s,l)) -> + pr_with_comments ~loc (pr.pr_extend 1 s l), lcall | TacAlias (loc,(kn,l)) -> - pr_with_comments loc (pr.pr_alias (level_of inherited) kn l), latom + pr_with_comments ~loc (pr.pr_alias (level_of inherited) kn l), latom ) in if prec_less prec inherited then strm @@ -1078,7 +1078,7 @@ type 'a extra_genarg_printer = | TacNumgoals -> keyword "numgoals" | (TacCall _|Tacexp _ | TacGeneric _) as a -> - hov 0 (keyword "ltac:" ++ surround (pr_tac ltop (TacArg (Loc.ghost,a)))) + hov 0 (keyword "ltac:" ++ surround (pr_tac ltop (TacArg (Loc.tag a)))) in pr_tac @@ -1087,7 +1087,7 @@ type 'a extra_genarg_printer = if Int.equal n 0 then (List.rev acc, (ty,None)) else match Loc.obj ty with Glob_term.GProd(na,Explicit,a,b) -> - strip_ty (([Loc.ghost,na],(a,None))::acc) (n-1) b + strip_ty (([Loc.tag na],(a,None))::acc) (n-1) b | _ -> error "Cannot translate fix tactic: not enough products" in strip_ty [] n ty @@ -1158,7 +1158,7 @@ type 'a extra_genarg_printer = if n=0 then (List.rev acc, EConstr.of_constr ty) else match Term.kind_of_term ty with Term.Prod(na,a,b) -> - strip_ty (([Loc.ghost,na],EConstr.of_constr a)::acc) (n-1) b + strip_ty (([Loc.tag na],EConstr.of_constr a)::acc) (n-1) b | _ -> error "Cannot translate fix tactic: not enough products" in strip_ty [] n ty @@ -1253,7 +1253,7 @@ let () = wit_clause_dft_concl (pr_clauses (Some true) pr_lident) (pr_clauses (Some true) pr_lident) - (pr_clauses (Some true) (fun id -> pr_lident (Loc.ghost,id))) + (pr_clauses (Some true) (fun id -> pr_lident (Loc.tag id))) ; Genprint.register_print0 wit_constr diff --git a/plugins/ltac/pptactic.mli b/plugins/ltac/pptactic.mli index 729338fb9..23570392d 100644 --- a/plugins/ltac/pptactic.mli +++ b/plugins/ltac/pptactic.mli @@ -21,7 +21,7 @@ open Ppextend type 'a grammar_tactic_prod_item_expr = | TacTerm of string -| TacNonTerm of Loc.t * 'a * Names.Id.t +| TacNonTerm of ('a * Names.Id.t) Loc.located type 'a raw_extra_genarg_printer = (constr_expr -> std_ppcmds) -> diff --git a/plugins/ltac/profile_ltac.ml b/plugins/ltac/profile_ltac.ml index bcb28f77c..eb97a0e70 100644 --- a/plugins/ltac/profile_ltac.ml +++ b/plugins/ltac/profile_ltac.ml @@ -129,7 +129,7 @@ let to_ltacprof_results xml = let feedback_results results = Feedback.(feedback - (Custom (Loc.dummy_loc, "ltacprof_results", of_ltacprof_results results))) + (Custom (None, "ltacprof_results", of_ltacprof_results results))) (* ************** pretty printing ************************************* *) @@ -250,7 +250,7 @@ let string_of_call ck = | Tacexpr.LtacVarCall (id, t) -> Nameops.pr_id id | Tacexpr.LtacAtomCall te -> (Pptactic.pr_glob_tactic (Global.env ()) - (Tacexpr.TacAtom (Loc.ghost, te))) + (Tacexpr.TacAtom (Loc.tag te))) | Tacexpr.LtacConstrInterp (c, _) -> pr_glob_constr_env (Global.env ()) c | Tacexpr.LtacMLCall te -> diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml index 40484bfea..19c2eaf0a 100644 --- a/plugins/ltac/rewrite.ml +++ b/plugins/ltac/rewrite.ml @@ -1787,11 +1787,11 @@ let rec strategy_of_ast = function (* By default the strategy for "rewrite_db" is top-down *) -let mkappc s l = Loc.tag @@ CAppExpl ((None,(Libnames.Ident (Loc.ghost,Id.of_string s)),None),l) +let mkappc s l = Loc.tag @@ CAppExpl ((None,(Libnames.Ident (Loc.tag @@ Id.of_string s)),None),l) let declare_an_instance n s args = - (((Loc.ghost,Name n),None), Explicit, - Loc.tag @@ CAppExpl ((None, Qualid (Loc.ghost, qualid_of_string s),None), + (((Loc.tag @@ Name n),None), Explicit, + Loc.tag @@ CAppExpl ((None, Qualid (Loc.tag @@ qualid_of_string s),None), args)) let declare_instance a aeq n s = declare_an_instance n s [a;aeq] @@ -1804,17 +1804,17 @@ let anew_instance global binders instance fields = let declare_instance_refl global binders a aeq n lemma = let instance = declare_instance a aeq (add_suffix n "_Reflexive") "Coq.Classes.RelationClasses.Reflexive" in anew_instance global binders instance - [(Ident (Loc.ghost,Id.of_string "reflexivity"),lemma)] + [(Ident (Loc.tag @@ Id.of_string "reflexivity"),lemma)] let declare_instance_sym global binders a aeq n lemma = let instance = declare_instance a aeq (add_suffix n "_Symmetric") "Coq.Classes.RelationClasses.Symmetric" in anew_instance global binders instance - [(Ident (Loc.ghost,Id.of_string "symmetry"),lemma)] + [(Ident (Loc.tag @@ Id.of_string "symmetry"),lemma)] let declare_instance_trans global binders a aeq n lemma = let instance = declare_instance a aeq (add_suffix n "_Transitive") "Coq.Classes.RelationClasses.Transitive" in anew_instance global binders instance - [(Ident (Loc.ghost,Id.of_string "transitivity"),lemma)] + [(Ident (Loc.tag @@ Id.of_string "transitivity"),lemma)] let declare_relation ?(binders=[]) a aeq n refl symm trans = init_setoid (); @@ -1838,16 +1838,16 @@ let declare_relation ?(binders=[]) a aeq n refl symm trans = let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.PreOrder" in ignore( anew_instance global binders instance - [(Ident (Loc.ghost,Id.of_string "PreOrder_Reflexive"), lemma1); - (Ident (Loc.ghost,Id.of_string "PreOrder_Transitive"),lemma3)]) + [(Ident (Loc.tag @@ Id.of_string "PreOrder_Reflexive"), lemma1); + (Ident (Loc.tag @@ Id.of_string "PreOrder_Transitive"),lemma3)]) | (None, Some lemma2, Some lemma3) -> let _lemma_sym = declare_instance_sym global binders a aeq n lemma2 in let _lemma_trans = declare_instance_trans global binders a aeq n lemma3 in let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.PER" in ignore( anew_instance global binders instance - [(Ident (Loc.ghost,Id.of_string "PER_Symmetric"), lemma2); - (Ident (Loc.ghost,Id.of_string "PER_Transitive"),lemma3)]) + [(Ident (Loc.tag @@ Id.of_string "PER_Symmetric"), lemma2); + (Ident (Loc.tag @@ Id.of_string "PER_Transitive"),lemma3)]) | (Some lemma1, Some lemma2, Some lemma3) -> let _lemma_refl = declare_instance_refl global binders a aeq n lemma1 in let _lemma_sym = declare_instance_sym global binders a aeq n lemma2 in @@ -1855,9 +1855,9 @@ let declare_relation ?(binders=[]) a aeq n refl symm trans = let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.Equivalence" in ignore( anew_instance global binders instance - [(Ident (Loc.ghost,Id.of_string "Equivalence_Reflexive"), lemma1); - (Ident (Loc.ghost,Id.of_string "Equivalence_Symmetric"), lemma2); - (Ident (Loc.ghost,Id.of_string "Equivalence_Transitive"), lemma3)]) + [(Ident (Loc.tag @@ Id.of_string "Equivalence_Reflexive"), lemma1); + (Ident (Loc.tag @@ Id.of_string "Equivalence_Symmetric"), lemma2); + (Ident (Loc.tag @@ Id.of_string "Equivalence_Transitive"), lemma3)]) let cHole = Loc.tag @@ CHole (None, Misctypes.IntroAnonymous, None) @@ -1959,17 +1959,16 @@ let add_setoid global binders a aeq t n = let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.Equivalence" in ignore( anew_instance global binders instance - [(Ident (Loc.ghost,Id.of_string "Equivalence_Reflexive"), mkappc "Seq_refl" [a;aeq;t]); - (Ident (Loc.ghost,Id.of_string "Equivalence_Symmetric"), mkappc "Seq_sym" [a;aeq;t]); - (Ident (Loc.ghost,Id.of_string "Equivalence_Transitive"), mkappc "Seq_trans" [a;aeq;t])]) + [(Ident (Loc.tag @@ Id.of_string "Equivalence_Reflexive"), mkappc "Seq_refl" [a;aeq;t]); + (Ident (Loc.tag @@ Id.of_string "Equivalence_Symmetric"), mkappc "Seq_sym" [a;aeq;t]); + (Ident (Loc.tag @@ Id.of_string "Equivalence_Transitive"), mkappc "Seq_trans" [a;aeq;t])]) let make_tactic name = let open Tacexpr in - let loc = Loc.ghost in let tacpath = Libnames.qualid_of_string name in - let tacname = Qualid (loc, tacpath) in - TacArg (loc, TacCall (Loc.tag ~loc (tacname, []))) + let tacname = Qualid (Loc.tag tacpath) in + TacArg (Loc.tag @@ TacCall (Loc.tag (tacname, []))) let add_morphism_infer glob m n = init_setoid (); @@ -2012,9 +2011,9 @@ let add_morphism glob binders m s n = let poly = Flags.is_universe_polymorphism () in let instance_id = add_suffix n "_Proper" in let instance = - (((Loc.ghost,Name instance_id),None), Explicit, + (((Loc.tag @@ Name instance_id),None), Explicit, Loc.tag @@ CAppExpl ( - (None, Qualid (Loc.ghost, Libnames.qualid_of_string "Coq.Classes.Morphisms.Proper"),None), + (None, Qualid (Loc.tag @@ Libnames.qualid_of_string "Coq.Classes.Morphisms.Proper"),None), [cHole; s; m])) in let tac = Tacinterp.interp (make_tactic "add_morphism_tactic") in diff --git a/plugins/ltac/tacentries.ml b/plugins/ltac/tacentries.ml index 39121ac94..f608515aa 100644 --- a/plugins/ltac/tacentries.ml +++ b/plugins/ltac/tacentries.ml @@ -22,7 +22,7 @@ open Nameops type 'a grammar_tactic_prod_item_expr = 'a Pptactic.grammar_tactic_prod_item_expr = | TacTerm of string -| TacNonTerm of Loc.t * 'a * Names.Id.t +| TacNonTerm of ('a * Names.Id.t) Loc.located type raw_argument = string * string option type argument = Genarg.ArgT.any Extend.user_symbol @@ -178,7 +178,7 @@ let add_tactic_entry (kn, ml, tg) state = in let map = function | TacTerm s -> GramTerminal s - | TacNonTerm (loc, s, _) -> + | TacNonTerm (loc, (s, _)) -> let EntryName (typ, e) = prod_item_of_symbol tg.tacgram_level s in GramNonTerminal (loc, typ, e) in @@ -206,7 +206,7 @@ let register_tactic_notation_entry name entry = let interp_prod_item = function | TacTerm s -> TacTerm s - | TacNonTerm (loc, (nt, sep), id) -> + | TacNonTerm (loc, ((nt, sep), id)) -> let symbol = parse_user_entry nt sep in let interp s = function | None -> @@ -224,7 +224,7 @@ let interp_prod_item = function end in let symbol = interp_entry_name interp symbol in - TacNonTerm (loc, symbol, id) + TacNonTerm (loc, (symbol, id)) let make_fresh_key = let id = Summary.ref ~name:"TACTIC-NOTATION-COUNTER" 0 in @@ -300,7 +300,7 @@ let inTacticGrammar : tactic_grammar_obj -> obj = let cons_production_parameter = function | TacTerm _ -> None -| TacNonTerm (_, _, id) -> Some id +| TacNonTerm (_, (_, id)) -> Some id let add_glob_tactic_notation local ~level prods forml ids tac = let parule = { @@ -338,10 +338,10 @@ let extend_atomic_tactic name entries = in let empty_value = function | TacTerm s -> raise NonEmptyArgument - | TacNonTerm (_, symb, _) -> + | TacNonTerm (_, (symb, _)) -> let EntryName (typ, e) = prod_item_of_symbol 0 symb in let Genarg.Rawwit wit = typ in - let inj x = TacArg (Loc.ghost, TacGeneric (Genarg.in_gen typ x)) in + let inj x = TacArg (Loc.tag @@ TacGeneric (Genarg.in_gen typ x)) in let default = epsilon_value inj e in match default with | None -> raise NonEmptyArgument @@ -355,7 +355,7 @@ let extend_atomic_tactic name entries = | Some (id, args) -> let args = List.map (fun a -> Tacexp a) args in let entry = { mltac_name = name; mltac_index = i } in - let body = TacML (Loc.ghost, entry, args) in + let body = TacML (Loc.tag (entry, args)) in Tacenv.register_ltac false false (Names.Id.of_string id) body in List.iteri add_atomic entries @@ -366,12 +366,12 @@ let add_ml_tactic_notation name ~level prods = let open Tacexpr in let get_id = function | TacTerm s -> None - | TacNonTerm (_, _, id) -> Some id + | TacNonTerm (_, (_, id)) -> Some id in let ids = List.map_filter get_id prods in let entry = { mltac_name = name; mltac_index = len - i - 1 } in - let map id = Reference (Misctypes.ArgVar (Loc.ghost, id)) in - let tac = TacML (Loc.ghost, entry, List.map map ids) in + let map id = Reference (Misctypes.ArgVar (Loc.tag id)) in + let tac = TacML (Loc.tag (entry, List.map map ids)) in add_glob_tactic_notation false ~level prods true ids tac in List.iteri iter (List.rev prods); diff --git a/plugins/ltac/tacentries.mli b/plugins/ltac/tacentries.mli index 069504473..48598f7f4 100644 --- a/plugins/ltac/tacentries.mli +++ b/plugins/ltac/tacentries.mli @@ -20,7 +20,7 @@ val register_ltac : locality_flag -> Tacexpr.tacdef_body list -> unit type 'a grammar_tactic_prod_item_expr = 'a Pptactic.grammar_tactic_prod_item_expr = | TacTerm of string -| TacNonTerm of Loc.t * 'a * Names.Id.t +| TacNonTerm of ('a * Names.Id.t) Loc.located type raw_argument = string * string option (** An argument type as provided in Tactic notations, i.e. a string like diff --git a/plugins/ltac/tacexpr.mli b/plugins/ltac/tacexpr.mli index e8e99fcfe..bf760e7bb 100644 --- a/plugins/ltac/tacexpr.mli +++ b/plugins/ltac/tacexpr.mli @@ -265,7 +265,7 @@ and 'a gen_tactic_expr = | TacArg of 'a gen_tactic_arg located | TacSelect of goal_selector * 'a gen_tactic_expr (* For ML extensions *) - | TacML of Loc.t * ml_tactic_entry * 'a gen_tactic_arg list + | TacML of (ml_tactic_entry * 'a gen_tactic_arg list) Loc.located (* For syntax extensions *) | TacAlias of (KerName.t * 'a gen_tactic_arg list) Loc.located diff --git a/plugins/ltac/tacintern.ml b/plugins/ltac/tacintern.ml index df845b1a4..787a5f50b 100644 --- a/plugins/ltac/tacintern.ml +++ b/plugins/ltac/tacintern.ml @@ -72,7 +72,7 @@ let intern_name l ist = function let strict_check = ref false -let adjust_loc loc = if !strict_check then Loc.ghost else loc +let adjust_loc loc = if !strict_check then None else Some loc (* Globalize a name which must be bound -- actually just check it is bound *) let intern_hyp ist (loc,id as locid) = @@ -293,7 +293,7 @@ let intern_evaluable_reference_or_by_notation ist = function | AN r -> intern_evaluable_global_reference ist r | ByNotation (loc,(ntn,sc)) -> evaluable_of_global_reference ist.genv - (Notation.interp_notation_as_global_reference loc + (Notation.interp_notation_as_global_reference ~loc (function ConstRef _ | VarRef _ -> true | _ -> false) ntn sc) (* Globalize a reduction expression *) @@ -550,7 +550,7 @@ and intern_tactic_seq onlytac ist = function | TacAtom (loc,t) -> let lf = ref ist.ltacvars in let t = intern_atomic lf ist t in - !lf, TacAtom (adjust_loc loc, t) + !lf, TacAtom (Loc.tag ?loc:(adjust_loc loc) t) | TacFun tacfun -> ist.ltacvars, TacFun (intern_tactic_fun ist tacfun) | TacLetIn (isrec,l,u) -> let ltacvars = Id.Set.union (extract_let_names l) ist.ltacvars in @@ -627,9 +627,9 @@ and intern_tactic_seq onlytac ist = function | TacAlias (loc,(s,l)) -> let l = List.map (intern_tacarg !strict_check false ist) l in ist.ltacvars, TacAlias (Loc.tag ~loc (s,l)) - | TacML (loc,opn,l) -> + | TacML (loc,(opn,l)) -> let _ignore = Tacenv.interp_ml_tactic opn in - ist.ltacvars, TacML (adjust_loc loc, opn,List.map (intern_tacarg !strict_check false ist) l) + ist.ltacvars, TacML (loc, (opn,List.map (intern_tacarg !strict_check false ist) l)) and intern_tactic_as_arg loc onlytac ist a = match intern_tacarg !strict_check onlytac ist a with diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml index 64eda28ed..e969b86f6 100644 --- a/plugins/ltac/tacinterp.ml +++ b/plugins/ltac/tacinterp.ml @@ -201,8 +201,6 @@ end let print_top_val env v = Pptactic.pr_value Pptactic.ltop v -let dloc = Loc.ghost - let catching_error call_trace fail (e, info) = let inner_trace = Option.default [] (Exninfo.get info ltac_trace_info) @@ -326,7 +324,7 @@ let coerce_to_tactic loc id v = | _ -> fail () else fail () -let intro_pattern_of_ident id = (Loc.ghost, IntroNaming (IntroIdentifier id)) +let intro_pattern_of_ident id = (Loc.tag @@ IntroNaming (IntroIdentifier id)) let value_of_ident id = in_gen (topwit wit_intro_pattern) (intro_pattern_of_ident id) @@ -370,22 +368,22 @@ let debugging_exception_step ist signal_anomaly e pp = debugging_step ist (fun () -> pp() ++ spc() ++ str "raised the exception" ++ fnl() ++ explain_exc e) -let error_ltac_variable loc id env v s = - user_err ~loc (str "Ltac variable " ++ pr_id id ++ +let error_ltac_variable ?loc id env v s = + user_err ?loc (str "Ltac variable " ++ pr_id id ++ strbrk " is bound to" ++ spc () ++ pr_value env v ++ spc () ++ strbrk "which cannot be coerced to " ++ str s ++ str".") (* Raise Not_found if not in interpretation sign *) let try_interp_ltac_var coerce ist env (loc,id) = let v = Id.Map.find id ist.lfun in - try coerce v with CannotCoerceTo s -> error_ltac_variable loc id env v s + try coerce v with CannotCoerceTo s -> error_ltac_variable ~loc id env v s let interp_ltac_var coerce ist env locid = try try_interp_ltac_var coerce ist env locid with Not_found -> anomaly (str "Detected '" ++ Id.print (snd locid) ++ str "' as ltac var at interning time") let interp_ident ist env sigma id = - try try_interp_ltac_var (coerce_var_to_ident false env sigma) ist (Some (env,sigma)) (dloc,id) + try try_interp_ltac_var (coerce_var_to_ident false env sigma) ist (Some (env,sigma)) (Loc.tag id) with Not_found -> id (* Interprets an optional identifier, bound or fresh *) @@ -531,7 +529,7 @@ let extract_ids ids lfun = if has_type v (topwit wit_intro_pattern) then let (_, ipat) = out_gen (topwit wit_intro_pattern) v in if Id.List.mem id ids then accu - else accu @ intropattern_ids (dloc, ipat) + else accu @ intropattern_ids (Loc.tag ipat) else accu in Id.Map.fold fold lfun [] @@ -541,7 +539,7 @@ let default_fresh_id = Id.of_string "H" let interp_fresh_id ist env sigma l = let extract_ident ist env sigma id = try try_interp_ltac_var (coerce_to_ident_not_fresh env sigma) - ist (Some (env,sigma)) (dloc,id) + ist (Some (env,sigma)) (Loc.tag id) with Not_found -> id in let ids = List.map_filter (function ArgVar (_, id) -> Some id | _ -> None) l in let avoid = match TacStore.get ist.extra f_avoid_ids with @@ -977,14 +975,14 @@ let interp_binding_name ist sigma = function (* If a name is bound, it has to be a quantified hypothesis *) (* user has to use other names for variables if these ones clash with *) (* a name intented to be used as a (non-variable) identifier *) - try try_interp_ltac_var (coerce_to_quantified_hypothesis sigma) ist None(dloc,id) + try try_interp_ltac_var (coerce_to_quantified_hypothesis sigma) ist None(Loc.tag id) with Not_found -> NamedHyp id let interp_declared_or_quantified_hypothesis ist env sigma = function | AnonHyp n -> AnonHyp n | NamedHyp id -> try try_interp_ltac_var - (coerce_to_decl_or_quant_hyp env sigma) ist (Some (env,sigma)) (dloc,id) + (coerce_to_decl_or_quant_hyp env sigma) ist (Some (env,sigma)) (Loc.tag id) with Not_found -> NamedHyp id let interp_binding ist env sigma (loc,(b,c)) = @@ -1012,14 +1010,14 @@ let interp_open_constr_with_bindings ist env sigma (c,bl) = sigma, (c, bl) let loc_of_bindings = function -| NoBindings -> Loc.ghost -| ImplicitBindings l -> loc_of_glob_constr (fst (List.last l)) -| ExplicitBindings l -> fst (List.last l) +| NoBindings -> None +| ImplicitBindings l -> Some (loc_of_glob_constr (fst (List.last l))) +| ExplicitBindings l -> Some (fst (List.last l)) let interp_open_constr_with_bindings_loc ist ((c,_),bl as cb) = let loc1 = loc_of_glob_constr c in let loc2 = loc_of_bindings bl in - let loc = if Loc.is_ghost loc2 then loc1 else Loc.merge loc1 loc2 in + let loc = Loc.opt_merge loc1 loc2 in let f = { delayed = fun env sigma -> let sigma = Sigma.to_evar_map sigma in let (sigma, c) = interp_open_constr_with_bindings ist env sigma cb in @@ -1288,8 +1286,8 @@ and eval_tactic ist tac : unit Proofview.tactic = match tac with in Ftactic.run tac (fun () -> Proofview.tclUNIT ()) - | TacML (loc,opn,l) -> - push_trace (loc,LtacMLCall tac) ist >>= fun trace -> + | TacML (loc,(opn,l)) -> + push_trace (Loc.tag ~loc @@ LtacMLCall tac) ist >>= fun trace -> let ist = { ist with extra = TacStore.set ist.extra f_trace trace; } in let tac = Tacenv.interp_ml_tactic opn in let args = Ftactic.List.map_right (fun a -> interp_tacarg ist a) l in @@ -1308,7 +1306,7 @@ and force_vrec ist v : Val.t Ftactic.t = | v -> Ftactic.return (of_tacvalue v) else Ftactic.return v -and interp_ltac_reference loc' mustbetac ist r : Val.t Ftactic.t = +and interp_ltac_reference ?loc' mustbetac ist r : Val.t Ftactic.t = match r with | ArgVar (loc,id) -> let v = @@ -1322,7 +1320,7 @@ and interp_ltac_reference loc' mustbetac ist r : Val.t Ftactic.t = end | ArgArg (loc,r) -> let ids = extract_ids [] ist.lfun in - let loc_info = ((if Loc.is_ghost loc' then loc else loc'),LtacNameCall r) in + let loc_info = (Option.default loc loc',LtacNameCall r) in let extra = TacStore.set ist.extra f_avoid_ids ids in push_trace loc_info ist >>= fun trace -> let extra = TacStore.set extra f_trace trace in @@ -1333,7 +1331,7 @@ and interp_ltac_reference loc' mustbetac ist r : Val.t Ftactic.t = and interp_tacarg ist arg : Val.t Ftactic.t = match arg with | TacGeneric arg -> interp_genarg ist arg - | Reference r -> interp_ltac_reference dloc false ist r + | Reference r -> interp_ltac_reference false ist r | ConstrMayEval c -> Ftactic.s_enter { s_enter = begin fun gl -> let sigma = project gl in @@ -1342,16 +1340,16 @@ and interp_tacarg ist arg : Val.t Ftactic.t = Sigma.Unsafe.of_pair (Ftactic.return (Value.of_constr c_interp), sigma) end } | TacCall (loc,(r,[])) -> - interp_ltac_reference loc true ist r + interp_ltac_reference true ist r | TacCall (loc,(f,l)) -> let (>>=) = Ftactic.bind in - interp_ltac_reference loc true ist f >>= fun fv -> + interp_ltac_reference true ist f >>= fun fv -> Ftactic.List.map (fun a -> interp_tacarg ist a) l >>= fun largs -> interp_app loc ist fv largs | TacFreshId l -> Ftactic.enter { enter = begin fun gl -> let id = interp_fresh_id ist (pf_env gl) (project gl) l in - Ftactic.return (in_gen (topwit wit_intro_pattern) (dloc, IntroNaming (IntroIdentifier id))) + Ftactic.return (in_gen (topwit wit_intro_pattern) (Loc.tag @@ IntroNaming (IntroIdentifier id))) end } | TacPretype c -> Ftactic.s_enter { s_enter = begin fun gl -> @@ -1442,7 +1440,7 @@ and interp_letrec ist llc u = Proofview.tclUNIT () >>= fun () -> (* delay for the effects of [lref], just in case. *) let lref = ref ist.lfun in let fold accu ((_, id), b) = - let v = of_tacvalue (VRec (lref, TacArg (dloc, b))) in + let v = of_tacvalue (VRec (lref, TacArg (Loc.tag b))) in Id.Map.add id v accu in let lfun = List.fold_left fold ist.lfun llc in @@ -1768,7 +1766,7 @@ and interp_atomic ist tac : unit Proofview.tactic = (* We try to fully-typecheck the term *) let (sigma,c_interp) = interp_constr ist env sigma c in let let_tac b na c cl eqpat = - let id = Option.default (Loc.ghost,IntroAnonymous) eqpat in + let id = Option.default (Loc.tag IntroAnonymous) eqpat in let with_eq = if b then None else Some (true,id) in Tactics.letin_tac with_eq na c None cl in @@ -1780,7 +1778,7 @@ and interp_atomic ist tac : unit Proofview.tactic = else (* We try to keep the pattern structure as much as possible *) let let_pat_tac b na c cl eqpat = - let id = Option.default (Loc.ghost,IntroAnonymous) eqpat in + let id = Option.default (Loc.tag IntroAnonymous) eqpat in let with_eq = if b then None else Some (true,id) in Tactics.letin_pat_tac with_eq na c cl in @@ -2132,7 +2130,7 @@ let lift_constr_tac_to_ml_tac vars tac = let c = Id.Map.find id ist.lfun in try Some (coerce_to_closed_constr env c) with CannotCoerceTo ty -> - error_ltac_variable Loc.ghost dummy_id (Some (env,sigma)) c ty + error_ltac_variable dummy_id (Some (env,sigma)) c ty in let args = List.map_filter map vars in tac args ist diff --git a/plugins/ltac/tacinterp.mli b/plugins/ltac/tacinterp.mli index 1e5f6bd42..6cd5a63b3 100644 --- a/plugins/ltac/tacinterp.mli +++ b/plugins/ltac/tacinterp.mli @@ -111,7 +111,7 @@ val interp_int : interp_sign -> Id.t Loc.located -> int val interp_int_or_var : interp_sign -> int or_var -> int -val error_ltac_variable : Loc.t -> Id.t -> +val error_ltac_variable : ?loc:Loc.t -> Id.t -> (Environ.env * Evd.evar_map) option -> value -> string -> 'a (** Transforms a constr-expecting tactic into a tactic finding its arguments in diff --git a/plugins/ltac/tacsubst.ml b/plugins/ltac/tacsubst.ml index c92dd23a0..0ee6e8a85 100644 --- a/plugins/ltac/tacsubst.ml +++ b/plugins/ltac/tacsubst.ml @@ -234,7 +234,7 @@ and subst_tactic subst (t:glob_tactic_expr) = match t with | TacAlias (_,(s,l)) -> let s = subst_kn subst s in TacAlias (Loc.tag (s,List.map (subst_tacarg subst) l)) - | TacML (_loc,opn,l) -> TacML (_loc, opn,List.map (subst_tacarg subst) l) + | TacML (loc,(opn,l)) -> TacML (loc, (opn,List.map (subst_tacarg subst) l)) and subst_tactic_fun subst (var,body) = (var,subst_tactic subst body) diff --git a/plugins/ltac/tactic_debug.ml b/plugins/ltac/tactic_debug.ml index dffeade29..ac8534bdc 100644 --- a/plugins/ltac/tactic_debug.ml +++ b/plugins/ltac/tactic_debug.ml @@ -352,7 +352,7 @@ let explain_ltac_call_trace last trace loc = Pptactic.pr_glob_tactic (Global.env()) t ++ str ")" | Tacexpr.LtacAtomCall te -> quote (Pptactic.pr_glob_tactic (Global.env()) - (Tacexpr.TacAtom (Loc.ghost,te))) + (Tacexpr.TacAtom (Loc.tag te))) | Tacexpr.LtacConstrInterp (c, { Pretyping.ltac_constrs = vars }) -> quote (Printer.pr_glob_constr_env (Global.env()) c) ++ (if not (Id.Map.is_empty vars) then @@ -389,14 +389,15 @@ let skip_extensions trace = let finer_loc loc1 loc2 = Loc.merge loc1 loc2 = loc2 -let extract_ltac_trace trace eloc = +let extract_ltac_trace ?loc trace = let trace = skip_extensions trace in - let (loc,c),tail = List.sep_last trace in + let (tloc,c),tail = List.sep_last trace in + let loc = Option.default tloc loc in if is_defined_ltac trace then (* We entered a user-defined tactic, we display the trace with location of the call *) - let msg = hov 0 (explain_ltac_call_trace c tail eloc ++ fnl()) in - Some msg, if finer_loc eloc loc then eloc else loc + let msg = hov 0 (explain_ltac_call_trace c tail loc ++ fnl()) in + (if finer_loc loc tloc then loc else tloc), Some msg else (* We entered a primitive tactic, we don't display trace but report on the finest location *) @@ -411,14 +412,14 @@ let extract_ltac_trace trace eloc = else aux best_loc tail | [] -> best_loc in - aux eloc trace in - None, best_loc + aux loc trace in + best_loc, None let get_ltac_trace (_, info) = let ltac_trace = Exninfo.get info ltac_trace_info in - let loc = Option.default Loc.ghost (Loc.get_loc info) in + let loc = Loc.get_loc info in match ltac_trace with | None -> None - | Some trace -> Some (extract_ltac_trace trace loc) + | Some trace -> Some (extract_ltac_trace ?loc trace) let () = ExplainErr.register_additional_error_info get_ltac_trace diff --git a/plugins/ltac/tactic_debug.mli b/plugins/ltac/tactic_debug.mli index 7745d9b7b..38d8caca6 100644 --- a/plugins/ltac/tactic_debug.mli +++ b/plugins/ltac/tactic_debug.mli @@ -78,4 +78,4 @@ val db_breakpoint : debug_info -> Id.t Loc.located message_token list -> unit Proofview.NonLogical.t val extract_ltac_trace : - Tacexpr.ltac_trace -> Loc.t -> Pp.std_ppcmds option * Loc.t + ?loc:Loc.t -> Tacexpr.ltac_trace -> Pp.std_ppcmds option Loc.located diff --git a/plugins/ltac/tauto.ml b/plugins/ltac/tauto.ml index 30eed08d7..df186cc46 100644 --- a/plugins/ltac/tauto.ml +++ b/plugins/ltac/tauto.ml @@ -89,7 +89,6 @@ let _ = (** Base tactics *) -let loc = Loc.ghost let idtac = Proofview.tclUNIT () let fail = Proofview.tclINDEPENDENT (tclFAIL 0 (Pp.mt ())) @@ -207,7 +206,7 @@ let u_iff = make_unfold "iff" let u_not = make_unfold "not" let reduction_not_iff _ ist = - let make_reduce c = TacAtom (loc, TacReduce (Genredexpr.Unfold c, Locusops.allHypsAndConcl)) in + let make_reduce c = TacAtom (Loc.tag @@ TacReduce (Genredexpr.Unfold c, Locusops.allHypsAndConcl)) in let tac = match !negation_unfolding, unfold_iff () with | true, true -> make_reduce [u_not; u_iff] | true, false -> make_reduce [u_not] @@ -260,11 +259,11 @@ let tauto_power_flags = { } let with_flags flags _ ist = - let f = (loc, Id.of_string "f") in - let x = (loc, Id.of_string "x") in + let f = (Loc.tag @@ Id.of_string "f") in + let x = (Loc.tag @@ Id.of_string "x") in let arg = Val.Dyn (tag_tauto_flags, flags) in let ist = { ist with lfun = Id.Map.add (snd x) arg ist.lfun } in - eval_tactic_ist ist (TacArg (loc, TacCall (loc, (ArgVar f, [Reference (ArgVar x)])))) + eval_tactic_ist ist (TacArg (Loc.tag @@ TacCall (Loc.tag (ArgVar f, [Reference (ArgVar x)])))) let register_tauto_tactic tac name0 args = let ids = List.map (fun id -> Id.of_string id) args in @@ -272,7 +271,7 @@ let register_tauto_tactic tac name0 args = let name = { mltac_plugin = tauto_plugin; mltac_tactic = name0; } in let entry = { mltac_name = name; mltac_index = 0 } in let () = Tacenv.register_ml_tactic name [| tac |] in - let tac = TacFun (ids, TacML (loc, entry, [])) in + let tac = TacFun (ids, TacML (Loc.tag (entry, []))) in let obj () = Tacenv.register_ltac true true (Id.of_string name0) tac in Mltop.declare_cache_obj obj tauto_plugin diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml index 4b87e6e2e..79d17db25 100644 --- a/plugins/micromega/coq_micromega.ml +++ b/plugins/micromega/coq_micromega.ml @@ -1997,7 +1997,7 @@ let micromega_gen let intro_vars = Tacticals.New.tclTHENLIST (List.map intro vars) in let intro_props = Tacticals.New.tclTHENLIST (List.map intro props) in - let ipat_of_name id = Some (Loc.ghost, Misctypes.IntroNaming (Misctypes.IntroIdentifier id)) in + let ipat_of_name id = Some (Loc.tag @@ Misctypes.IntroNaming (Misctypes.IntroIdentifier id)) in let goal_name = Tactics.fresh_id [] (Names.Id.of_string "__arith") gl in let env' = List.map (fun (id,i) -> Term.mkVar id,i) vars in @@ -2113,7 +2113,7 @@ let micromega_genr prover tac = let intro_vars = Tacticals.New.tclTHENLIST (List.map intro vars) in let intro_props = Tacticals.New.tclTHENLIST (List.map intro props) in - let ipat_of_name id = Some (Loc.ghost, Misctypes.IntroNaming (Misctypes.IntroIdentifier id)) in + let ipat_of_name id = Some (Loc.tag @@ Misctypes.IntroNaming (Misctypes.IntroIdentifier id)) in let goal_name = Tactics.fresh_id [] (Names.Id.of_string "__arith") gl in let env' = List.map (fun (id,i) -> Term.mkVar id,i) vars in diff --git a/plugins/quote/g_quote.ml4 b/plugins/quote/g_quote.ml4 index 432625e4d..980f03db3 100644 --- a/plugins/quote/g_quote.ml4 +++ b/plugins/quote/g_quote.ml4 @@ -19,15 +19,14 @@ open Tacarg DECLARE PLUGIN "quote_plugin" -let loc = Loc.ghost let cont = Id.of_string "cont" let x = Id.of_string "x" let make_cont (k : Val.t) (c : EConstr.t) = let c = Tacinterp.Value.of_constr c in - let tac = TacCall (loc, (ArgVar (loc, cont), [Reference (ArgVar (loc, x))])) in + let tac = TacCall (Loc.tag (ArgVar (Loc.tag cont), [Reference (ArgVar (Loc.tag x))])) in let ist = { lfun = Id.Map.add cont k (Id.Map.singleton x c); extra = TacStore.empty; } in - Tacinterp.eval_tactic_ist ist (TacArg (loc, tac)) + Tacinterp.eval_tactic_ist ist (TacArg (Loc.tag tac)) TACTIC EXTEND quote [ "quote" ident(f) ] -> [ quote f [] ] diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index e3e14bcf3..e2a6ad55c 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -128,11 +128,11 @@ let closed_term_ast l = mltac_name = tacname; mltac_index = 0; } in - let l = List.map (fun gr -> ArgArg(Loc.ghost,gr)) l in + let l = List.map (fun gr -> ArgArg(Loc.tag gr)) l in TacFun([Name(Id.of_string"t")], - TacML(Loc.ghost,tacname, + TacML(Loc.tag (tacname, [TacGeneric (Genarg.in_gen (Genarg.glbwit Stdarg.wit_constr) (Loc.tag @@ GVar(Id.of_string"t"),None)); - TacGeneric (Genarg.in_gen (Genarg.glbwit (Genarg.wit_list Stdarg.wit_ref)) l)])) + TacGeneric (Genarg.in_gen (Genarg.glbwit (Genarg.wit_list Stdarg.wit_ref)) l)]))) (* let _ = add_tacdef false ((Loc.ghost,Id.of_string"ring_closed_term" *) @@ -170,7 +170,7 @@ let ltac_lcall tac args = let ltac_apply (f : Value.t) (args: Tacinterp.Value.t list) = let fold arg (i, vars, lfun) = let id = Id.of_string ("x" ^ string_of_int i) in - let x = Reference (ArgVar (Loc.ghost, id)) in + let x = Reference (ArgVar (Loc.tag id)) in (succ i, x :: vars, Id.Map.add id arg lfun) in let (_, args, lfun) = List.fold_right fold args (0, [], Id.Map.empty) in @@ -205,7 +205,7 @@ let get_res = let exec_tactic env evd n f args = let fold arg (i, vars, lfun) = let id = Id.of_string ("x" ^ string_of_int i) in - let x = Reference (ArgVar (Loc.ghost, id)) in + let x = Reference (ArgVar (Loc.tag id)) in (succ i, x :: vars, Id.Map.add id (Value.of_constr arg) lfun) in let (_, args, lfun) = List.fold_right fold args (0, [], Id.Map.empty) in @@ -213,7 +213,7 @@ let exec_tactic env evd n f args = (** Build the getter *) let lid = List.init n (fun i -> Id.of_string("x"^string_of_int i)) in let n = Genarg.in_gen (Genarg.glbwit Stdarg.wit_int) n in - let get_res = TacML (Loc.ghost, get_res, [TacGeneric n]) in + let get_res = TacML (Loc.tag (get_res, [TacGeneric n])) in let getter = Tacexp (TacFun (List.map (fun n -> Name n) lid, get_res)) in (** Evaluate the whole result *) let gl = dummy_goal env evd in @@ -603,7 +603,7 @@ let interp_power env evd pow = match pow with | None -> let t = ArgArg(Loc.tag (Lazy.force ltac_inv_morph_nothing)) in - (TacArg(Loc.ghost,TacCall(Loc.tag (t,[]))), plapp evd coq_None [|carrier|]) + (TacArg(Loc.tag (TacCall(Loc.tag (t,[])))), plapp evd coq_None [|carrier|]) | Some (tac, spec) -> let tac = match tac with diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index f8cccf714..a6a6bd6f9 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -151,12 +151,12 @@ let dC t = CastConv t 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 = 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) +let mkCHole ~loc = Loc.tag ?loc @@ CHole (None, IntroAnonymous, None) +let mkCLambda ?loc name ty t = Loc.tag ?loc @@ + CLambdaN ([[Loc.tag ?loc name], Default Explicit, ty], t) +let mkCLetIn ?loc name bo t = Loc.tag ?loc @@ + CLetIn ((Loc.tag ?loc name), bo, None, t) +let mkCCast ?loc t ty = Loc.tag ?loc @@ CCast (t, dC ty) (** Constructors for rawconstr *) let mkRHole = Loc.tag @@ GHole (InternalHole, IntroAnonymous, None) let mkRApp f args = if args = [] then f else Loc.tag @@ GApp (f, args) @@ -943,8 +943,8 @@ let glob_cpattern gs p = let name = Name (id_of_string ("_ssrpat_" ^ s)) in k, (mkRCast mkRHole (mkRLambda name mkRHole (mkRApp mkRHole l)), None) in let bind_in t1 t2 = - let d = Loc.ghost in let n = Name (destCVar t1) in - fst (glob (mkCCast d (mkCHole d) (mkCLambda d n (mkCHole d) t2))) in + let mkCHole = mkCHole ~loc:None in let n = Name (destCVar t1) in + fst (glob (mkCCast mkCHole (mkCLambda n mkCHole t2))) in let check_var t2 = if not (isCVar t2) then loc_error (constr_loc t2) "Only identifiers are allowed here" in match p with @@ -1187,27 +1187,27 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty = pp(lazy(str"decoded as: " ++ pr_pattern_w_ids red)); let red = match redty with None -> red | Some ty -> let ty = ' ', ty in match red with - | T t -> T (combineCG t ty (mkCCast (loc_ofCG t)) mkRCast) + | T t -> T (combineCG t ty (mkCCast ~loc:(loc_ofCG t)) mkRCast) | X_In_T (x,t) -> let ty = pf_intern_term ist gl ty in E_As_X_In_T (mkG (mkRCast mkRHole ty), x, t) | E_In_X_In_T (e,x,t) -> let ty = mkG (pf_intern_term ist gl ty) in - E_In_X_In_T (combineCG e ty (mkCCast (loc_ofCG t)) mkRCast, x, t) + E_In_X_In_T (combineCG e ty (mkCCast ~loc:(loc_ofCG t)) mkRCast, x, t) | E_As_X_In_T (e,x,t) -> let ty = mkG (pf_intern_term ist gl ty) in - E_As_X_In_T (combineCG e ty (mkCCast (loc_ofCG t)) mkRCast, x, t) + E_As_X_In_T (combineCG e ty (mkCCast ~loc:(loc_ofCG t)) mkRCast, x, t) | red -> red in pp(lazy(str"typed as: " ++ pr_pattern_w_ids red)); - let mkXLetIn loc x (a,(g,c)) = match c with - | Some b -> a,(g,Some (mkCLetIn loc x (mkCHole loc) b)) - | None -> a,(Loc.tag ~loc @@ GLetIn (x, Loc.tag ~loc @@ GHole (BinderType x, IntroAnonymous, None), None, g), None) in + let mkXLetIn ?loc x (a,(g,c)) = match c with + | Some b -> a,(g,Some (mkCLetIn ?loc x (mkCHole ?loc) b)) + | None -> a,(Loc.tag ?loc @@ GLetIn (x, Loc.tag ?loc @@ GHole (BinderType x, IntroAnonymous, None), None, g), None) in match red with | T t -> let sigma, t = interp_term ist gl t in sigma, T t | In_T t -> let sigma, t = interp_term ist gl t in sigma, In_T t | X_In_T (x, rp) | In_X_In_T (x, rp) -> let mk x p = match red with X_In_T _ -> X_In_T(x,p) | _ -> In_X_In_T(x,p) in - let rp = mkXLetIn Loc.ghost (Name x) rp in + let rp = mkXLetIn (Name x) rp in let sigma, rp = interp_term ist gl rp in let _, h, _, rp = destLetIn rp in let sigma = cleanup_XinE h x rp sigma in @@ -1216,7 +1216,7 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty = | E_In_X_In_T(e, x, rp) | E_As_X_In_T (e, x, rp) -> let mk e x p = match red with E_In_X_In_T _ ->E_In_X_In_T(e,x,p)|_->E_As_X_In_T(e,x,p) in - let rp = mkXLetIn Loc.ghost (Name x) rp in + let rp = mkXLetIn (Name x) rp in let sigma, rp = interp_term ist gl rp in let _, h, _, rp = destLetIn rp in let sigma = cleanup_XinE h x rp sigma in @@ -1426,7 +1426,7 @@ let () = let () = Tacenv.register_ml_tactic name [|mltac|] in let tac = TacFun ([Name (Id.of_string "pattern")], - TacML (Loc.ghost, { mltac_name = name; mltac_index = 0 }, [])) in + TacML (Loc.tag ({ mltac_name = name; mltac_index = 0 }, []))) in let obj () = Tacenv.register_ltac true false (Id.of_string "ssrpattern") tac in Mltop.declare_cache_obj obj "ssrmatching_plugin" diff --git a/plugins/syntax/ascii_syntax.ml b/plugins/syntax/ascii_syntax.ml index dc0b87793..ed977c416 100644 --- a/plugins/syntax/ascii_syntax.ml +++ b/plugins/syntax/ascii_syntax.ml @@ -37,24 +37,24 @@ let glob_Ascii = lazy (make_reference "Ascii") open Lazy -let interp_ascii loc p = +let interp_ascii ?loc p = let rec aux n p = if Int.equal n 0 then [] else let mp = p mod 2 in - (Loc.tag ~loc @@ GRef ((if Int.equal mp 0 then glob_false else glob_true),None)) + (Loc.tag ?loc @@ GRef ((if Int.equal mp 0 then glob_false else glob_true),None)) :: (aux (n-1) (p/2)) in - Loc.tag ~loc @@ GApp (Loc.tag ~loc @@ GRef(force glob_Ascii,None), aux 8 p) + Loc.tag ?loc @@ GApp (Loc.tag ?loc @@ GRef(force glob_Ascii,None), aux 8 p) -let interp_ascii_string dloc s = +let interp_ascii_string ?loc s = let p = if Int.equal (String.length s) 1 then int_of_char s.[0] else if Int.equal (String.length s) 3 && is_digit s.[0] && is_digit s.[1] && is_digit s.[2] then int_of_string s else - user_err ~loc:dloc ~hdr:"interp_ascii_string" + user_err ?loc ~hdr:"interp_ascii_string" (str "Expects a single character or a three-digits ascii code.") in - interp_ascii dloc p + interp_ascii ?loc p let uninterp_ascii r = let rec uninterp_bool_list n = function diff --git a/plugins/syntax/nat_syntax.ml b/plugins/syntax/nat_syntax.ml index 90d643b7f..5cdd82024 100644 --- a/plugins/syntax/nat_syntax.ml +++ b/plugins/syntax/nat_syntax.ml @@ -33,21 +33,21 @@ let warn_large_nat = strbrk "may vary from 5000 to 70000 depending on your system " ++ strbrk "limits and on the command executed).") -let nat_of_int loc n = +let nat_of_int ?loc n = if is_pos_or_zero n then begin if less_than threshold n then warn_large_nat (); - let ref_O = Loc.tag ~loc @@ GRef (glob_O, None) in - let ref_S = Loc.tag ~loc @@ GRef (glob_S, None) in + let ref_O = Loc.tag ?loc @@ GRef (glob_O, None) in + let ref_S = Loc.tag ?loc @@ GRef (glob_S, None) in let rec mk_nat acc n = if n <> zero then - mk_nat (Loc.tag ~loc @@ GApp (ref_S, [acc])) (sub_1 n) + mk_nat (Loc.tag ?loc @@ GApp (ref_S, [acc])) (sub_1 n) else acc in mk_nat ref_O n end else - user_err ~hdr:"nat_of_int" + user_err ?loc ~hdr:"nat_of_int" (str "Cannot interpret a negative number as a number of type nat") (************************************************************************) diff --git a/plugins/syntax/numbers_syntax.ml b/plugins/syntax/numbers_syntax.ml index 8876d464a..3ee64ba7e 100644 --- a/plugins/syntax/numbers_syntax.ml +++ b/plugins/syntax/numbers_syntax.ml @@ -86,10 +86,10 @@ exception Non_closed (* parses a *non-negative* integer (from bigint.ml) into an int31 wraps modulo 2^31 *) -let int31_of_pos_bigint loc n = - let ref_construct = Loc.tag ~loc @@ GRef (int31_construct, None) in - let ref_0 = Loc.tag ~loc @@ GRef (int31_0, None) in - let ref_1 = Loc.tag ~loc @@ GRef (int31_1, None) in +let int31_of_pos_bigint ?loc n = + let ref_construct = Loc.tag ?loc @@ GRef (int31_construct, None) in + let ref_0 = Loc.tag ?loc @@ GRef (int31_0, None) in + let ref_1 = Loc.tag ?loc @@ GRef (int31_1, None) in let rec args counter n = if counter <= 0 then [] @@ -97,16 +97,16 @@ let int31_of_pos_bigint loc n = let (q,r) = div2_with_rest n in (if r then ref_1 else ref_0)::(args (counter-1) q) in - Loc.tag ~loc @@ GApp (ref_construct, List.rev (args 31 n)) + Loc.tag ?loc @@ GApp (ref_construct, List.rev (args 31 n)) -let error_negative dloc = - CErrors.user_err ~loc:dloc ~hdr:"interp_int31" (Pp.str "int31 are only non-negative numbers.") +let error_negative ?loc = + CErrors.user_err ?loc ~hdr:"interp_int31" (Pp.str "int31 are only non-negative numbers.") -let interp_int31 dloc n = +let interp_int31 ?loc n = if is_pos_or_zero n then - int31_of_pos_bigint dloc n + int31_of_pos_bigint ?loc n else - error_negative dloc + error_negative ?loc (* Pretty prints an int31 *) @@ -162,40 +162,40 @@ let height bi = in hght 0 base (* n must be a non-negative integer (from bigint.ml) *) -let word_of_pos_bigint loc hght n = - let ref_W0 = Loc.tag ~loc @@ GRef (zn2z_W0, None) in - let ref_WW = Loc.tag ~loc @@ GRef (zn2z_WW, None) in +let word_of_pos_bigint ?loc hght n = + let ref_W0 = Loc.tag ?loc @@ GRef (zn2z_W0, None) in + let ref_WW = Loc.tag ?loc @@ GRef (zn2z_WW, None) in let rec decomp hgt n = if hgt <= 0 then - int31_of_pos_bigint loc n + int31_of_pos_bigint ?loc n else if equal n zero then - Loc.tag ~loc @@ GApp (ref_W0, [Loc.tag ~loc @@ GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None)]) + Loc.tag ?loc @@ GApp (ref_W0, [Loc.tag ?loc @@ GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None)]) else let (h,l) = split_at hgt n in - Loc.tag ~loc @@ GApp (ref_WW, [Loc.tag ~loc @@ GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None); + Loc.tag ?loc @@ GApp (ref_WW, [Loc.tag ?loc @@ GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None); decomp (hgt-1) h; decomp (hgt-1) l]) in decomp hght n -let bigN_of_pos_bigint loc n = +let bigN_of_pos_bigint ?loc n = let h = height n in - let ref_constructor = Loc.tag ~loc @@ GRef (bigN_constructor h, None) in - let word = word_of_pos_bigint loc h n in + let ref_constructor = Loc.tag ?loc @@ GRef (bigN_constructor h, None) in + let word = word_of_pos_bigint ?loc h n in let args = if h < n_inlined then [word] - else [Nat_syntax_plugin.Nat_syntax.nat_of_int loc (of_int (h-n_inlined));word] + else [Nat_syntax_plugin.Nat_syntax.nat_of_int ?loc (of_int (h-n_inlined));word] in - Loc.tag ~loc @@ GApp (ref_constructor, args) + Loc.tag ?loc @@ GApp (ref_constructor, args) -let bigN_error_negative loc = - CErrors.user_err ~loc ~hdr:"interp_bigN" (Pp.str "bigN are only non-negative numbers.") +let bigN_error_negative ?loc = + CErrors.user_err ?loc ~hdr:"interp_bigN" (Pp.str "bigN are only non-negative numbers.") -let interp_bigN dloc n = +let interp_bigN ?loc n = if is_pos_or_zero n then - bigN_of_pos_bigint dloc n + bigN_of_pos_bigint ?loc n else - bigN_error_negative dloc + bigN_error_negative ?loc (* Pretty prints a bigN *) @@ -256,13 +256,13 @@ let _ = Notation.declare_numeral_interpreter bigN_scope (*** Parsing for bigZ in digital notation ***) -let interp_bigZ loc n = - let ref_pos = Loc.tag ~loc @@ GRef (bigZ_pos, None) in - let ref_neg = Loc.tag ~loc @@ GRef (bigZ_neg, None) in +let interp_bigZ ?loc n = + let ref_pos = Loc.tag ?loc @@ GRef (bigZ_pos, None) in + let ref_neg = Loc.tag ?loc @@ GRef (bigZ_neg, None) in if is_pos_or_zero n then - Loc.tag ~loc @@ GApp (ref_pos, [bigN_of_pos_bigint loc n]) + Loc.tag ?loc @@ GApp (ref_pos, [bigN_of_pos_bigint ?loc n]) else - Loc.tag ~loc @@ GApp (ref_neg, [bigN_of_pos_bigint loc (neg n)]) + Loc.tag ?loc @@ GApp (ref_neg, [bigN_of_pos_bigint ?loc (neg n)]) (* pretty printing functions for bigZ *) let bigint_of_bigZ = function @@ -292,9 +292,9 @@ let _ = Notation.declare_numeral_interpreter bigZ_scope true) (*** Parsing for bigQ in digital notation ***) -let interp_bigQ loc n = - let ref_z = Loc.tag ~loc @@ GRef (bigQ_z, None) in - Loc.tag ~loc @@ GApp (ref_z, [interp_bigZ loc n]) +let interp_bigQ ?loc n = + let ref_z = Loc.tag ?loc @@ GRef (bigQ_z, None) in + Loc.tag ?loc @@ GApp (ref_z, [interp_bigZ ?loc n]) let uninterp_bigQ rc = try match rc with diff --git a/plugins/syntax/r_syntax.ml b/plugins/syntax/r_syntax.ml index 1af3f6c5b..b7041d045 100644 --- a/plugins/syntax/r_syntax.ml +++ b/plugins/syntax/r_syntax.ml @@ -41,7 +41,7 @@ let glob_xI = ConstructRef path_of_xI let glob_xO = ConstructRef path_of_xO let glob_xH = ConstructRef path_of_xH -let pos_of_bignat dloc x = +let pos_of_bignat ?loc x = let ref_xI = Loc.tag @@ GRef (glob_xI, None) in let ref_xH = Loc.tag @@ GRef (glob_xH, None) in let ref_xO = Loc.tag @@ GRef (glob_xO, None) in @@ -77,11 +77,11 @@ let glob_ZERO = ConstructRef path_of_ZERO let glob_POS = ConstructRef path_of_POS let glob_NEG = ConstructRef path_of_NEG -let z_of_int dloc n = +let z_of_int ?loc n = if not (Bigint.equal n zero) then let sgn, n = if is_pos_or_zero n then glob_POS, n else glob_NEG, Bigint.neg n in - Loc.tag @@ GApp(Loc.tag @@ GRef (sgn,None), [pos_of_bignat dloc n]) + Loc.tag @@ GApp(Loc.tag @@ GRef (sgn,None), [pos_of_bignat ?loc n]) else Loc.tag @@ GRef (glob_ZERO, None) @@ -107,8 +107,8 @@ let make_path dir id = Globnames.encode_con dir (Id.of_string id) let glob_IZR = ConstRef (make_path (make_dir rdefinitions) "IZR") -let r_of_int dloc z = - Loc.tag @@ GApp (Loc.tag @@ GRef(glob_IZR,None), [z_of_int dloc z]) +let r_of_int ?loc z = + Loc.tag @@ GApp (Loc.tag @@ GRef(glob_IZR,None), [z_of_int ?loc z]) (**********************************************************************) (* Printing R via scopes *) @@ -128,6 +128,6 @@ let uninterp_r p = let _ = Notation.declare_numeral_interpreter "R_scope" (r_path,["Coq";"Reals";"Rdefinitions"]) r_of_int - ([Loc.tag @@ GRef (glob_IZR,None)], + ([Loc.tag @@ GRef (glob_IZR, None)], uninterp_r, false) diff --git a/plugins/syntax/string_syntax.ml b/plugins/syntax/string_syntax.ml index 539670722..49cb9355c 100644 --- a/plugins/syntax/string_syntax.ml +++ b/plugins/syntax/string_syntax.ml @@ -33,12 +33,12 @@ let glob_EmptyString = lazy (make_reference "EmptyString") open Lazy -let interp_string loc s = +let interp_string ?loc s = let le = String.length s in let rec aux n = - if n = le then Loc.tag ~loc @@ GRef (force glob_EmptyString, None) else - Loc.tag ~loc @@ GApp (Loc.tag ~loc @@ GRef (force glob_String, None), - [interp_ascii loc (int_of_char s.[n]); aux (n+1)]) + if n = le then Loc.tag ?loc @@ GRef (force glob_EmptyString, None) else + Loc.tag ?loc @@ GApp (Loc.tag ?loc @@ GRef (force glob_String, None), + [interp_ascii ?loc (int_of_char s.[n]); aux (n+1)]) in aux 0 let uninterp_string r = diff --git a/plugins/syntax/z_syntax.ml b/plugins/syntax/z_syntax.ml index a00525f91..96c1f3e39 100644 --- a/plugins/syntax/z_syntax.ml +++ b/plugins/syntax/z_syntax.ml @@ -44,25 +44,25 @@ let glob_xI = ConstructRef path_of_xI let glob_xO = ConstructRef path_of_xO let glob_xH = ConstructRef path_of_xH -let pos_of_bignat loc x = - let ref_xI = Loc.tag ~loc @@ GRef (glob_xI, None) in - let ref_xH = Loc.tag ~loc @@ GRef (glob_xH, None) in - let ref_xO = Loc.tag ~loc @@ GRef (glob_xO, None) in +let pos_of_bignat ?loc x = + let ref_xI = Loc.tag ?loc @@ GRef (glob_xI, None) in + let ref_xH = Loc.tag ?loc @@ GRef (glob_xH, None) in + let ref_xO = Loc.tag ?loc @@ GRef (glob_xO, None) in let rec pos_of x = match div2_with_rest x with - | (q,false) -> Loc.tag ~loc @@ GApp (ref_xO,[pos_of q]) - | (q,true) when not (Bigint.equal q zero) -> Loc.tag ~loc @@ GApp (ref_xI,[pos_of q]) + | (q,false) -> Loc.tag ?loc @@ GApp (ref_xO,[pos_of q]) + | (q,true) when not (Bigint.equal q zero) -> Loc.tag ?loc @@ GApp (ref_xI,[pos_of q]) | (q,true) -> ref_xH in pos_of x -let error_non_positive dloc = - user_err ~loc:dloc ~hdr:"interp_positive" +let error_non_positive ?loc = + user_err ?loc ~hdr:"interp_positive" (str "Only strictly positive numbers in type \"positive\".") -let interp_positive dloc n = - if is_strictly_pos n then pos_of_bignat dloc n - else error_non_positive dloc +let interp_positive ?loc n = + if is_strictly_pos n then pos_of_bignat ?loc n + else error_non_positive ?loc (**********************************************************************) (* Printing positive via scopes *) @@ -106,18 +106,18 @@ let glob_Npos = ConstructRef path_of_Npos let n_path = make_path binnums "N" -let n_of_binnat loc pos_or_neg n = Loc.tag ~loc @@ +let n_of_binnat ?loc pos_or_neg n = Loc.tag ?loc @@ if not (Bigint.equal n zero) then - GApp(Loc.tag @@ GRef (glob_Npos,None), [pos_of_bignat loc n]) + GApp(Loc.tag @@ GRef (glob_Npos,None), [pos_of_bignat ?loc n]) else GRef(glob_N0, None) -let error_negative dloc = - user_err ~loc:dloc ~hdr:"interp_N" (str "No negative numbers in type \"N\".") +let error_negative ?loc = + user_err ?loc ~hdr:"interp_N" (str "No negative numbers in type \"N\".") -let n_of_int dloc n = - if is_pos_or_zero n then n_of_binnat dloc true n - else error_negative dloc +let n_of_int ?loc n = + if is_pos_or_zero n then n_of_binnat ?loc true n + else error_negative ?loc (**********************************************************************) (* Printing N via scopes *) @@ -157,13 +157,13 @@ let glob_ZERO = ConstructRef path_of_ZERO let glob_POS = ConstructRef path_of_POS let glob_NEG = ConstructRef path_of_NEG -let z_of_int loc n = +let z_of_int ?loc n = if not (Bigint.equal n zero) then let sgn, n = if is_pos_or_zero n then glob_POS, n else glob_NEG, Bigint.neg n in - Loc.tag ~loc @@ GApp(Loc.tag ~loc @@ GRef(sgn,None), [pos_of_bignat loc n]) + Loc.tag ?loc @@ GApp(Loc.tag ?loc @@ GRef(sgn,None), [pos_of_bignat ?loc n]) else - Loc.tag ~loc @@ GRef(glob_ZERO, None) + Loc.tag ?loc @@ GRef(glob_ZERO, None) (**********************************************************************) (* Printing Z via scopes *) diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 347c49f44..f5e2e52a1 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -122,7 +122,7 @@ type 'a equation = { patterns : cases_pattern list; rhs : 'a rhs; alias_stack : Name.t list; - eqn_loc : Loc.t; + eqn_loc : Loc.t option; used : bool ref } type 'a matrix = 'a equation list @@ -251,7 +251,7 @@ type 'a pattern_matching_problem = tomatch : tomatch_stack; history : pattern_continuation; mat : 'a matrix; - caseloc : Loc.t; + caseloc : Loc.t option; casestyle : case_style; typing_function: type_constraint -> env -> evar_map ref -> 'a option -> unsafe_judgment } @@ -281,8 +281,8 @@ let inductive_template evdref env tmloc ind = let arsign = inductive_alldecls_env env indu in let indu = on_snd EInstance.make indu in let hole_source i = match tmloc with - | Some loc -> (loc, Evar_kinds.TomatchTypeParameter (ind,i)) - | None -> (Loc.ghost, Evar_kinds.TomatchTypeParameter (ind,i)) in + | Some loc -> Loc.tag ~loc @@ Evar_kinds.TomatchTypeParameter (ind,i) + | None -> Loc.tag @@ Evar_kinds.TomatchTypeParameter (ind,i) in let (_,evarl,_) = List.fold_right (fun decl (subst,evarl,n) -> @@ -351,7 +351,7 @@ let coerce_row typing_fun evdref env pats (tomatch,(_,indopt)) = let loc = loc_of_glob_constr tomatch in let tycon,realnames = find_tomatch_tycon evdref env (Some loc) indopt in let j = typing_fun tycon env evdref tomatch in - let evd, j = Coercion.inh_coerce_to_base (loc_of_glob_constr tomatch) env !evdref j in + let evd, j = Coercion.inh_coerce_to_base ~loc:(loc_of_glob_constr tomatch) env !evdref j in evdref := evd; let typ = nf_evar !evdref j.uj_type in let t = @@ -370,7 +370,7 @@ let coerce_to_indtype typing_fun evdref env matx tomatchl = (************************************************************************) (* Utils *) -let mkExistential env ?(src=(Loc.ghost,Evar_kinds.InternalHole)) evdref = +let mkExistential env ?(src=(Loc.tag Evar_kinds.InternalHole)) evdref = let e, u = e_new_type_evar env evdref univ_flexible_alg ~src:src in e let evd_comb2 f evdref x y = @@ -402,7 +402,7 @@ let adjust_tomatch_to_pattern pb ((current,typ),deps,dep) = let _ = e_cumul pb.env pb.evdref indt typ in current else - (evd_comb2 (Coercion.inh_conv_coerce_to true Loc.ghost pb.env) + (evd_comb2 (Coercion.inh_conv_coerce_to true pb.env) pb.evdref (make_judge current typ) indt).uj_val in let sigma = !(pb.evdref) in (current,try_find_ind pb.env sigma indt names)) @@ -469,11 +469,11 @@ let prepend_pattern tms eqn = {eqn with patterns = tms@eqn.patterns } exception NotAdjustable -let rec adjust_local_defs loc = function +let rec adjust_local_defs ?loc = function | (pat :: pats, LocalAssum _ :: decls) -> - pat :: adjust_local_defs loc (pats,decls) + pat :: adjust_local_defs ?loc (pats,decls) | (pats, LocalDef _ :: decls) -> - (Loc.tag ~loc @@ PatVar Anonymous) :: adjust_local_defs loc (pats,decls) + (Loc.tag ?loc @@ PatVar Anonymous) :: adjust_local_defs ?loc (pats,decls) | [], [] -> [] | _ -> raise NotAdjustable @@ -489,14 +489,14 @@ let check_and_adjust_constructor env ind cstrs = function if Int.equal (List.length args) nb_args_constr then pat else try - let args' = adjust_local_defs loc (args, List.rev ci.cs_args) + let args' = adjust_local_defs ~loc (args, List.rev ci.cs_args) in Loc.tag ~loc @@ PatCstr (cstr, args', alias) with NotAdjustable -> error_wrong_numarg_constructor ~loc env cstr nb_args_constr else (* Try to insert a coercion *) try - Coercion.inh_pattern_coerce_to loc env pat ind' ind + Coercion.inh_pattern_coerce_to ~loc env pat ind' ind with Not_found -> error_bad_constructor ~loc env cstr ind @@ -510,7 +510,7 @@ let check_all_variables env sigma typ mat = let check_unused_pattern env eqn = if not !(eqn.used) then - raise_pattern_matching_error ~loc:eqn.eqn_loc (env, Evd.empty, UnusedClause eqn.patterns) + raise_pattern_matching_error ?loc:eqn.eqn_loc (env, Evd.empty, UnusedClause eqn.patterns) let set_used_pattern eqn = eqn.used := true @@ -978,7 +978,7 @@ let add_assert_false_case pb tomatch = avoid_ids = []; it = None }; alias_stack = Anonymous::aliasnames; - eqn_loc = Loc.ghost; + eqn_loc = None; used = ref false } ] let adjust_impossible_cases pb pred tomatch submat = @@ -1545,7 +1545,7 @@ let matx_of_eqns env eqns = it = Some initial_rhs } in { patterns = initial_lpat; alias_stack = []; - eqn_loc = loc; + eqn_loc = Some loc; used = ref false; rhs = rhs } in List.map build_eqn eqns @@ -1629,11 +1629,11 @@ let rec list_assoc_in_triple x = function * similarly for each ti. *) -let abstract_tycon loc env evdref subst tycon extenv t = +let abstract_tycon ?loc env evdref subst tycon extenv t = let t = nf_betaiota !evdref t in (* it helps in some cases to remove K-redex*) let src = match EConstr.kind !evdref t with - | Evar (evk,_) -> (loc,Evar_kinds.SubEvar evk) - | _ -> (loc,Evar_kinds.CasesType true) in + | Evar (evk,_) -> (Loc.tag ?loc @@ Evar_kinds.SubEvar evk) + | _ -> (Loc.tag ?loc @@ Evar_kinds.CasesType true) in let subst0,t0 = adjust_to_extended_env_and_remove_deps env extenv !evdref subst t in (* We traverse the type T of the original problem Xi looking for subterms that match the non-constructor part of the constraints (this part @@ -1687,7 +1687,7 @@ let abstract_tycon loc env evdref subst tycon extenv t = in aux (0,extenv,subst0) t0 -let build_tycon loc env tycon_env s subst tycon extenv evdref t = +let build_tycon ?loc env tycon_env s subst tycon extenv evdref t = let t,tt = match t with | None -> (* This is the situation we are building a return predicate and @@ -1695,10 +1695,10 @@ let build_tycon loc env tycon_env s subst tycon extenv evdref t = let n = Context.Rel.length (rel_context env) in let n' = Context.Rel.length (rel_context tycon_env) in let impossible_case_type, u = - e_new_type_evar (reset_context env) evdref univ_flexible_alg ~src:(loc,Evar_kinds.ImpossibleCase) in + e_new_type_evar (reset_context env) evdref univ_flexible_alg ~src:(Loc.tag ?loc Evar_kinds.ImpossibleCase) in (lift (n'-n) impossible_case_type, mkSort u) | Some t -> - let t = abstract_tycon loc tycon_env evdref subst tycon extenv t in + let t = abstract_tycon ?loc tycon_env evdref subst tycon extenv t in let evd,tt = Typing.type_of extenv !evdref t in evdref := evd; (t,tt) in @@ -1786,7 +1786,7 @@ let build_inversion_problem loc env sigma tms t = let main_eqn = { patterns = patl; alias_stack = []; - eqn_loc = Loc.ghost; + eqn_loc = None; used = ref false; rhs = { rhs_env = pb_env; (* we assume all vars are used; in practice we discard dependent @@ -1806,7 +1806,7 @@ let build_inversion_problem loc env sigma tms t = else [ { patterns = List.map (fun _ -> Loc.tag @@ PatVar Anonymous) patl; alias_stack = []; - eqn_loc = Loc.ghost; + eqn_loc = None; used = ref false; rhs = { rhs_env = pb_env; rhs_vars = []; @@ -1827,7 +1827,7 @@ let build_inversion_problem loc env sigma tms t = mat = main_eqn :: catch_all_eqn; caseloc = loc; casestyle = RegularStyle; - typing_function = build_tycon loc env pb_env s subst} in + typing_function = build_tycon ?loc env pb_env s subst} in let pred = (compile pb).uj_val in (!evdref,pred) @@ -1880,10 +1880,10 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign = | _ -> assert false in List.rev (buildrec 0 (tomatchl,tmsign)) -let inh_conv_coerce_to_tycon loc env evdref j tycon = +let inh_conv_coerce_to_tycon ?loc env evdref j tycon = match tycon with | Some p -> - let (evd',j) = Coercion.inh_conv_coerce_to true loc env !evdref j p in + let (evd',j) = Coercion.inh_conv_coerce_to ?loc true env !evdref j p in evdref := evd'; j | None -> j @@ -1966,7 +1966,7 @@ let noccur_with_meta sigma n m term = in try (occur_rec n term; true) with LocalOccur -> false -let prepare_predicate loc typing_fun env sigma tomatchs arsign tycon pred = +let prepare_predicate ?loc typing_fun env sigma tomatchs arsign tycon pred = let refresh_tycon sigma t = (** If we put the typing constraint in the term, it has to be refreshed to preserve the invariant that no algebraic universe @@ -1997,7 +1997,7 @@ let prepare_predicate loc typing_fun env sigma tomatchs arsign tycon pred = | None -> let sigma = Sigma.Unsafe.of_evar_map sigma in let Sigma ((t, _), sigma, _) = - new_type_evar env sigma univ_flexible_alg ~src:(loc, Evar_kinds.CasesType false) in + new_type_evar env sigma univ_flexible_alg ~src:(Loc.tag ?loc @@ Evar_kinds.CasesType false) in let sigma = Sigma.to_evar_map sigma in sigma, t in @@ -2438,7 +2438,7 @@ let context_of_arsign l = l ([], 0) in x -let compile_program_cases loc style (typing_function, evdref) tycon env +let compile_program_cases ?loc style (typing_function, evdref) tycon env (predopt, tomatchl, eqns) = let typing_fun tycon env = function | Some t -> typing_function tycon env evdref t @@ -2545,9 +2545,9 @@ let compile_program_cases loc style (typing_function, evdref) tycon env (**************************************************************************) (* Main entry of the matching compilation *) -let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, eqns) = +let compile_cases ?loc style (typing_fun, evdref) tycon env (predopt, tomatchl, eqns) = if predopt == None && Flags.is_program_mode () && Program.is_program_cases () then - compile_program_cases loc style (typing_fun, evdref) + compile_program_cases ?loc style (typing_fun, evdref) tycon env (predopt, tomatchl, eqns) else @@ -2564,7 +2564,7 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e with the type of arguments to match; if none is provided, we build alternative possible predicates *) let arsign = extract_arity_signature env tomatchs tomatchl in - let preds = prepare_predicate loc typing_fun env !evdref tomatchs arsign tycon predopt in + let preds = prepare_predicate ?loc typing_fun env !evdref tomatchs arsign tycon predopt in let compile_for_one_predicate (sigma,nal,pred) = (* We push the initial terms to match and push their alias to rhs' envs *) @@ -2614,7 +2614,7 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e let j = compile pb in (* We coerce to the tycon (if an elim predicate was provided) *) - let j = inh_conv_coerce_to_tycon loc env myevdref j tycon in + let j = inh_conv_coerce_to_tycon ?loc env myevdref j tycon in evdref := !myevdref; j in diff --git a/pretyping/cases.mli b/pretyping/cases.mli index 6c2b5bf68..b16342db4 100644 --- a/pretyping/cases.mli +++ b/pretyping/cases.mli @@ -38,7 +38,7 @@ val irrefutable : env -> cases_pattern -> bool (** {6 Compilation primitive. } *) val compile_cases : - Loc.t -> case_style -> + ?loc:Loc.t -> case_style -> (type_constraint -> env -> evar_map ref -> glob_constr -> unsafe_judgment) * evar_map ref -> type_constraint -> env -> glob_constr option * tomatch_tuples * cases_clauses -> @@ -65,7 +65,7 @@ type 'a equation = { patterns : cases_pattern list; rhs : 'a rhs; alias_stack : Name.t list; - eqn_loc : Loc.t; + eqn_loc : Loc.t option; used : bool ref } type 'a matrix = 'a equation list @@ -106,14 +106,14 @@ type 'a pattern_matching_problem = tomatch : tomatch_stack; history : pattern_continuation; mat : 'a matrix; - caseloc : Loc.t; + caseloc : Loc.t option; casestyle : case_style; typing_function: type_constraint -> env -> evar_map ref -> 'a option -> unsafe_judgment } val compile : 'a pattern_matching_problem -> unsafe_judgment -val prepare_predicate : Loc.t -> +val prepare_predicate : ?loc:Loc.t -> (Evarutil.type_constraint -> Environ.env -> Evd.evar_map ref -> glob_constr -> unsafe_judgment) -> Environ.env -> diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index b2c1d0116..acccfc125 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -75,25 +75,25 @@ let apply_coercion_args env evd check isproj argl funj = !evdref, res (* appliquer le chemin de coercions de patterns p *) -let apply_pattern_coercion loc pat p = +let apply_pattern_coercion ?loc pat p = List.fold_left (fun pat (co,n) -> let f i = - if i (None, v) in aux t -and coerce loc env evdref (x : EConstr.constr) (y : EConstr.constr) +and coerce ?loc env evdref (x : EConstr.constr) (y : EConstr.constr) : (EConstr.constr -> EConstr.constr) option = let open Context.Rel.Declaration in @@ -183,7 +183,7 @@ and coerce loc env evdref (x : EConstr.constr) (y : EConstr.constr) let args = List.rev (restargs @ mkRel 1 :: List.map (lift 1) tele) in let pred = mkLambda (n, eqT, applist (lift 1 c, args)) in let eq = papp evdref coq_eq_ind [| eqT; hdx; hdy |] in - let evar = make_existential loc env evdref eq in + let evar = make_existential ?loc env evdref eq in let eq_app x = papp evdref coq_eq_rect [| eqT; hdx; pred; x; hdy; evar|] in @@ -326,7 +326,7 @@ and coerce loc env evdref (x : EConstr.constr) (y : EConstr.constr) Some (fun x -> let cx = app_opt env evdref c x in - let evar = make_existential loc env evdref (mkApp (p, [| cx |])) + let evar = make_existential ?loc env evdref (mkApp (p, [| cx |])) in (papp evdref sig_intro [| u; p; cx; evar |])) | None -> @@ -340,9 +340,9 @@ let app_coercion env evdref coercion v = let v' = Typing.e_solve_evars env evdref (f v) in whd_betaiota !evdref v' -let coerce_itf loc env evd v t c1 = +let coerce_itf ?loc env evd v t c1 = let evdref = ref evd in - let coercion = coerce loc env evdref t c1 in + let coercion = coerce ?loc env evdref t c1 in let t = Option.map (app_coercion env evdref coercion) v in !evdref, t @@ -410,16 +410,16 @@ let type_judgment env sigma j = | Sort s -> {utj_val = j.uj_val; utj_type = ESorts.kind sigma s } | _ -> error_not_a_type env sigma j -let inh_tosort_force loc env evd j = +let inh_tosort_force ?loc env evd j = try let t,p = lookup_path_to_sort_from env evd j.uj_type in let evd,j1 = apply_coercion env evd p j t in let j2 = on_judgment_type (whd_evar evd) j1 in (evd,type_judgment env evd j2) with Not_found | NoCoercion -> - error_not_a_type ~loc env evd j + error_not_a_type ?loc env evd j -let inh_coerce_to_sort loc env evd j = +let inh_coerce_to_sort ?loc env evd j = let typ = whd_all env evd j.uj_type in match EConstr.kind evd typ with | Sort s -> (evd,{ utj_val = j.uj_val; utj_type = ESorts.kind evd s }) @@ -427,9 +427,9 @@ let inh_coerce_to_sort loc env evd j = let (evd',s) = Evardefine.define_evar_as_sort env evd ev in (evd',{ utj_val = j.uj_val; utj_type = s }) | _ -> - inh_tosort_force loc env evd j + inh_tosort_force ?loc env evd j -let inh_coerce_to_base loc env evd j = +let inh_coerce_to_base ?loc env evd j = if Flags.is_program_mode () then let evdref = ref evd in let ct, typ' = mu env evdref j.uj_type in @@ -439,7 +439,7 @@ let inh_coerce_to_base loc env evd j = in !evdref, res else (evd, j) -let inh_coerce_to_prod loc env evd t = +let inh_coerce_to_prod ?loc env evd t = if Flags.is_program_mode () then let evdref = ref evd in let _, typ' = mu env evdref t in @@ -466,7 +466,7 @@ let inh_coerce_to_fail env evd rigidonly v t c1 = try (the_conv_x_leq env t' c1 evd, v') with UnableToUnify _ -> raise NoCoercion -let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = +let rec inh_conv_coerce_to_fail ?loc env evd rigidonly v t c1 = try (the_conv_x_leq env t c1 evd, v) with UnableToUnify (best_failed_evd,e) -> try inh_coerce_to_fail env evd rigidonly v t c1 @@ -488,49 +488,50 @@ let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = let open Context.Rel.Declaration in let env1 = push_rel (LocalAssum (name,u1)) env in let (evd', v1) = - inh_conv_coerce_to_fail loc env1 evd rigidonly + inh_conv_coerce_to_fail ?loc env1 evd rigidonly (Some (mkRel 1)) (lift 1 u1) (lift 1 t1) in let v1 = Option.get v1 in let v2 = Option.map (fun v -> beta_applist evd' (lift 1 v,[v1])) v in let t2 = match v2 with | None -> subst_term evd' v1 t2 | Some v2 -> Retyping.get_type_of env1 evd' v2 in - let (evd'',v2') = inh_conv_coerce_to_fail loc env1 evd' rigidonly v2 t2 u2 in + let (evd'',v2') = inh_conv_coerce_to_fail ?loc env1 evd' rigidonly v2 t2 u2 in (evd'', Option.map (fun v2' -> mkLambda (name, u1, v2')) v2') | _ -> raise (NoCoercionNoUnifier (best_failed_evd,e)) (* Look for cj' obtained from cj by inserting coercions, s.t. cj'.typ = t *) -let inh_conv_coerce_to_gen resolve_tc rigidonly loc env evd cj t = +let inh_conv_coerce_to_gen ?loc resolve_tc rigidonly env evd cj t = let (evd', val') = try - inh_conv_coerce_to_fail loc env evd rigidonly (Some cj.uj_val) cj.uj_type t + inh_conv_coerce_to_fail ?loc env evd rigidonly (Some cj.uj_val) cj.uj_type t with NoCoercionNoUnifier (best_failed_evd,e) -> try if Flags.is_program_mode () then - coerce_itf loc env evd (Some cj.uj_val) cj.uj_type t + coerce_itf ?loc env evd (Some cj.uj_val) cj.uj_type t else raise NoSubtacCoercion with | NoSubtacCoercion when not resolve_tc || not !use_typeclasses_for_conversion -> - error_actual_type ~loc env best_failed_evd cj t e + error_actual_type ?loc env best_failed_evd cj t e | NoSubtacCoercion -> let evd' = saturate_evd env evd in try if evd' == evd then - error_actual_type ~loc env best_failed_evd cj t e + error_actual_type ?loc env best_failed_evd cj t e else - inh_conv_coerce_to_fail loc env evd' rigidonly (Some cj.uj_val) cj.uj_type t + inh_conv_coerce_to_fail ?loc env evd' rigidonly (Some cj.uj_val) cj.uj_type t with NoCoercionNoUnifier (_evd,_error) -> - error_actual_type ~loc env best_failed_evd cj t e + error_actual_type ?loc env best_failed_evd cj t e in let val' = match val' with Some v -> v | None -> assert(false) in (evd',{ uj_val = val'; uj_type = t }) -let inh_conv_coerce_to resolve_tc = inh_conv_coerce_to_gen resolve_tc false -let inh_conv_coerce_rigid_to resolve_tc = inh_conv_coerce_to_gen resolve_tc true +let inh_conv_coerce_to ?loc resolve_tc = inh_conv_coerce_to_gen ?loc resolve_tc false -let inh_conv_coerces_to loc env evd t t' = +let inh_conv_coerce_rigid_to ?loc resolve_tc = inh_conv_coerce_to_gen resolve_tc ?loc true + +let inh_conv_coerces_to ?loc env evd t t' = try - fst (inh_conv_coerce_to_fail loc env evd true None t t') + fst (inh_conv_coerce_to_fail ?loc env evd true None t t') with NoCoercion -> evd (* Maybe not enough information to unify *) diff --git a/pretyping/coercion.mli b/pretyping/coercion.mli index bc63d092d..25ee82bbf 100644 --- a/pretyping/coercion.mli +++ b/pretyping/coercion.mli @@ -26,17 +26,17 @@ val inh_app_fun : bool -> (** [inh_coerce_to_sort env isevars j] coerces [j] to a type; i.e. it inserts a coercion into [j], if needed, in such a way it gets as type a sort; it fails if no coercion is applicable *) -val inh_coerce_to_sort : Loc.t -> +val inh_coerce_to_sort : ?loc:Loc.t -> env -> evar_map -> unsafe_judgment -> evar_map * unsafe_type_judgment (** [inh_coerce_to_base env isevars j] coerces [j] to its base type; i.e. it inserts a coercion into [j], if needed, in such a way it gets as type its base type (the notion depends on the coercion system) *) -val inh_coerce_to_base : Loc.t -> +val inh_coerce_to_base : ?loc:Loc.t -> env -> evar_map -> unsafe_judgment -> evar_map * unsafe_judgment (** [inh_coerce_to_prod env isevars t] coerces [t] to a product type *) -val inh_coerce_to_prod : Loc.t -> +val inh_coerce_to_prod : ?loc:Loc.t -> env -> evar_map -> types -> evar_map * types (** [inh_conv_coerce_to resolve_tc Loc.t env isevars j t] coerces [j] to an @@ -44,20 +44,20 @@ val inh_coerce_to_prod : Loc.t -> a way [t] and [j.uj_type] are convertible; it fails if no coercion is applicable. resolve_tc=false disables resolving type classes (as the last resort before failing) *) -val inh_conv_coerce_to : bool -> Loc.t -> +val inh_conv_coerce_to : ?loc:Loc.t -> bool -> env -> evar_map -> unsafe_judgment -> types -> evar_map * unsafe_judgment -val inh_conv_coerce_rigid_to : bool -> Loc.t -> +val inh_conv_coerce_rigid_to : ?loc:Loc.t -> bool -> env -> evar_map -> unsafe_judgment -> types -> evar_map * unsafe_judgment (** [inh_conv_coerces_to loc env isevars t t'] checks if an object of type [t] is coercible to an object of type [t'] adding evar constraints if needed; it fails if no coercion exists *) -val inh_conv_coerces_to : Loc.t -> +val inh_conv_coerces_to : ?loc:Loc.t -> env -> evar_map -> types -> types -> evar_map (** [inh_pattern_coerce_to loc env isevars pat ind1 ind2] coerces the Cases pattern [pat] typed in [ind1] into a pattern typed in [ind2]; raises [Not_found] if no coercion found *) val inh_pattern_coerce_to : - Loc.t -> env -> cases_pattern -> inductive -> inductive -> cases_pattern + ?loc:Loc.t -> env -> cases_pattern -> inductive -> inductive -> cases_pattern diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index 05d6a1ad4..3079d1052 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -29,8 +29,6 @@ open Misctypes open Decl_kinds open Context.Named.Declaration -let dl = Loc.ghost - (** Should we keep details of universes during detyping ? *) let print_universes = Flags.univ_print diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 4bb66b8e9..95680e18a 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -888,7 +888,7 @@ and conv_record trs env evd (ctx,(h,h2),c,bs,(params,params1),(us,us2),(sk1,sk2) let test i = evar_conv_x trs env i CUMUL ty (substl ks b) in (i,t2::ks, m-1, test) else - let dloc = (Loc.ghost,Evar_kinds.InternalHole) in + let dloc = Loc.tag Evar_kinds.InternalHole in let i = Sigma.Unsafe.of_evar_map i in let Sigma (ev, i', _) = Evarutil.new_evar env i ~src:dloc (substl ks b) in let i' = Sigma.to_evar_map i' in @@ -1214,7 +1214,7 @@ let apply_conversion_problem_heuristic ts env evd pbty t1 t2 = let error_cannot_unify env evd pb ?reason t1 t2 = Pretype_errors.error_cannot_unify - ~loc:(loc_of_conv_pb evd pb) env + ?loc:(loc_of_conv_pb evd pb) env evd ?reason (t1, t2) let check_problems_are_solved env evd = diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml index 25ece5b8e..cba1780ba 100644 --- a/pretyping/glob_ops.ml +++ b/pretyping/glob_ops.ml @@ -22,7 +22,7 @@ let cases_predicate_names tml = | (tm,(na,None)) -> [na] | (tm,(na,Some (_,(_,nal)))) -> na::nal) tml) -let mkGApp loc p t = Loc.tag ~loc @@ +let mkGApp ?loc p t = Loc.tag ?loc @@ match snd p with | GApp (f,l) -> GApp (f,l@[t]) | _ -> GApp (p,[t]) diff --git a/pretyping/glob_ops.mli b/pretyping/glob_ops.mli index 55e6b6533..ac2118ba7 100644 --- a/pretyping/glob_ops.mli +++ b/pretyping/glob_ops.mli @@ -25,7 +25,7 @@ val cases_pattern_loc : cases_pattern -> Loc.t val cases_predicate_names : tomatch_tuples -> Name.t list (** Apply one argument to a glob_constr *) -val mkGApp : Loc.t -> glob_constr -> glob_constr -> glob_constr +val mkGApp : ?loc:Loc.t -> glob_constr -> glob_constr -> glob_constr val map_glob_constr : (glob_constr -> glob_constr) -> glob_constr -> glob_constr diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 5f9f4bb08..fe15cb490 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -133,7 +133,7 @@ let nf_fix sigma (nas, cs, ts) = let inj c = EConstr.to_constr sigma c in (nas, Array.map inj cs, Array.map inj ts) -let search_guard loc env possible_indexes fixdefs = +let search_guard ?loc env possible_indexes fixdefs = (* Standard situation with only one possibility for each fix. *) (* We treat it separately in order to get proper error msg. *) let is_singleton = function [_] -> true | _ -> false in @@ -143,7 +143,7 @@ let search_guard loc env possible_indexes fixdefs = (try check_fix env fix with reraise -> let (e, info) = CErrors.push reraise in - let info = Loc.add_loc info loc in + let info = Option.cata (fun loc -> Loc.add_loc info loc) info loc in iraise (e, info)); indexes else @@ -166,7 +166,7 @@ let search_guard loc env possible_indexes fixdefs = with TypeError _ -> ()) (List.combinations possible_indexes); let errmsg = "Cannot guess decreasing argument of fix." in - user_err ~loc ~hdr:"search_guard" (Pp.str errmsg) + user_err ?loc ~hdr:"search_guard" (Pp.str errmsg) with Found indexes -> indexes) (* To force universe name declaration before use *) @@ -384,10 +384,10 @@ let process_inference_flags flags env initial_sigma (sigma,c) = let allow_anonymous_refs = ref false (* coerce to tycon if any *) -let inh_conv_coerce_to_tycon resolve_tc loc env evdref j = function +let inh_conv_coerce_to_tycon ?loc resolve_tc env evdref j = function | None -> j | Some t -> - evd_comb2 (Coercion.inh_conv_coerce_to resolve_tc loc env.ExtraEnv.env) evdref j t + evd_comb2 (Coercion.inh_conv_coerce_to ?loc resolve_tc env.ExtraEnv.env) evdref j t let check_instance loc subst = function | [] -> () @@ -568,18 +568,18 @@ let (f_genarg_interp, genarg_interp_hook) = Hook.make () (* the type constraint tycon *) let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdref (lvar : ltac_var_map) (loc, t) = - let inh_conv_coerce_to_tycon = inh_conv_coerce_to_tycon resolve_tc in + let inh_conv_coerce_to_tycon ?loc = inh_conv_coerce_to_tycon ?loc resolve_tc in let pretype_type = pretype_type k0 resolve_tc in let pretype = pretype k0 resolve_tc in let open Context.Rel.Declaration in match t with | GRef (ref,u) -> - inh_conv_coerce_to_tycon loc env evdref + inh_conv_coerce_to_tycon ~loc env evdref (pretype_ref loc evdref env ref u) tycon | GVar id -> - inh_conv_coerce_to_tycon loc env evdref + inh_conv_coerce_to_tycon ~loc env evdref (pretype_id (fun e r l t -> pretype tycon e r l t) k0 loc env evdref lvar id) tycon @@ -594,7 +594,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let args = pretype_instance k0 resolve_tc env evdref lvar loc hyps evk inst in let c = mkEvar (evk, args) in let j = (Retyping.get_judgment_of env.ExtraEnv.env !evdref c) in - inh_conv_coerce_to_tycon loc env evdref j tycon + inh_conv_coerce_to_tycon ~loc env evdref j tycon | GPatVar (someta,n) -> let env = ltac_interp_name_env k0 lvar env !evdref in @@ -674,7 +674,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre { uj_val = it_mkLambda_or_LetIn j.uj_val ctxt; uj_type = it_mkProd_or_LetIn j.uj_type ctxt }) ctxtv vdef in - Typing.check_type_fixpoint loc env.ExtraEnv.env evdref names ftys vdefj; + Typing.check_type_fixpoint ~loc env.ExtraEnv.env evdref names ftys vdefj; let nf c = nf_evar !evdref c in let ftys = Array.map nf ftys in (** FIXME *) let fdefs = Array.map (fun x -> nf (j_val x)) vdefj in @@ -696,7 +696,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let fixdecls = (names,ftys,fdefs) in let indexes = search_guard - loc env.ExtraEnv.env possible_indexes (nf_fix !evdref fixdecls) + ~loc env.ExtraEnv.env possible_indexes (nf_fix !evdref fixdecls) in make_judge (mkFix ((indexes,i),fixdecls)) ftys.(i) | GCoFix i -> @@ -709,11 +709,11 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre iraise (e, info)); make_judge (mkCoFix cofix) ftys.(i) in - inh_conv_coerce_to_tycon loc env evdref fixj tycon + inh_conv_coerce_to_tycon ~loc env evdref fixj tycon | GSort s -> let j = pretype_sort loc evdref s in - inh_conv_coerce_to_tycon loc env evdref j tycon + inh_conv_coerce_to_tycon ~loc env evdref j tycon | GApp (f,args) -> let fj = pretype empty_tycon env evdref lvar f in @@ -792,7 +792,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre else resj | _ -> resj in - inh_conv_coerce_to_tycon loc env evdref resj tycon + inh_conv_coerce_to_tycon ~loc env evdref resj tycon | GLambda(name,bk,c1,c2) -> let tycon' = evd_comb1 @@ -800,7 +800,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre match tycon with | None -> evd, tycon | Some ty -> - let evd, ty' = Coercion.inh_coerce_to_prod loc env.ExtraEnv.env evd ty in + let evd, ty' = Coercion.inh_coerce_to_prod ~loc env.ExtraEnv.env evd ty in evd, Some ty') evdref tycon in @@ -814,7 +814,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let j' = pretype rng (push_rel !evdref var env) evdref lvar c2 in let name = ltac_interp_name lvar name in let resj = judge_of_abstraction env.ExtraEnv.env (orelse_name name name') j j' in - inh_conv_coerce_to_tycon loc env evdref resj tycon + inh_conv_coerce_to_tycon ~loc env evdref resj tycon | GProd(name,bk,c1,c2) -> let j = pretype_type empty_valcon env evdref lvar c1 in @@ -838,7 +838,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let (e, info) = CErrors.push e in let info = Loc.add_loc info loc in iraise (e, info) in - inh_conv_coerce_to_tycon loc env evdref resj tycon + inh_conv_coerce_to_tycon ~loc env evdref resj tycon | GLetIn(name,c1,t,c2) -> let tycon1 = @@ -1020,10 +1020,10 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre mkCase (ci, pred, cj.uj_val, [|b1;b2|]) in let cj = { uj_val = v; uj_type = p } in - inh_conv_coerce_to_tycon loc env evdref cj tycon + inh_conv_coerce_to_tycon ~loc env evdref cj tycon | GCases (sty,po,tml,eqns) -> - Cases.compile_cases loc sty + Cases.compile_cases ~loc sty ((fun vtyc env evdref -> pretype vtyc (make_env env !evdref) evdref lvar),evdref) tycon env.ExtraEnv.env (* loc *) (po,tml,eqns) @@ -1032,7 +1032,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre match k with | CastCoerce -> let cj = pretype empty_tycon env evdref lvar c in - evd_comb1 (Coercion.inh_coerce_to_base loc env.ExtraEnv.env) evdref cj + evd_comb1 (Coercion.inh_coerce_to_base ~loc env.ExtraEnv.env) evdref cj | CastConv t | CastVM t | CastNative t -> let k = (match k with CastVM _ -> VMcast | CastNative _ -> NATIVEcast | _ -> DEFAULTcast) in let tj = pretype_type empty_valcon env evdref lvar t in @@ -1067,7 +1067,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre in let v = mkCast (cj.uj_val, k, tval) in { uj_val = v; uj_type = tval } - in inh_conv_coerce_to_tycon loc env evdref cj tycon + in inh_conv_coerce_to_tycon ~loc env evdref cj tycon and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update = let f decl (subst,update) = @@ -1128,7 +1128,7 @@ and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function | c -> let j = pretype k0 resolve_tc empty_tycon env evdref lvar c in let loc = loc_of_glob_constr c in - let tj = evd_comb1 (Coercion.inh_coerce_to_sort loc env.ExtraEnv.env) evdref j in + let tj = evd_comb1 (Coercion.inh_coerce_to_sort ~loc env.ExtraEnv.env) evdref j in match valcon with | None -> tj | Some v -> diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli index f13c10b05..79fafcf1a 100644 --- a/pretyping/pretyping.mli +++ b/pretyping/pretyping.mli @@ -24,7 +24,7 @@ open Misctypes (** An auxiliary function for searching for fixpoint guard indexes *) val search_guard : - Loc.t -> env -> int list list -> rec_declaration -> int array + ?loc:Loc.t -> env -> int list list -> rec_declaration -> int array type typing_constraint = OfType of types | IsType | WithoutTypeConstraint diff --git a/pretyping/typing.ml b/pretyping/typing.ml index c2a030bcd..41d994553 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -152,13 +152,13 @@ let e_judge_of_case env evdref ci pj cj lfj = { uj_val = mkCase (ci, pj.uj_val, cj.uj_val, Array.map j_val lfj); uj_type = rslty } -let check_type_fixpoint loc env evdref lna lar vdefj = +let check_type_fixpoint ?loc env evdref lna lar vdefj = let lt = Array.length vdefj in if Int.equal (Array.length lar) lt then for i = 0 to lt-1 do if not (Evarconv.e_cumul env evdref (vdefj.(i)).uj_type (lift lt lar.(i))) then - error_ill_typed_rec_body ~loc env !evdref + error_ill_typed_rec_body ?loc env !evdref i lna vdefj lar done @@ -361,7 +361,7 @@ and execute_recdef env evdref (names,lar,vdef) = let env1 = push_rec_types (names,lara,vdef) env in let vdefj = execute_array env1 evdref vdef in let vdefv = Array.map j_val vdefj in - let _ = check_type_fixpoint Loc.ghost env1 evdref names lara vdefj in + let _ = check_type_fixpoint env1 evdref names lara vdefj in (names,lara,vdefv) and execute_array env evdref = Array.map (execute env evdref) diff --git a/pretyping/typing.mli b/pretyping/typing.mli index 91134b499..1f3ba34e5 100644 --- a/pretyping/typing.mli +++ b/pretyping/typing.mli @@ -44,7 +44,7 @@ val check_allowed_sort : env -> evar_map -> pinductive -> constr -> constr -> (** Raise an error message if bodies have types not unifiable with the expected ones *) -val check_type_fixpoint : Loc.t -> env -> evar_map ref -> +val check_type_fixpoint : ?loc:Loc.t -> env -> evar_map ref -> Names.Name.t array -> types array -> unsafe_judgment array -> unit val judge_of_prop : unsafe_judgment diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 532cc8baa..9678e3a42 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -1250,7 +1250,7 @@ let applyHead env (type r) (evd : r Sigma.t) n c = let sigma = Sigma.to_evar_map evd in match EConstr.kind sigma (whd_all env sigma cty) with | Prod (_,c1,c2) -> - let Sigma (evar, evd', q) = Evarutil.new_evar env evd ~src:(Loc.ghost,Evar_kinds.GoalEvar) c1 in + let Sigma (evar, evd', q) = Evarutil.new_evar env evd ~src:(Loc.tag Evar_kinds.GoalEvar) c1 in apprec (n-1) (mkApp(c,[|evar|])) (subst1 evar c2) (p +> q) evd' | _ -> error "Apply_Head_Then" in @@ -1265,7 +1265,7 @@ let is_mimick_head sigma ts f = let try_to_coerce env evd c cty tycon = let j = make_judge c cty in - let (evd',j') = inh_conv_coerce_rigid_to true Loc.ghost env evd j tycon in + let (evd',j') = inh_conv_coerce_rigid_to true env evd j tycon in let evd' = Evarconv.solve_unif_constraints_with_heuristics env evd' in let evd' = Evd.map_metas_fvalue (fun c -> EConstr.Unsafe.to_constr (nf_evar evd' (EConstr.of_constr c))) evd' in (evd',j'.uj_val) diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml index 117e1900d..fddd54a9f 100644 --- a/printing/ppconstr.ml +++ b/printing/ppconstr.ml @@ -145,9 +145,9 @@ let tag_var = tag Tag.variable if !Flags.beautify && not (Int.equal n 0) then comment (CLexer.extract_comments n) else mt() - let pr_with_comments loc pp = pr_located (fun x -> x) (loc,pp) + let pr_with_comments ?loc pp = pr_located (fun x -> x) (Loc.tag ?loc pp) - let pr_sep_com sep f c = pr_with_comments (constr_loc c) (sep() ++ f c) + let pr_sep_com sep f c = pr_with_comments ~loc:(constr_loc c) (sep() ++ f c) let pr_univ l = match l with @@ -302,7 +302,7 @@ let tag_var = tag Tag.variable assert false in let loc = cases_pattern_expr_loc (loc, p) in - pr_with_comments loc + pr_with_comments ~loc (sep() ++ if prec_less prec inh then strm else surround strm) let pr_patt = pr_patt mt @@ -310,7 +310,7 @@ let tag_var = tag Tag.variable let pr_eqn pr (loc,(pl,rhs)) = let pl = List.map snd pl in spc() ++ hov 4 - (pr_with_comments loc + (pr_with_comments ~loc (str "| " ++ hov 0 (prlist_with_sep pr_bar (prlist_with_sep sep_v (pr_patt ltop)) pl ++ str " =>") ++ @@ -730,7 +730,7 @@ let tag_var = tag Tag.variable return (pr_delimiters sc (pr mt (ldelim,E) a), ldelim) in let loc = constr_loc a in - pr_with_comments loc + pr_with_comments ~loc (sep() ++ if prec_less prec inherited then strm else surround strm) type term_pr = { diff --git a/printing/ppconstr.mli b/printing/ppconstr.mli index f92caf426..482c994c2 100644 --- a/printing/ppconstr.mli +++ b/printing/ppconstr.mli @@ -35,7 +35,7 @@ val pr_or_var : ('a -> std_ppcmds) -> 'a or_var -> std_ppcmds val pr_lident : Id.t located -> std_ppcmds val pr_lname : Name.t located -> std_ppcmds -val pr_with_comments : Loc.t -> std_ppcmds -> std_ppcmds +val pr_with_comments : ?loc:Loc.t -> std_ppcmds -> std_ppcmds val pr_com_at : int -> std_ppcmds val pr_sep_com : (unit -> std_ppcmds) -> diff --git a/printing/pputils.ml b/printing/pputils.ml index e90b54685..4ae5035a2 100644 --- a/printing/pputils.ml +++ b/printing/pputils.ml @@ -15,7 +15,7 @@ open Locus open Genredexpr let pr_located pr (loc, x) = - if !Flags.beautify && loc <> Loc.ghost then + if !Flags.beautify && not (Loc.is_ghost loc) then let (b, e) = Loc.unloc loc in (* Side-effect: order matters *) let before = Pp.comment (CLexer.extract_comments b) in diff --git a/printing/prettyp.ml b/printing/prettyp.ml index 96b0f49d4..025514902 100644 --- a/printing/prettyp.ml +++ b/printing/prettyp.ml @@ -539,7 +539,7 @@ let gallina_print_constant_with_infos sp = let gallina_print_syntactic_def kn = let qid = Nametab.shortest_qualid_of_syndef Id.Set.empty kn and (vars,a) = Syntax_def.search_syntactic_definition kn in - let c = Notation_ops.glob_constr_of_notation_constr Loc.ghost a in + let c = Notation_ops.glob_constr_of_notation_constr a in hov 2 (hov 4 (str "Notation " ++ pr_qualid qid ++ @@ -752,7 +752,7 @@ let print_any_name = function let print_name = function | ByNotation (loc,(ntn,sc)) -> print_any_name - (Term (Notation.interp_notation_as_global_reference loc (fun _ -> true) + (Term (Notation.interp_notation_as_global_reference ~loc (fun _ -> true) ntn sc)) | AN ref -> print_any_name (locate_any_name ref) @@ -800,7 +800,7 @@ let print_about_any loc k = let print_about = function | ByNotation (loc,(ntn,sc)) -> print_about_any loc - (Term (Notation.interp_notation_as_global_reference loc (fun _ -> true) + (Term (Notation.interp_notation_as_global_reference ~loc (fun _ -> true) ntn sc)) | AN ref -> print_about_any (loc_of_reference ref) (locate_any_name ref) diff --git a/printing/printer.ml b/printing/printer.ml index 35ddf2e8c..279295a03 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -196,10 +196,10 @@ let qualid_of_global env r = let safe_gen f env sigma c = let orig_extern_ref = Constrextern.get_extern_reference () in - let extern_ref loc vars r = - try orig_extern_ref loc vars r + let extern_ref ?loc vars r = + try orig_extern_ref ?loc vars r with e when CErrors.noncritical e -> - Libnames.Qualid (loc, qualid_of_global env r) + Libnames.Qualid (Loc.tag ?loc @@ qualid_of_global env r) in Constrextern.set_extern_reference extern_ref; try diff --git a/proofs/clenv.ml b/proofs/clenv.ml index 17a9651cd..251e0d27d 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -681,7 +681,7 @@ let define_with_type sigma env ev c = let t = Retyping.get_type_of env sigma ev in let ty = Retyping.get_type_of env sigma c in let j = Environ.make_judge c ty in - let (sigma, j) = Coercion.inh_conv_coerce_to true (Loc.ghost) env sigma j t in + let (sigma, j) = Coercion.inh_conv_coerce_to true env sigma j t in let (ev, _) = destEvar sigma ev in let sigma = Evd.define ev (EConstr.Unsafe.to_constr j.Environ.uj_val) sigma in sigma diff --git a/proofs/goal.ml b/proofs/goal.ml index 9046f4534..5cc9d0df9 100644 --- a/proofs/goal.ml +++ b/proofs/goal.ml @@ -69,7 +69,7 @@ module V82 = struct Evd.evar_concl = concl; Evd.evar_filter = Evd.Filter.identity; Evd.evar_body = Evd.Evar_empty; - Evd.evar_source = (Loc.ghost,Evar_kinds.GoalEvar); + Evd.evar_source = (Loc.tag Evar_kinds.GoalEvar); Evd.evar_candidates = None; Evd.evar_extra = extra } in diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml index 5f4a7766f..d659eba24 100644 --- a/proofs/proof_global.ml +++ b/proofs/proof_global.ml @@ -287,13 +287,13 @@ let set_used_variables l = match entry with | LocalAssum (x,_) -> if Id.Set.mem x all_safe then orig - else (ctx, all_safe, (Loc.ghost,x)::to_clear) + else (ctx, all_safe, (Loc.tag x)::to_clear) | LocalDef (x,bo, ty) as decl -> if Id.Set.mem x all_safe then orig else let vars = Id.Set.union (vars_of env bo) (vars_of env ty) in if Id.Set.subset vars all_safe then (decl :: ctx, Id.Set.add x all_safe, to_clear) - else (ctx, all_safe, (Loc.ghost,x) :: to_clear) in + else (ctx, all_safe, (Loc.tag x) :: to_clear) in let ctx, _, to_clear = Environ.fold_named_context aux env ~init:(ctx,ctx_set,[]) in let to_clear = if !proof_using_auto_clear then to_clear else [] in diff --git a/proofs/proof_using.ml b/proofs/proof_using.ml index 2c489d6de..f0854e9aa 100644 --- a/proofs/proof_using.ml +++ b/proofs/proof_using.ml @@ -76,7 +76,7 @@ and full_set env = List.fold_right Id.Set.add (List.map NamedDecl.get_id (named_context env)) Id.Set.empty let process_expr env e ty = - let ty_expr = SsSingl(Loc.ghost, Id.of_string "Type") in + let ty_expr = SsSingl(Loc.tag @@ Id.of_string "Type") in let v_ty = process_expr env ty_expr ty in let s = Id.Set.union v_ty (process_expr env e ty) in Id.Set.elements s diff --git a/proofs/refine.ml b/proofs/refine.ml index 1ee6e0ca5..d423a658a 100644 --- a/proofs/refine.ml +++ b/proofs/refine.ml @@ -137,7 +137,7 @@ let with_type env evd c t = let my_type = Retyping.get_type_of env evd c in let j = Environ.make_judge c my_type in let (evd,j') = - Coercion.inh_conv_coerce_to true (Loc.ghost) env evd j t + Coercion.inh_conv_coerce_to true env evd j t in evd , j'.Environ.uj_val diff --git a/stm/stm.ml b/stm/stm.ml index e823373f7..38415c1dd 100644 --- a/stm/stm.ml +++ b/stm/stm.ml @@ -23,9 +23,9 @@ open Feedback open Vernacexpr open Vernac_classifier -let execution_error state_id loc msg = +let execution_error ?loc state_id msg = feedback ~id:state_id - (Message (Error, Some loc, msg)) + (Message (Error, loc, msg)) module Hooks = struct @@ -72,7 +72,7 @@ let async_proofs_workers_extra_env = ref [||] type aast = { verbose : bool; - loc : Loc.t; + loc : Loc.t option; indentation : int; strlen : int; mutable expr : vernac_expr; (* mutable: Proof using hinted by aux file *) @@ -801,9 +801,9 @@ end = struct (* {{{ *) match Stateid.get info with | Some _ -> (e, info) | None -> - let loc = Option.default Loc.ghost (Loc.get_loc info) in + let loc = Loc.get_loc info in let (e, info) = Hooks.(call_process_error_once (e, info)) in - execution_error id loc (iprint (e, info)); + execution_error ?loc id (iprint (e, info)); (e, Stateid.add info ~valid id) let same_env { system = s1 } { system = s2 } = @@ -949,7 +949,7 @@ let stm_vernac_interp ?proof id ?route { verbose; loc; expr } = the whole document state, such as backtrack, etc... so we start to design the stm command interpreter now *) set_id_for_feedback ?route id; - Aux_file.record_in_aux_set_at loc; + Aux_file.record_in_aux_set_at ?loc; (* We need to check if a command should be filtered from * vernac_entries, as it cannot handle it. This should go away in * future refactorings. @@ -968,7 +968,7 @@ let stm_vernac_interp ?proof id ?route { verbose; loc; expr } = | VernacShow ShowScript -> ShowScript.show_script () | expr -> stm_pperr_endline Pp.(fun () -> str "interpreting " ++ Ppvernac.pr_vernac expr); - try Vernacentries.interp ?verbosely:(Some verbose) ?proof (loc, expr) + try Vernacentries.interp ?verbosely:(Some verbose) ?proof (Loc.tag ?loc expr) with e -> let e = CErrors.push e in Exninfo.iraise Hooks.(call_process_error_once e) @@ -1111,12 +1111,12 @@ end (* }}} *) let hints = ref Aux_file.empty_aux_file let set_compilation_hints file = hints := Aux_file.load_aux_file_for file -let get_hint_ctx loc = - let s = Aux_file.get !hints loc "context_used" in +let get_hint_ctx ?loc = + let s = Aux_file.get ?loc !hints "context_used" in match Str.split (Str.regexp ";") s with | ids :: _ -> let ids = List.map Names.Id.of_string (Str.split (Str.regexp " ") ids) in - let ids = List.map (fun id -> Loc.ghost, id) ids in + let ids = List.map (fun id -> Loc.tag id) ids in begin match ids with | [] -> SsEmpty | x :: xs -> @@ -1125,15 +1125,15 @@ let get_hint_ctx loc = | _ -> raise Not_found let get_hint_bp_time proof_name = - try float_of_string (Aux_file.get !hints Loc.ghost proof_name) + try float_of_string (Aux_file.get !hints proof_name) with Not_found -> 1.0 -let record_pb_time proof_name loc time = +let record_pb_time ?loc proof_name time = let proof_build_time = Printf.sprintf "%.3f" time in - Aux_file.record_in_aux_at loc "proof_build_time" proof_build_time; + Aux_file.record_in_aux_at ?loc "proof_build_time" proof_build_time; if proof_name <> "" then begin - Aux_file.record_in_aux_at Loc.ghost proof_name proof_build_time; - hints := Aux_file.set !hints Loc.ghost proof_name proof_build_time + Aux_file.record_in_aux_at proof_name proof_build_time; + hints := Aux_file.set !hints proof_name proof_build_time end exception RemoteException of Pp.std_ppcmds @@ -1222,7 +1222,7 @@ module rec ProofTask : sig t_drop : bool; t_states : competence; t_assign : Proof_global.closed_proof_output Future.assignement -> unit; - t_loc : Loc.t; + t_loc : Loc.t option; t_uuid : Future.UUID.t; t_name : string } @@ -1240,8 +1240,9 @@ module rec ProofTask : sig and type request := request val build_proof_here : + ?loc:Loc.t -> drop_pt:bool -> - Stateid.t * Stateid.t -> Loc.t -> Stateid.t -> + Stateid.t * Stateid.t -> Stateid.t -> Proof_global.closed_proof_output Future.computation (* If set, only tasks overlapping with this list are processed *) @@ -1259,7 +1260,7 @@ end = struct (* {{{ *) t_drop : bool; t_states : competence; t_assign : Proof_global.closed_proof_output Future.assignement -> unit; - t_loc : Loc.t; + t_loc : Loc.t option; t_uuid : Future.UUID.t; t_name : string } @@ -1328,7 +1329,7 @@ end = struct (* {{{ *) RespBuiltProof (pl, time) -> feedback (InProgress ~-1); t_assign (`Val pl); - record_pb_time t_name t_loc time; + record_pb_time ?loc:t_loc t_name time; if !Flags.async_proofs_full || t_drop then `Stay(t_states,[States t_states]) else `End @@ -1349,16 +1350,16 @@ end = struct (* {{{ *) let info = Stateid.add ~valid:start Exninfo.null start in let e = (RemoteException (Pp.strbrk s), info) in t_assign (`Exn e); - execution_error start Loc.ghost (Pp.strbrk s); + execution_error start (Pp.strbrk s); feedback (InProgress ~-1) - let build_proof_here ~drop_pt (id,valid) loc eop = + let build_proof_here ?loc ~drop_pt (id,valid) eop = Future.create (State.exn_on id ~valid) (fun () -> let wall_clock1 = Unix.gettimeofday () in if !Flags.batch_mode then Reach.known_state ~cache:`No eop else Reach.known_state ~cache:`Shallow eop; let wall_clock2 = Unix.gettimeofday () in - Aux_file.record_in_aux_at loc "proof_build_time" + Aux_file.record_in_aux_at ?loc "proof_build_time" (Printf.sprintf "%.3f" (wall_clock2 -. wall_clock1)); let p = Proof_global.return_proof ~allow_partial:drop_pt () in if drop_pt then feedback ~id Complete; @@ -1370,7 +1371,7 @@ end = struct (* {{{ *) VCS.print (); let proof, future_proof, time = let wall_clock = Unix.gettimeofday () in - let fp = build_proof_here ~drop_pt:drop exn_info loc stop in + let fp = build_proof_here ?loc ~drop_pt:drop exn_info stop in let proof = Future.force fp in proof, fp, Unix.gettimeofday () -. wall_clock in (* We typecheck the proof with the kernel (in the worker) to spot @@ -1452,7 +1453,7 @@ end = struct (* {{{ *) msg_error(Pp.strbrk("Marshalling error: "^s^". "^ "The system state could not be sent to the worker process. "^ "Falling back to local, lazy, evaluation.")); - t_assign(`Comp(build_proof_here ~drop_pt t_exn_info t_loc t_stop)); + t_assign(`Comp(build_proof_here ?loc:t_loc ~drop_pt t_exn_info t_stop)); feedback (InProgress ~-1) end (* }}} *) @@ -1462,7 +1463,7 @@ and Slaves : sig (* (eventually) remote calls *) val build_proof : - loc:Loc.t -> drop_pt:bool -> + ?loc:Loc.t -> drop_pt:bool -> exn_info:(Stateid.t * Stateid.t) -> block_start:Stateid.t -> block_stop:Stateid.t -> name:string -> future_proof * cancel_switch @@ -1542,7 +1543,7 @@ end = struct (* {{{ *) | { step = `Fork (( { loc }, _, _, _), _) } | { step = `Qed ( { qast = { loc } }, _) } | { step = `Sideff (`Ast ( { loc }, _)) } -> - let start, stop = Loc.unloc loc in + let start, stop = Option.cata Loc.unloc (0,0) loc in msg_error Pp.( str"File " ++ str name ++ str ": proof of " ++ str r_name ++ str ": chars " ++ int start ++ str "-" ++ int stop ++ @@ -1598,10 +1599,10 @@ end = struct (* {{{ *) let info_tasks l = CList.map_i (fun i ({ Stateid.loc; name }, _) -> let time1 = - try float_of_string (Aux_file.get !hints loc "proof_build_time") + try float_of_string (Aux_file.get ?loc !hints "proof_build_time") with Not_found -> 0.0 in let time2 = - try float_of_string (Aux_file.get !hints loc "proof_check_time") + try float_of_string (Aux_file.get ?loc !hints "proof_check_time") with Not_found -> 0.0 in name, max (time1 +. time2) 0.0001,i) 0 l @@ -1622,7 +1623,7 @@ end = struct (* {{{ *) BuildProof { t_states = s2 } -> overlap_rel s1 s2 | _ -> 0) - let build_proof ~loc ~drop_pt ~exn_info ~block_start ~block_stop ~name:pname = + let build_proof ?loc ~drop_pt ~exn_info ~block_start ~block_stop ~name:pname = let id, valid as t_exn_info = exn_info in let cancel_switch = ref false in if TaskQueue.n_workers (Option.get !queue) = 0 then @@ -1637,7 +1638,7 @@ end = struct (* {{{ *) TaskQueue.enqueue_task (Option.get !queue) (task,cancel_switch); f, cancel_switch end else - ProofTask.build_proof_here ~drop_pt t_exn_info loc block_stop, cancel_switch + ProofTask.build_proof_here ?loc ~drop_pt t_exn_info block_stop, cancel_switch else let f, t_assign = Future.create_delegate ~name:pname (State.exn_on id ~valid) in let t_uuid = Future.uuid f in @@ -2026,7 +2027,7 @@ let collect_proof keep cur hd brkind id = !Flags.compilation_mode = Flags.BuildVio -> assert (VCS.Branch.equal hd hd'||VCS.Branch.equal hd VCS.edit_branch); (try - let name, hint = name ids, get_hint_ctx loc in + let name, hint = name ids, get_hint_ctx ?loc in let t, v = proof_no_using last in v.expr <- VernacProof(t, Some hint); `ASync (parent last,proof_using_ast last,accn,name,delegate name) @@ -2122,7 +2123,7 @@ let known_state ?(redefine_qed=false) ~cache id = feedback ~id:id Feedback.AddedAxiom; fst (Pfedit.solve Vernacexpr.SelectAll None tac p), ()); Option.iter (fun expr -> stm_vernac_interp id { - verbose = true; loc = Loc.ghost; expr; indentation = 0; + verbose = true; loc = None; expr; indentation = 0; strlen = 0 }) recovery_command | _ -> assert false @@ -2244,7 +2245,7 @@ let known_state ?(redefine_qed=false) ~cache id = ^"the proof's statement to avoid that.")); let fp, cancel = Slaves.build_proof - ~loc ~drop_pt ~exn_info ~block_start ~block_stop ~name in + ?loc ~drop_pt ~exn_info ~block_start ~block_stop ~name in Future.replace ofp fp; qed.fproof <- Some (fp, cancel); (* We don't generate a new state, but we still need @@ -2256,10 +2257,10 @@ let known_state ?(redefine_qed=false) ~cache id = let fp, cancel = if delegate then Slaves.build_proof - ~loc ~drop_pt ~exn_info ~block_start ~block_stop ~name + ?loc ~drop_pt ~exn_info ~block_start ~block_stop ~name else - ProofTask.build_proof_here - ~drop_pt exn_info loc block_stop, ref false + ProofTask.build_proof_here ?loc + ~drop_pt exn_info block_stop, ref false in qed.fproof <- Some (fp, cancel); let proof = @@ -2279,7 +2280,7 @@ let known_state ?(redefine_qed=false) ~cache id = log_processing_sync id name reason; reach eop; let wall_clock = Unix.gettimeofday () in - record_pb_time name x.loc (wall_clock -. !wall_clock_last_fork); + record_pb_time name ?loc:x.loc (wall_clock -. !wall_clock_last_fork); let proof = match keep with | VtDrop -> None @@ -2296,7 +2297,7 @@ let known_state ?(redefine_qed=false) ~cache id = let wall_clock2 = Unix.gettimeofday () in stm_vernac_interp id ?proof x; let wall_clock3 = Unix.gettimeofday () in - Aux_file.record_in_aux_at x.loc "proof_check_time" + Aux_file.record_in_aux_at ?loc:x.loc "proof_check_time" (Printf.sprintf "%.3f" (wall_clock3 -. wall_clock2)); Proof_global.discard_all () ), `Yes, true @@ -2658,7 +2659,7 @@ let get_ast id = | { step = `Cmd { cast = { loc; expr } } } | { step = `Fork (({ loc; expr }, _, _, _), _) } | { step = `Qed ({ qast = { loc; expr } }, _) } -> - Some (expr, loc) + Some (Loc.tag ?loc expr) | _ -> None let stop_worker n = Slaves.cancel_worker n @@ -2749,7 +2750,7 @@ let add ~ontop ?newtip verb (loc, ast) = CWarnings.set_current_loc loc; (* XXX: Classifiy vernac should be moved inside process transaction *) let clas = classify_vernac ast in - let aast = { verbose = verb; indentation; strlen; loc; expr = ast } in + let aast = { verbose = verb; indentation; strlen; loc = Some loc; expr = ast } in match process_transaction ?newtip aast clas with | `Ok -> VCS.cur_tip (), `NewTip | `Unfocus qed_id -> qed_id, `Unfocus (VCS.cur_tip ()) @@ -2770,7 +2771,7 @@ let query ~at ?(report_with=(Stateid.dummy,default_route)) s = let indentation, strlen = compute_indentation at loc in CWarnings.set_current_loc loc; let clas = classify_vernac ast in - let aast = { verbose = true; indentation; strlen; loc; expr = ast } in + let aast = { verbose = true; indentation; strlen; loc = Some loc; expr = ast } in match clas with | VtStm (w,_), _ -> ignore(process_transaction aast (VtStm (w,false), VtNow)) diff --git a/stm/stm.mli b/stm/stm.mli index 30e9629c6..699b4f5ff 100644 --- a/stm/stm.mli +++ b/stm/stm.mli @@ -87,7 +87,7 @@ val get_current_state : unit -> Stateid.t val init : unit -> unit (* This returns the node at that position *) -val get_ast : Stateid.t -> (Vernacexpr.vernac_expr * Loc.t) option +val get_ast : Stateid.t -> (Vernacexpr.vernac_expr Loc.located) option (* Filename *) val set_compilation_hints : string -> unit diff --git a/stm/vio_checking.ml b/stm/vio_checking.ml index a6237daa0..b2ea3341b 100644 --- a/stm/vio_checking.ml +++ b/stm/vio_checking.ml @@ -107,7 +107,7 @@ let schedule_vio_compilation j fs = let long_f_dot_v = Loadpath.locate_file (f^".v") in let aux = Aux_file.load_aux_file_for long_f_dot_v in let eta = - try float_of_string (Aux_file.get aux Loc.ghost "vo_compile_time") + try float_of_string (Aux_file.get aux "vo_compile_time") with Not_found -> 0.0 in jobs := (f, eta) :: !jobs) fs; diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml index 63f923dfd..768d6860d 100644 --- a/tactics/contradiction.ml +++ b/tactics/contradiction.ml @@ -31,7 +31,7 @@ let absurd c = let env = Proofview.Goal.env gl in let sigma = Sigma.to_evar_map sigma in let j = Retyping.get_judgment_of env sigma c in - let sigma, j = Coercion.inh_coerce_to_sort Loc.ghost env sigma j in + let sigma, j = Coercion.inh_coerce_to_sort env sigma j in let t = j.Environ.utj_val in let tac = Tacticals.New.tclTHENLIST [ diff --git a/tactics/equality.ml b/tactics/equality.ml index 7ae7446c8..ace7c30e9 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -1462,7 +1462,7 @@ let simpleInjClause with_evars = function | Some c -> onInductionArg (fun clear_flag -> onEquality with_evars (injEq ~old:true with_evars clear_flag None)) c let injConcl = injClause None false None -let injHyp clear_flag id = injClause None false (Some (clear_flag,ElimOnIdent (Loc.ghost,id))) +let injHyp clear_flag id = injClause None false (Some (clear_flag,ElimOnIdent (Loc.tag id))) let decompEqThen ntac (lbeq,_,(t,t1,t2) as u) clause = Proofview.Goal.enter { enter = begin fun gl -> diff --git a/tactics/inv.ml b/tactics/inv.ml index 904a17417..f2147db40 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -283,7 +283,7 @@ let generalizeRewriteIntros as_mode tac depids id = end } let error_too_many_names pats = - let loc = Loc.join_loc (fst (List.hd pats)) (fst (List.last pats)) in + let loc = Loc.merge (fst (List.hd pats)) (fst (List.last pats)) in Proofview.tclENV >>= fun env -> tclZEROMSG ~loc ( str "Unexpected " ++ diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index 90b7d6581..074c8b9de 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -171,23 +171,23 @@ let fix_empty_or_and_pattern nv l = | IntroOrPattern [[]] -> IntroOrPattern (List.make nv []) | _ -> l -let check_or_and_pattern_size check_and loc names branchsigns = +let check_or_and_pattern_size ?loc check_and names branchsigns = let n = Array.length branchsigns in let msg p1 p2 = strbrk "a conjunctive pattern made of " ++ int p1 ++ (if p1 == p2 then mt () else str " or " ++ int p2) ++ str " patterns" in let err1 p1 p2 = - user_err ~loc (str "Expects " ++ msg p1 p2 ++ str ".") in + user_err ?loc (str "Expects " ++ msg p1 p2 ++ str ".") in let errn n = - user_err ~loc (str "Expects a disjunctive pattern with " ++ int n + user_err ?loc (str "Expects a disjunctive pattern with " ++ int n ++ str " branches.") in let err1' p1 p2 = - user_err ~loc (strbrk "Expects a disjunctive pattern with 1 branch or " ++ msg p1 p2 ++ str ".") in - let errforthcoming loc = - user_err ~loc (strbrk "Unexpected non atomic pattern.") in + user_err ?loc (strbrk "Expects a disjunctive pattern with 1 branch or " ++ msg p1 p2 ++ str ".") in + let errforthcoming ?loc = + user_err ?loc (strbrk "Unexpected non atomic pattern.") in match names with | IntroAndPattern l -> if not (Int.equal n 1) then errn n; let l' = List.filter (function _,IntroForthcoming _ -> true | _,IntroNaming _ | _,IntroAction _ -> false) l in - if l' != [] then errforthcoming (fst (List.hd l')); + if l' != [] then errforthcoming ~loc:(fst (List.hd l')); if check_and then let p1 = List.count (fun x -> x) branchsigns.(0) in let p2 = List.length branchsigns.(0) in @@ -195,7 +195,7 @@ let check_or_and_pattern_size check_and loc names branchsigns = if not (Int.equal p p1 || Int.equal p p2) then err1 p1 p2; if Int.equal p p1 then IntroAndPattern - (List.extend branchsigns.(0) (Loc.ghost,IntroNaming IntroAnonymous) l) + (List.extend branchsigns.(0) (Loc.tag @@ IntroNaming IntroAnonymous) l) else names else @@ -208,20 +208,20 @@ let check_or_and_pattern_size check_and loc names branchsigns = err1' p1 p2 else errn n; names -let get_and_check_or_and_pattern_gen check_and loc names branchsigns = - let names = check_or_and_pattern_size check_and loc names branchsigns in +let get_and_check_or_and_pattern_gen ?loc check_and names branchsigns = + let names = check_or_and_pattern_size ?loc check_and names branchsigns in match names with | IntroAndPattern l -> [|l|] | IntroOrPattern l -> Array.of_list l -let get_and_check_or_and_pattern = get_and_check_or_and_pattern_gen true +let get_and_check_or_and_pattern ?loc = get_and_check_or_and_pattern_gen ?loc true let compute_induction_names_gen check_and branchletsigns = function | None -> Array.make (Array.length branchletsigns) [] | Some (loc,names) -> let names = fix_empty_or_and_pattern (Array.length branchletsigns) names in - get_and_check_or_and_pattern_gen check_and loc names branchletsigns + get_and_check_or_and_pattern_gen check_and ~loc names branchletsigns let compute_induction_names = compute_induction_names_gen true diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli index 3b90ec514..d60196213 100644 --- a/tactics/tacticals.mli +++ b/tactics/tacticals.mli @@ -116,7 +116,7 @@ type branch_assumptions = private { error message if |pats| <> |branchsign|; extends them if no pattern is given for let-ins in the case of a conjunctive pattern *) val get_and_check_or_and_pattern : - Loc.t -> delayed_open_constr or_and_intro_pattern_expr -> + ?loc:Loc.t -> delayed_open_constr or_and_intro_pattern_expr -> bool list array -> intro_patterns array (** Tolerate "[]" to mean a disjunctive pattern of any length *) diff --git a/tactics/tactics.ml b/tactics/tactics.ml index e79258582..7dad90242 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -53,8 +53,6 @@ module NamedDecl = Context.Named.Declaration let inj_with_occurrences e = (AllOccurrences,e) -let dloc = Loc.ghost - let typ_of env sigma c = let open Retyping in try get_type_of ~lax:true env (Sigma.to_evar_map sigma) c @@ -427,11 +425,11 @@ let default_id env sigma decl = type name_flag = | NamingAvoid of Id.t list | NamingBasedOn of Id.t * Id.t list - | NamingMustBe of Loc.t * Id.t + | NamingMustBe of Id.t Loc.located let naming_of_name = function | Anonymous -> NamingAvoid [] - | Name id -> NamingMustBe (dloc,id) + | Name id -> NamingMustBe (Loc.tag id) let find_name mayrepl decl naming gl = match naming with | NamingAvoid idl -> @@ -469,7 +467,7 @@ let assert_before_gen b naming t = assert_before_then_gen b naming t (fun _ -> Proofview.tclUNIT ()) let assert_before na = assert_before_gen false (naming_of_name na) -let assert_before_replacing id = assert_before_gen true (NamingMustBe (dloc,id)) +let assert_before_replacing id = assert_before_gen true (NamingMustBe (Loc.tag id)) let assert_after_then_gen b naming t tac = let open Context.Rel.Declaration in @@ -488,7 +486,7 @@ let assert_after_gen b naming t = assert_after_then_gen b naming t (fun _ -> (Proofview.tclUNIT ())) let assert_after na = assert_after_gen false (naming_of_name na) -let assert_after_replacing id = assert_after_gen true (NamingMustBe (dloc,id)) +let assert_after_replacing id = assert_after_gen true (NamingMustBe (Loc.tag id)) (**************************************************************) (* Fixpoints and CoFixpoints *) @@ -969,7 +967,7 @@ let rec intro_then_gen name_flag move_flag force_flag dep_flag tac = end } let intro_gen n m f d = intro_then_gen n m f d (fun _ -> Proofview.tclUNIT ()) -let intro_mustbe_force id = intro_gen (NamingMustBe (dloc,id)) MoveLast true false +let intro_mustbe_force id = intro_gen (NamingMustBe (Loc.tag id)) MoveLast true false let intro_using id = intro_gen (NamingBasedOn (id,[])) MoveLast false false let intro_then = intro_then_gen (NamingAvoid []) MoveLast false false @@ -979,7 +977,7 @@ let intro_avoiding l = intro_gen (NamingAvoid l) MoveLast false false let intro_move_avoid idopt avoid hto = match idopt with | None -> intro_gen (NamingAvoid avoid) hto true false - | Some id -> intro_gen (NamingMustBe (dloc,id)) hto true false + | Some id -> intro_gen (NamingMustBe (Loc.tag id)) hto true false let intro_move idopt hto = intro_move_avoid idopt [] hto @@ -1139,7 +1137,7 @@ let try_intros_until tac = function let rec intros_move = function | [] -> Proofview.tclUNIT () | (hyp,destopt) :: rest -> - Tacticals.New.tclTHEN (intro_gen (NamingMustBe (dloc,hyp)) destopt false false) + Tacticals.New.tclTHEN (intro_gen (NamingMustBe (Loc.tag hyp)) destopt false false) (intros_move rest) let run_delayed env sigma c = @@ -1291,7 +1289,7 @@ let clenv_refine_in ?(sidecond_first=false) with_evars ?(with_classes=true) error_uninstantiated_metas new_hyp_typ clenv; let new_hyp_prf = clenv_value clenv in let exact_tac = Proofview.V82.tactic (Tacmach.refine_no_check new_hyp_prf) in - let naming = NamingMustBe (dloc,targetid) in + let naming = NamingMustBe (Loc.tag targetid) in let with_clear = do_replace (Some id) naming in Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARS (clear_metas clenv.evd)) @@ -1798,13 +1796,13 @@ let apply_with_delayed_bindings_gen b e l = (one k f) (aux cbl) in aux l -let apply_with_bindings cb = apply_with_bindings_gen false false [None,(dloc,cb)] +let apply_with_bindings cb = apply_with_bindings_gen false false [None,(Loc.tag cb)] -let eapply_with_bindings cb = apply_with_bindings_gen false true [None,(dloc,cb)] +let eapply_with_bindings cb = apply_with_bindings_gen false true [None,(Loc.tag cb)] -let apply c = apply_with_bindings_gen false false [None,(dloc,(c,NoBindings))] +let apply c = apply_with_bindings_gen false false [None,(Loc.tag (c,NoBindings))] -let eapply c = apply_with_bindings_gen false true [None,(dloc,(c,NoBindings))] +let eapply c = apply_with_bindings_gen false true [None,(Loc.tag (c,NoBindings))] let apply_list = function | c::l -> apply_with_bindings (c,ImplicitBindings l) @@ -2216,7 +2214,7 @@ let constructor_tac with_evars expctdnumopt i lbind = (Proofview.Goal.env gl) sigma (fst mind, i) in let cons = mkConstructU (cons, EInstance.make u) in - let apply_tac = general_apply true false with_evars None (dloc,(cons,lbind)) in + let apply_tac = general_apply true false with_evars None (Loc.tag (cons,lbind)) in let tac = (Tacticals.New.tclTHENLIST [ @@ -2301,7 +2299,7 @@ let my_find_eq_data_decompose gl t = -> None | Constr_matching.PatternMatchingFailure -> None -let intro_decomp_eq loc l thin tac id = +let intro_decomp_eq ?loc l thin tac id = Proofview.Goal.enter { enter = begin fun gl -> let c = mkVar id in let t = Tacmach.New.pf_unsafe_type_of gl c in @@ -2309,13 +2307,13 @@ let intro_decomp_eq loc l thin tac id = match my_find_eq_data_decompose gl t with | Some (eq,u,eq_args) -> !intro_decomp_eq_function - (fun n -> tac ((dloc,id)::thin) (Some (true,n)) l) + (fun n -> tac ((Loc.tag id)::thin) (Some (true,n)) l) (eq,t,eq_args) (c, t) | None -> Tacticals.New.tclZEROMSG (str "Not a primitive equality here.") end } -let intro_or_and_pattern loc with_evars bracketed ll thin tac id = +let intro_or_and_pattern ?loc with_evars bracketed ll thin tac id = Proofview.Goal.enter { enter = begin fun gl -> let c = mkVar id in let t = Tacmach.New.pf_unsafe_type_of gl c in @@ -2323,7 +2321,7 @@ let intro_or_and_pattern loc with_evars bracketed ll thin tac id = let branchsigns = compute_constructor_signatures false ind in let nv_with_let = Array.map List.length branchsigns in let ll = fix_empty_or_and_pattern (Array.length branchsigns) ll in - let ll = get_and_check_or_and_pattern loc ll branchsigns in + let ll = get_and_check_or_and_pattern ?loc ll branchsigns in Tacticals.New.tclTHENLASTn (Tacticals.New.tclTHEN (simplest_ecase c) (clear [id])) (Array.map2 (fun n l -> tac thin (Some (bracketed,n)) l) @@ -2372,8 +2370,8 @@ let rewrite_hyp_then assert_style with_evars thin l2r id tac = Tacticals.New.tclTHENFIRST eqtac (tac thin) end } -let prepare_naming loc = function - | IntroIdentifier id -> NamingMustBe (loc,id) +let prepare_naming ?loc = function + | IntroIdentifier id -> NamingMustBe (Loc.tag ?loc id) | IntroAnonymous -> NamingAvoid [] | IntroFresh id -> NamingBasedOn (id,[]) @@ -2415,7 +2413,7 @@ let make_tmp_naming avoid l = function case of IntroFresh, we should use check_thin_clash_then anyway to prevent the case of an IntroFresh precisely using the wild_id *) | IntroWildcard -> NamingBasedOn (wild_id,avoid@explicit_intro_names l) - | pat -> NamingAvoid(avoid@explicit_intro_names ((dloc,IntroAction pat)::l)) + | pat -> NamingAvoid(avoid@explicit_intro_names ((Loc.tag @@ IntroAction pat)::l)) let fit_bound n = function | None -> true @@ -2451,7 +2449,7 @@ let rec intro_patterns_core with_evars b avoid ids thin destopt bound n tac = | [] -> (* Behave as IntroAnonymous *) intro_patterns_core with_evars b avoid ids thin destopt bound n tac - [dloc,IntroNaming IntroAnonymous] + [Loc.tag @@ IntroNaming IntroAnonymous] | (loc,pat) :: l -> if exceed_bound n bound then error_unexpected_extra_pattern loc bound pat else match pat with @@ -2463,7 +2461,7 @@ let rec intro_patterns_core with_evars b avoid ids thin destopt bound n tac = | IntroAction pat -> intro_then_gen (make_tmp_naming avoid l pat) destopt true false - (intro_pattern_action loc with_evars (b || not (List.is_empty l)) false + (intro_pattern_action ~loc with_evars (b || not (List.is_empty l)) false pat thin destopt (fun thin bound' -> intro_patterns_core with_evars b avoid ids thin destopt bound' 0 (fun ids thin -> @@ -2488,21 +2486,21 @@ and intro_pattern_naming loc with_evars b avoid ids pat thin destopt bound n tac destopt true false (fun id -> intro_patterns_core with_evars b avoid (id::ids) thin destopt bound n tac l) -and intro_pattern_action loc with_evars b style pat thin destopt tac id = +and intro_pattern_action ?loc with_evars b style pat thin destopt tac id = match pat with | IntroWildcard -> - tac ((loc,id)::thin) None [] + tac ((Loc.tag ?loc id)::thin) None [] | IntroOrAndPattern ll -> - intro_or_and_pattern loc with_evars b ll thin tac id + intro_or_and_pattern ?loc with_evars b ll thin tac id | IntroInjection l' -> - intro_decomp_eq loc l' thin tac id + intro_decomp_eq ?loc l' thin tac id | IntroRewrite l2r -> rewrite_hyp_then style with_evars thin l2r id (fun thin -> tac thin None []) | IntroApplyOn ((loc',f),(loc,pat)) -> let naming,tac_ipat = - prepare_intros_loc loc with_evars (IntroIdentifier id) destopt pat in + prepare_intros ~loc with_evars (IntroIdentifier id) destopt pat in let doclear = - if naming = NamingMustBe (loc,id) then + if naming = NamingMustBe (Loc.tag ~loc id) then Proofview.tclUNIT () (* apply_in_once do a replacement *) else clear [id] in @@ -2513,18 +2511,18 @@ and intro_pattern_action loc with_evars b style pat thin destopt tac id = apply_in_delayed_once false true true with_evars naming id (None,(loc',f)) (fun id -> Tacticals.New.tclTHENLIST [doclear; tac_ipat id; tac thin None []]) -and prepare_intros_loc loc with_evars dft destopt = function +and prepare_intros ?loc with_evars dft destopt = function | IntroNaming ipat -> - prepare_naming loc ipat, + prepare_naming ?loc ipat, (fun id -> move_hyp id destopt) | IntroAction ipat -> - prepare_naming loc dft, + prepare_naming ?loc dft, (let tac thin bound = intro_patterns_core with_evars true [] [] thin destopt bound 0 (fun _ l -> clear_wildcards l) in fun id -> - intro_pattern_action loc with_evars true true ipat [] destopt tac id) - | IntroForthcoming _ -> user_err ~loc + intro_pattern_action ?loc with_evars true true ipat [] destopt tac id) + | IntroForthcoming _ -> user_err ?loc (str "Introduction pattern for one hypothesis expected.") let intro_patterns_bound_to with_evars n destopt = @@ -2536,7 +2534,7 @@ let intro_patterns_to with_evars destopt = [] [] [] destopt None 0 (fun _ l -> clear_wildcards l) let intro_pattern_to with_evars destopt pat = - intro_patterns_to with_evars destopt [dloc,pat] + intro_patterns_to with_evars destopt [Loc.tag pat] let intro_patterns with_evars = intro_patterns_to with_evars MoveLast @@ -2549,20 +2547,20 @@ let intros_patterns with_evars = function (* Forward reasoning *) (**************************) -let prepare_intros with_evars dft destopt = function - | None -> prepare_naming dloc dft, (fun _id -> Proofview.tclUNIT ()) - | Some (loc,ipat) -> prepare_intros_loc loc with_evars dft destopt ipat +let prepare_intros_opt with_evars dft destopt = function + | None -> prepare_naming dft, (fun _id -> Proofview.tclUNIT ()) + | Some (loc,ipat) -> prepare_intros ~loc with_evars dft destopt ipat let ipat_of_name = function | Anonymous -> None - | Name id -> Some (dloc, IntroNaming (IntroIdentifier id)) + | Name id -> Some (Loc.tag @@ IntroNaming (IntroIdentifier id)) let head_ident sigma c = let c = fst (decompose_app sigma (snd (decompose_lam_assum sigma c))) in if isVar sigma c then Some (destVar sigma c) else None let assert_as first hd ipat t = - let naming,tac = prepare_intros false IntroAnonymous MoveLast ipat in + let naming,tac = prepare_intros_opt false IntroAnonymous MoveLast ipat in let repl = do_replace hd naming in let tac = if repl then (fun id -> Proofview.tclUNIT ()) else tac in if first then assert_before_then_gen repl naming t tac @@ -2580,10 +2578,10 @@ let general_apply_in sidecond_first with_delta with_destruct with_evars if with_evars then MoveLast (* evars would depend on the whole context *) else get_previous_hyp_position id gl in let naming,ipat_tac = - prepare_intros with_evars (IntroIdentifier id) destopt ipat in + prepare_intros_opt with_evars (IntroIdentifier id) destopt ipat in let lemmas_target, last_lemma_target = let last,first = List.sep_last lemmas in - List.map (fun lem -> (NamingMustBe (dloc,id),lem)) first, (naming,last) + List.map (fun lem -> (NamingMustBe (Loc.tag id),lem)) first, (naming,last) in (* We chain apply_in_once, ending with an intro pattern *) List.fold_right tac lemmas_target (tac last_lemma_target ipat_tac) id @@ -2661,7 +2659,7 @@ let letin_tac_gen with_eq (id,depdecls,lastlhyp,ccl,c) ty = let tac = Tacticals.New.tclTHENLIST [ convert_concl_no_check newcl DEFAULTcast; - intro_gen (NamingMustBe (dloc,id)) (decode_hyp lastlhyp) true false; + intro_gen (NamingMustBe (Loc.tag id)) (decode_hyp lastlhyp) true false; Tacticals.New.tclMAP convert_hyp_no_check depdecls; eq_tac ] in @@ -2961,7 +2959,7 @@ let specialize (c,lbind) ipat = | Var id when Id.List.mem id (Tacmach.New.pf_ids_of_hyps gl) -> (* Like assert (id:=id args) but with the concept of specialization *) let naming,tac = - prepare_intros false (IntroIdentifier id) MoveLast ipat in + prepare_intros_opt false (IntroIdentifier id) MoveLast ipat in let repl = do_replace (Some id) naming in Tacticals.New.tclTHENFIRST (assert_before_then_gen repl naming typ tac) @@ -2975,7 +2973,7 @@ let specialize (c,lbind) ipat = | Some (loc,ipat) -> (* Like pose proof with extra support for "with" bindings *) (* even though the "with" bindings forces full application *) - let naming,tac = prepare_intros_loc loc false IntroAnonymous MoveLast ipat in + let naming,tac = prepare_intros ~loc false IntroAnonymous MoveLast ipat in Tacticals.New.tclTHENFIRST (assert_before_then_gen false naming typ tac) (exact_no_check term) @@ -3063,7 +3061,7 @@ let intropattern_of_name gl avoid = function | Name id -> IntroNaming (IntroIdentifier (new_fresh_id avoid id gl)) let rec consume_pattern avoid na isdep gl = function - | [] -> ((dloc, intropattern_of_name gl avoid na), []) + | [] -> ((Loc.tag @@ intropattern_of_name gl avoid na), []) | (loc,IntroForthcoming true)::names when not isdep -> consume_pattern avoid na isdep gl names | (loc,IntroForthcoming _)::names as fullpat -> @@ -3140,7 +3138,7 @@ let induct_discharge with_evars dests avoid' tac (avoid,ra) names = let (recpat,names) = match names with | [loc,IntroNaming (IntroIdentifier id) as pat] -> let id' = next_ident_away (add_prefix "IH" id) avoid in - (pat, [dloc, IntroNaming (IntroIdentifier id')]) + (pat, [Loc.tag @@ IntroNaming (IntroIdentifier id')]) | _ -> consume_pattern avoid (Name recvarname) deprec gl names in let dest = get_recarg_dest dests in dest_intro_patterns with_evars avoid thin dest [recpat] (fun ids thin -> @@ -5026,14 +5024,14 @@ module Simple = struct let intro x = intro_move (Some x) MoveLast let apply c = - apply_with_bindings_gen false false [None,(Loc.ghost,(c,NoBindings))] + apply_with_bindings_gen false false [None,(Loc.tag (c,NoBindings))] let eapply c = - apply_with_bindings_gen false true [None,(Loc.ghost,(c,NoBindings))] + apply_with_bindings_gen false true [None,(Loc.tag (c,NoBindings))] let elim c = elim false None (c,NoBindings) None let case c = general_case_analysis false None (c,NoBindings) let apply_in id c = - apply_in false false id [None,(Loc.ghost, (c, NoBindings))] None + apply_in false false id [None,(Loc.tag (c, NoBindings))] None end diff --git a/toplevel/coqtop.ml b/toplevel/coqtop.ml index f5f43ff66..b7065993f 100644 --- a/toplevel/coqtop.ml +++ b/toplevel/coqtop.ml @@ -185,7 +185,7 @@ let load_vernacular sid = let load_vernacular_obj = ref ([] : string list) let add_vernac_obj s = load_vernacular_obj := s :: !load_vernacular_obj let load_vernac_obj () = - let map dir = Qualid (Loc.ghost, qualid_of_string dir) in + let map dir = Qualid (Loc.tag @@ qualid_of_string dir) in Vernacentries.vernac_require None None (List.rev_map map !load_vernacular_obj) let require_prelude () = @@ -200,7 +200,7 @@ let require_list = ref ([] : string list) let add_require s = require_list := s :: !require_list let require () = let () = if !load_init then silently require_prelude () in - let map dir = Qualid (Loc.ghost, qualid_of_string dir) in + let map dir = Qualid (Loc.tag @@ qualid_of_string dir) in Vernacentries.vernac_require None (Some false) (List.rev_map map !require_list) let add_compat_require v = diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml index 3359a1672..d6bcd2f15 100644 --- a/toplevel/vernac.ml +++ b/toplevel/vernac.ml @@ -146,9 +146,11 @@ let rec interp_vernac sid po (loc,com) = interp com with reraise -> let (reraise, info) = CErrors.push reraise in - let loc' = Option.default Loc.ghost (Loc.get_loc info) in - if Loc.is_ghost loc' then iraise (reraise, Loc.add_loc info loc) - else iraise (reraise, info) + let info = begin + match Loc.get_loc info with + | None -> Loc.add_loc info loc + | Some _ -> info + end in iraise (reraise, info) (* Load a vernac file. CErrors are annotated with file and location *) and load_vernac verbosely sid file = @@ -281,7 +283,7 @@ let compile verbosely f = let wall_clock2 = Unix.gettimeofday () in check_pending_proofs (); Library.save_library_to ldir long_f_dot_vo (Global.opaque_tables ()); - Aux_file.record_in_aux_at Loc.ghost "vo_compile_time" + Aux_file.record_in_aux_at "vo_compile_time" (Printf.sprintf "%.3f" (wall_clock2 -. wall_clock1)); Aux_file.stop_aux_file (); if !Flags.xml_export then Hook.get f_xml_end_library (); diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml index b9c4c6cc5..fd454b67e 100644 --- a/vernac/auto_ind_decl.ml +++ b/vernac/auto_ind_decl.ml @@ -60,8 +60,6 @@ exception NonSingletonProp of inductive exception DecidabilityMutualNotSupported exception NoDecidabilityCoInductive -let dl = Loc.ghost - let constr_of_global g = lazy (Universes.constr_of_global g) (* Some pre declaration of constant we are going to use *) @@ -88,12 +86,12 @@ let destruct_on c = destruct false None c None None let destruct_on_using c id = destruct false None c - (Some (dl,IntroOrPattern [[dl,IntroNaming IntroAnonymous]; - [dl,IntroNaming (IntroIdentifier id)]])) + (Some (Loc.tag @@ IntroOrPattern [[Loc.tag @@ IntroNaming IntroAnonymous]; + [Loc.tag @@ IntroNaming (IntroIdentifier id)]])) None let destruct_on_as c l = - destruct false None c (Some (dl,l)) None + destruct false None c (Some (Loc.tag l)) None (* reconstruct the inductive with the correct deBruijn indexes *) let mkFullInd (ind,u) n = @@ -611,8 +609,8 @@ repeat ( apply andb_prop in z;let z1:= fresh "Z" in destruct z as [z1 z]). Proofview.Goal.enter { enter = begin fun gl -> let fresht = fresh_id (Id.of_string "Z") gl in destruct_on_as (EConstr.mkVar freshz) - (IntroOrPattern [[dl,IntroNaming (IntroIdentifier fresht); - dl,IntroNaming (IntroIdentifier freshz)]]) + (IntroOrPattern [[Loc.tag @@ IntroNaming (IntroIdentifier fresht); + Loc.tag @@ IntroNaming (IntroIdentifier freshz)]]) end } ]); (* diff --git a/vernac/classes.ml b/vernac/classes.ml index ffe03bfb7..bdaa3ece6 100644 --- a/vernac/classes.ml +++ b/vernac/classes.ml @@ -408,7 +408,7 @@ let context poly l = let decl = (Discharge, poly, Definitional) in let nstatus = pi3 (Command.declare_assumption false decl (t, !uctx) [] [] impl - Vernacexpr.NoInline (Loc.ghost, id)) + Vernacexpr.NoInline (Loc.tag id)) in let () = uctx := Univ.ContextSet.empty in status && nstatus diff --git a/vernac/command.ml b/vernac/command.ml index 446afb578..fbaa09430 100644 --- a/vernac/command.ml +++ b/vernac/command.ml @@ -267,7 +267,7 @@ match local with (gr,inst,Lib.is_modtype_strict ()) let interp_assumption evdref env impls bl c = - let c = mkCProdN (local_binders_loc bl) bl c in + let c = mkCProdN ?loc:(local_binders_loc bl) bl c in let ty, impls = interp_type_evars_impls env evdref ~impls c in let ty = EConstr.Unsafe.to_constr ty in (ty, impls) @@ -917,7 +917,7 @@ let mkSubset name typ prop = [| typ; mkLambda (name, typ, prop) |]) let sigT = Lazy.from_fun build_sigma_type -let make_qref s = Qualid (Loc.ghost, qualid_of_string s) +let make_qref s = Qualid (Loc.tag @@ qualid_of_string s) let lt_ref = make_qref "Init.Peano.lt" let rec telescope l = @@ -1059,7 +1059,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = mkApp (EConstr.of_constr (Universes.constr_of_global (delayed_force fix_sub_ref)), [| argtyp ; wf_rel ; Evarutil.e_new_evar env evdref - ~src:(Loc.ghost, Evar_kinds.QuestionMark (Evar_kinds.Define false)) wf_proof; + ~src:(Loc.tag @@ Evar_kinds.QuestionMark (Evar_kinds.Define false)) wf_proof; prop |]) in let def = Typing.e_solve_evars env evdref def in @@ -1212,7 +1212,7 @@ let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ind let fixdefs = List.map Option.get fixdefs in let fixdecls = prepare_recursive_declaration fixnames fixtypes fixdefs in let env = Global.env() in - let indexes = search_guard Loc.ghost env indexes fixdecls in + let indexes = search_guard env indexes fixdecls in let fiximps = List.map (fun (n,r,p) -> r) fiximps in let vars = Universes.universes_of_constr (mkFix ((indexes,0),fixdecls)) in let fixdecls = @@ -1327,8 +1327,7 @@ let do_program_recursive local p fixkind fixl ntns = Array.of_list (List.map (subst_vars (List.rev fixnames)) fixdefs) in let indexes = - Pretyping.search_guard - Loc.ghost (Global.env ()) possible_indexes fixdecls in + Pretyping.search_guard (Global.env ()) possible_indexes fixdecls in List.iteri (fun i _ -> Inductive.check_fix env ((indexes,i),fixdecls)) diff --git a/vernac/explainErr.ml b/vernac/explainErr.ml index 04841c922..ed2dd274a 100644 --- a/vernac/explainErr.ml +++ b/vernac/explainErr.ml @@ -119,6 +119,6 @@ let process_vernac_interp_error ?(allow_uncaught=true) (exc, info) = in match e' with | None -> e - | Some (None, loc) -> (fst e, Loc.add_loc (snd e) loc) - | Some (Some msg, loc) -> + | Some (loc, None) -> (fst e, Loc.add_loc (snd e) loc) + | Some (loc, Some msg) -> (EvaluatedError (msg, Some (fst e)), Loc.add_loc (snd e) loc) diff --git a/vernac/explainErr.mli b/vernac/explainErr.mli index 370ad7e3b..9202729ce 100644 --- a/vernac/explainErr.mli +++ b/vernac/explainErr.mli @@ -18,4 +18,4 @@ val process_vernac_interp_error : ?allow_uncaught:bool -> Util.iexn -> Util.iexn val explain_exn_default : exn -> Pp.std_ppcmds -val register_additional_error_info : (Util.iexn -> (Pp.std_ppcmds option * Loc.t) option) -> unit +val register_additional_error_info : (Util.iexn -> (Pp.std_ppcmds option Loc.located) option) -> unit diff --git a/vernac/indschemes.ml b/vernac/indschemes.ml index 9ba4e46e4..2367177e2 100644 --- a/vernac/indschemes.ml +++ b/vernac/indschemes.ml @@ -379,7 +379,7 @@ requested | InType -> recs ^ "t_nodep") ) in let newid = add_suffix (basename_of_global (IndRef ind)) suffix in - let newref = (Loc.ghost,newid) in + let newref = Loc.tag newid in ((newref,isdep,ind,z)::l1),l2 in match t with diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml index 993a2c260..b51792b1e 100644 --- a/vernac/lemmas.ml +++ b/vernac/lemmas.ml @@ -81,7 +81,7 @@ let adjust_guardness_conditions const = function List.fold_left (fun e (_,c,cb,_) -> add c cb e) env l) env (Safe_typing.side_effects_of_private_constants eff) in let indexes = - search_guard Loc.ghost env + search_guard env possible_indexes fixdecls in (mkFix ((indexes,0),fixdecls), ctx), eff | _ -> (body, ctx), eff) } diff --git a/vernac/metasyntax.ml b/vernac/metasyntax.ml index 98b2c3729..a586fe9f2 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 = Loc.tag @@ CRef (Ident (Loc.ghost, Id.of_string x),None) +let inject_var x = Loc.tag @@ CRef (Ident (Loc.tag @@ Id.of_string x),None) let add_infix local ((loc,inf),modifiers) pr sc = check_infix_modifiers modifiers; diff --git a/vernac/obligations.ml b/vernac/obligations.ml index ea2401b5c..5e65855ef 100644 --- a/vernac/obligations.ml +++ b/vernac/obligations.ml @@ -558,8 +558,7 @@ let declare_mutual_definition l = List.map3 compute_possible_guardness_evidences wfl fixdefs fixtypes in let indexes = - Pretyping.search_guard - Loc.ghost (Global.env()) + Pretyping.search_guard (Global.env()) possible_indexes fixdecls in Some indexes, List.map_i (fun i _ -> @@ -674,7 +673,7 @@ let init_prog_info ?(opaque = false) sign n pl b t ctx deps fixkind assert(Int.equal (Array.length obls) 0); let n = Nameops.add_suffix n "_obligation" in [| { obl_name = n; obl_body = None; - obl_location = Loc.ghost, Evar_kinds.InternalHole; obl_type = t; + obl_location = Loc.tag Evar_kinds.InternalHole; obl_type = t; obl_status = false, Evar_kinds.Expand; obl_deps = Int.Set.empty; obl_tac = None } |], mkVar n diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml index 86406dbe5..75f403e33 100644 --- a/vernac/vernacentries.ml +++ b/vernac/vernacentries.ml @@ -643,7 +643,7 @@ let vernac_declare_module export (loc, id) binders_ast mty_ast = in Dumpglob.dump_moddef loc mp "mod"; if_verbose Feedback.msg_info (str "Module " ++ pr_id id ++ str " is declared"); - Option.iter (fun export -> vernac_import export [Ident (Loc.ghost,id)]) export + Option.iter (fun export -> vernac_import export [Ident (Loc.tag id)]) export let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l = (* We check the state of the system (in section, in module type) @@ -668,7 +668,7 @@ let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l = List.iter (fun (export,id) -> Option.iter - (fun export -> vernac_import export [Ident (Loc.ghost,id)]) export + (fun export -> vernac_import export [Ident (Loc.tag id)]) export ) argsexport | _::_ -> let binders_ast = List.map @@ -683,7 +683,7 @@ let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l = Dumpglob.dump_moddef loc mp "mod"; if_verbose Feedback.msg_info (str "Module " ++ pr_id id ++ str " is defined"); - Option.iter (fun export -> vernac_import export [Ident (Loc.ghost,id)]) + Option.iter (fun export -> vernac_import export [Ident (Loc.tag id)]) export let vernac_end_module export (loc,id as lid) = @@ -715,7 +715,7 @@ let vernac_declare_module_type (loc,id) binders_ast mty_sign mty_ast_l = List.iter (fun (export,id) -> Option.iter - (fun export -> vernac_import export [Ident (Loc.ghost,id)]) export + (fun export -> vernac_import export [Ident (Loc.tag id)]) export ) argsexport | _ :: _ -> @@ -1176,10 +1176,10 @@ let vernac_arguments locality reference args more_implicits nargs_for_red flags end; if scopes_specified || clear_scopes_flag then begin - let scopes = List.map (Option.map (fun (o,k) -> + let scopes = List.map (Option.map (fun (loc,k) -> try ignore (Notation.find_scope k); k with UserError _ -> - Notation.find_delimiters_scope o k)) scopes + Notation.find_delimiters_scope ~loc k)) scopes in vernac_arguments_scope locality reference scopes end; @@ -1708,7 +1708,7 @@ let interp_search_about_item env = | SearchString (s,sc) -> try let ref = - Notation.interp_notation_as_global_reference Loc.ghost + Notation.interp_notation_as_global_reference (fun _ -> true) s sc in GlobSearchSubPattern (Pattern.PRef ref) with UserError _ -> @@ -2065,15 +2065,15 @@ let interp ?proof ~loc locality poly c = | VernacShow s -> vernac_show s | VernacCheckGuard -> vernac_check_guard () | VernacProof (None, None) -> - Aux_file.record_in_aux_at loc "VernacProof" "tac:no using:no" + Aux_file.record_in_aux_at ~loc "VernacProof" "tac:no using:no" | VernacProof (Some tac, None) -> - Aux_file.record_in_aux_at loc "VernacProof" "tac:yes using:no"; + Aux_file.record_in_aux_at ~loc "VernacProof" "tac:yes using:no"; vernac_set_end_tac tac | VernacProof (None, Some l) -> - Aux_file.record_in_aux_at loc "VernacProof" "tac:no using:yes"; + Aux_file.record_in_aux_at ~loc "VernacProof" "tac:no using:yes"; vernac_set_used_variables l | VernacProof (Some tac, Some l) -> - Aux_file.record_in_aux_at loc "VernacProof" "tac:yes using:yes"; + Aux_file.record_in_aux_at ~loc "VernacProof" "tac:yes using:yes"; vernac_set_end_tac tac; vernac_set_used_variables l | VernacProofMode mn -> Proof_global.set_proof_mode mn diff --git a/vernac/vernacentries.mli b/vernac/vernacentries.mli index fb2899515..f75f7656d 100644 --- a/vernac/vernacentries.mli +++ b/vernac/vernacentries.mli @@ -18,7 +18,7 @@ val vernac_require : val interp : ?verbosely:bool -> ?proof:Proof_global.closed_proof -> - Loc.t * Vernacexpr.vernac_expr -> unit + Vernacexpr.vernac_expr Loc.located -> unit (** Prepare a "match" template for a given inductive type. For each branch of the match, we list the constructor name -- cgit v1.2.3 From d062642d6e3671bab8a0e6d70e346325558d2db3 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Sat, 8 Apr 2017 23:34:06 +0200 Subject: [location] Be consistent with type module qualification Thanks to @gasche --- intf/constrexpr.mli | 39 +++++++++++++++++++-------------------- 1 file changed, 19 insertions(+), 20 deletions(-) (limited to 'intf') diff --git a/intf/constrexpr.mli b/intf/constrexpr.mli index 92d0020a5..23223173e 100644 --- a/intf/constrexpr.mli +++ b/intf/constrexpr.mli @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Loc open Names open Libnames open Misctypes @@ -62,7 +61,7 @@ type cases_pattern_expr_r = | CPatRecord of (reference * cases_pattern_expr) list | CPatDelimiters of string * cases_pattern_expr | CPatCast of cases_pattern_expr * constr_expr -and cases_pattern_expr = cases_pattern_expr_r located +and cases_pattern_expr = cases_pattern_expr_r Loc.located and cases_pattern_notation_substitution = cases_pattern_expr list * (** for constr subterms *) @@ -70,14 +69,14 @@ and cases_pattern_notation_substitution = 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 + | CFix of Id.t Loc.located * fix_expr list + | CCoFix of Id.t Loc.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 + | CLetIn of Name.t Loc.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 + (constr_expr * explicitation Loc.located option) list | CRecord of (reference * constr_expr) list (* representation of the "let" and "match" constructs *) @@ -86,9 +85,9 @@ and constr_expr_r = * case_expr list * branch_expr list (* branches *) - | CLetTuple of Name.t located list * (Name.t located option * constr_expr option) * + | CLetTuple of Name.t Loc.located list * (Name.t Loc.located option * constr_expr option) * constr_expr * constr_expr - | CIf of constr_expr * (Name.t located option * constr_expr option) + | CIf of constr_expr * (Name.t Loc.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 @@ -99,24 +98,24 @@ and constr_expr_r = | 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 constr_expr = constr_expr_r Loc.located and case_expr = constr_expr (* expression that is being matched *) - * Name.t located option (* as-clause *) + * Name.t Loc.located option (* as-clause *) * cases_pattern_expr option (* in-clause *) and branch_expr = - (cases_pattern_expr list located list * constr_expr) Loc.located + (cases_pattern_expr list Loc.located list * constr_expr) Loc.located and binder_expr = - Name.t located list * binder_kind * constr_expr + Name.t Loc.located list * binder_kind * constr_expr and fix_expr = - Id.t located * (Id.t located option * recursion_order_expr) * + Id.t Loc.located * (Id.t Loc.located option * recursion_order_expr) * local_binder_expr list * constr_expr * constr_expr and cofix_expr = - Id.t located * local_binder_expr list * constr_expr * constr_expr + Id.t Loc.located * local_binder_expr list * constr_expr * constr_expr and recursion_order_expr = | CStructRec @@ -125,8 +124,8 @@ and recursion_order_expr = (** Anonymous defs allowed ?? *) and local_binder_expr = - | CLocalAssum of Name.t located list * binder_kind * constr_expr - | CLocalDef of Name.t located * constr_expr * constr_expr option + | CLocalAssum of Name.t Loc.located list * binder_kind * constr_expr + | CLocalDef of Name.t Loc.located * constr_expr * constr_expr option | CLocalPattern of (cases_pattern_expr * constr_expr option) Loc.located and constr_notation_substitution = @@ -134,7 +133,7 @@ and constr_notation_substitution = constr_expr list list * (** for recursive notations *) local_binder_expr list list (** for binders subexpressions *) -type typeclass_constraint = (Name.t located * Id.t located list option) * binding_kind * constr_expr +type typeclass_constraint = (Name.t Loc.located * Id.t Loc.located list option) * binding_kind * constr_expr and typeclass_context = typeclass_constraint list @@ -143,11 +142,11 @@ type constr_pattern_expr = constr_expr (** Concrete syntax for modules and module types *) type with_declaration_ast = - | CWith_Module of Id.t list located * qualid located - | CWith_Definition of Id.t list located * constr_expr + | CWith_Module of Id.t list Loc.located * qualid Loc.located + | CWith_Definition of Id.t list Loc.located * constr_expr type module_ast_r = | CMident of qualid | CMapply of module_ast * module_ast | CMwith of module_ast * with_declaration_ast -and module_ast = module_ast_r located +and module_ast = module_ast_r Loc.located -- cgit v1.2.3 From 054d2736c1c1b55cb7708ff0444af521cd0fe2ba Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Sat, 8 Apr 2017 23:19:35 +0200 Subject: [location] [ast] Switch Constrexpr AST to an extensible node type. Following @gasche idea, and the original intention of #402, we switch the main parsing AST of Coq from `'a Loc.located` to `'a CAst.ast` which is private and record-based. This provides significantly clearer code for the AST, and is robust wrt attributes. --- ide/texmacspp.ml | 16 ++-- interp/constrexpr_ops.ml | 61 +++++++------ interp/constrextern.ml | 87 +++++++++--------- interp/constrintern.ml | 133 ++++++++++++++-------------- interp/implicit_quantifiers.ml | 28 +++--- interp/notation.ml | 2 +- interp/topconstr.ml | 16 ++-- intf/constrexpr.mli | 6 +- lib/cAst.ml | 24 +++++ lib/cAst.mli | 22 +++++ lib/clib.mllib | 1 + parsing/egramcoq.ml | 16 ++-- parsing/g_constr.ml4 | 150 ++++++++++++++++---------------- parsing/g_proofs.ml4 | 2 +- parsing/g_vernac.ml4 | 12 +-- plugins/funind/glob_term_to_relation.ml | 23 ++--- plugins/funind/indfun.ml | 24 ++--- plugins/funind/merge.ml | 4 +- plugins/ltac/g_ltac.ml4 | 8 +- plugins/ltac/g_tactic.ml4 | 12 +-- plugins/ltac/pptactic.ml | 2 +- plugins/ltac/rewrite.ml | 12 +-- plugins/ltac/tacintern.ml | 10 +-- plugins/ltac/tacinterp.ml | 2 +- plugins/ssrmatching/ssrmatching.ml4 | 22 ++--- printing/ppconstr.ml | 58 ++++++------ printing/ppvernac.ml | 4 +- vernac/classes.ml | 8 +- vernac/command.ml | 8 +- vernac/metasyntax.ml | 4 +- vernac/record.ml | 6 +- 31 files changed, 423 insertions(+), 360 deletions(-) create mode 100644 lib/cAst.ml create mode 100644 lib/cAst.mli (limited to 'intf') diff --git a/ide/texmacspp.ml b/ide/texmacspp.ml index 071861e27..ddb62313f 100644 --- a/ide/texmacspp.ml +++ b/ide/texmacspp.ml @@ -218,8 +218,8 @@ let rec pp_bindlist bl = (List.map (fun (loc, name) -> xmlCst ?loc (string_of_name name)) loc_names) in - match e with - | _, CHole _ -> names + match e.CAst.v with + | CHole _ -> names | _ -> names @ [pp_expr e]) bl) in match tlist with @@ -232,7 +232,7 @@ and pp_local_binder lb = (* don't know what it is for now *) | CLocalDef ((loc, nam), ce, ty) -> let attrs = ["name", string_of_name nam] in let value = match ty with - Some t -> Loc.tag ?loc:(Loc.merge_opt (constr_loc ce) (constr_loc t)) @@ CCast (ce, CastConv t) + Some t -> CAst.make ?loc:(Loc.merge_opt (constr_loc ce) (constr_loc t)) @@ CCast (ce, CastConv t) | None -> ce in pp_expr ~attr:attrs value | CLocalAssum (namll, _, ce) -> @@ -302,7 +302,7 @@ and pp_cofixpoint_expr (((loc, id), pl), lbl, ce, ceo) = (* cofixpoint_expr *) end and pp_lident (loc, id) = xmlCst ?loc (Id.to_string id) and pp_simple_binder (idl, ce) = List.map pp_lident idl @ [pp_expr ce] -and pp_cases_pattern_expr (loc, cpe) = +and pp_cases_pattern_expr {loc ; CAst.v = cpe} = match cpe with | CPatAlias (cpe, id) -> xmlApply ?loc @@ -390,7 +390,7 @@ 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=[]) (loc, e) = +and pp_expr ?(attr=[]) { loc; CAst.v = e } = match e with | CRef (r, _) -> xmlCst ?loc:(Libnames.loc_of_reference r) ~attr (Libnames.string_of_reference r) @@ -469,13 +469,13 @@ and pp_expr ?(attr=[]) (loc, e) = [pp_branch_expr_list bel])) | CRecord _ -> assert false | CLetIn ((varloc, var), value, typ, body) -> - let (loc, value) = match typ with + let value = match typ with | Some t -> - Loc.tag ?loc:(Loc.merge_opt (constr_loc value) (constr_loc t)) (CCast (value, CastConv t)) + CAst.make ?loc:(Loc.merge_opt (constr_loc value) (constr_loc t)) (CCast (value, CastConv t)) | None -> value in xmlApply ?loc (xmlOperator ?loc "let" :: - [xmlCst ?loc:varloc (string_of_name var) ; pp_expr (Loc.tag ?loc value); pp_expr body]) + [xmlCst ?loc:varloc (string_of_name var) ; pp_expr value; pp_expr body]) | CLambdaN (bl, e) -> xmlApply ?loc (xmlOperator ?loc "lambda" :: [pp_bindlist bl] @ [pp_expr e]) diff --git a/interp/constrexpr_ops.ml b/interp/constrexpr_ops.ml index ce349a63f..b11972cd3 100644 --- a/interp/constrexpr_ops.ml +++ b/interp/constrexpr_ops.ml @@ -59,9 +59,9 @@ let explicitation_eq ex1 ex2 = match ex1, ex2 with let eq_located f (_, x) (_, y) = f x y -let rec cases_pattern_expr_eq (l1, p1) (l2, p2) = - if p1 == p2 then true - else match p1, p2 with +let rec cases_pattern_expr_eq p1 p2 = + if CAst.(p1.v == p2.v) then true + else match CAst.(p1.v, p2.v) with | CPatAlias(a1,i1), CPatAlias(a2,i2) -> Id.equal i1 i2 && cases_pattern_expr_eq a1 a2 | CPatCstr(c1,a1,b1), CPatCstr(c2,a2,b2) -> @@ -97,9 +97,9 @@ let eq_universes u1 u2 = | Some l, Some l' -> l = l' | _, _ -> false -let rec constr_expr_eq (_loc1, e1) (_loc2, e2) = - if e1 == e2 then true - else match e1, e2 with +let rec constr_expr_eq e1 e2 = + if CAst.(e1.v == e2.v) then true + else match CAst.(e1.v, e2.v) with | CRef (r1,u1), CRef (r2,u2) -> eq_reference r1 r2 && eq_universes u1 u2 | CFix(id1,fl1), CFix(id2,fl2) -> eq_located Id.equal id1 id2 && @@ -228,11 +228,11 @@ 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 (l,_) = l +let constr_loc c = CAst.(c.loc) -let cases_pattern_expr_loc (l,_) = l +let cases_pattern_expr_loc cp = CAst.(cp.loc) -let raw_cases_pattern_expr_loc (l, _) = l +let raw_cases_pattern_expr_loc pe = CAst.(pe.loc) let local_binder_loc = function | CLocalAssum ((loc,_)::_,_,t) @@ -247,18 +247,18 @@ let local_binders_loc bll = match bll with (** Pseudo-constructors *) -let mkIdentC id = Loc.tag @@ CRef (Ident (Loc.tag 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 mkIdentC id = CAst.make @@ CRef (Ident (Loc.tag id),None) +let mkRefC r = CAst.make @@ CRef (r,None) +let mkCastC (a,k) = CAst.make @@ CCast (a,k) +let mkLambdaC (idl,bk,a,b) = CAst.make @@ CLambdaN ([idl,bk,a],b) +let mkLetInC (id,a,t,b) = CAst.make @@ CLetIn (id,a,t,b) +let mkProdC (idl,bk,a,b) = CAst.make @@ CProdN ([idl,bk,a],b) let mkAppC (f,l) = let l = List.map (fun x -> (x,None)) l in - match f with - | _loc, CApp (g,l') -> Loc.tag @@ CApp (g, l' @ l) - | _ -> Loc.tag @@ CApp ((None, f), l) + match CAst.(f.v) with + | CApp (g,l') -> CAst.make @@ CApp (g, l' @ l) + | _ -> CAst.make @@ CApp ((None, f), l) let add_name_in_env env n = match snd n with @@ -276,7 +276,7 @@ let expand_binders ?loc mkC bl c = | CLocalDef ((loc1,_) as n, oty, b) -> let env, c = loop ?loc:(Loc.merge_opt loc1 loc) bl c in let env = add_name_in_env env n in - (env, Loc.tag ?loc @@ CLetIn (n,oty,b,c)) + (env, CAst.make ?loc @@ CLetIn (n,oty,b,c)) | CLocalAssum ((loc1,_)::_ as nl, bk, t) -> let env, c = loop ?loc:(Loc.merge_opt loc1 loc) bl c in let env = List.fold_left add_name_in_env env nl in @@ -288,10 +288,10 @@ let expand_binders ?loc mkC bl c = let id = (loc1, Name ni) in let ty = match ty with | Some ty -> ty - | None -> Loc.tag ?loc:loc1 @@ CHole (None, IntroAnonymous, None) + | None -> CAst.make ?loc:loc1 @@ CHole (None, IntroAnonymous, None) in - let e = Loc.tag @@ CRef (Libnames.Ident (loc1, ni), None) in - let c = Loc.tag ?loc @@ + let e = CAst.make @@ CRef (Libnames.Ident (loc1, ni), None) in + let c = CAst.make ?loc @@ CCases (LetPatternStyle, None, [(e,None,None)], [(Loc.tag ?loc:loc1 ([(loc1,[p])], c))]) @@ -302,11 +302,11 @@ let expand_binders ?loc mkC bl c = c let mkCProdN ?loc bll c = - let mk ?loc b c = Loc.tag ?loc @@ CProdN ([b],c) in + let mk ?loc b c = CAst.make ?loc @@ CProdN ([b],c) in expand_binders ?loc mk bll c let mkCLambdaN ?loc bll c = - let mk ?loc b c = Loc.tag ?loc @@ CLambdaN ([b],c) in + let mk ?loc b c = CAst.make ?loc @@ CLambdaN ([b],c) in expand_binders ?loc mk bll c (* Deprecated *) @@ -320,14 +320,13 @@ let coerce_reference_to_id = function (str "This expression should be a simple identifier.") let coerce_to_id = function - | _loc, CRef (Ident (loc,id),_) -> (loc,id) - | a -> CErrors.user_err ?loc:(constr_loc a) + | { CAst.v = CRef (Ident (loc,id),_); _ } -> (loc,id) + | { CAst.loc; _ } -> CErrors.user_err ?loc ~hdr:"coerce_to_id" (str "This expression should be a simple identifier.") let coerce_to_name = function - | _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.") + | { CAst.v = CRef (Ident (loc,id),_) } -> (loc,Name id) + | { CAst.loc; CAst.v = CHole (_,_,_) } -> (loc,Anonymous) + | { CAst.loc; _ } -> CErrors.user_err ?loc ~hdr:"coerce_to_name" + (str "This expression should be a name.") diff --git a/interp/constrextern.ml b/interp/constrextern.ml index 30b81ecc4..e8a5b5265 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -144,15 +144,15 @@ module PrintingConstructor = Goptions.MakeRefTable(PrintingRecordConstructor) let insert_delimiters e = function | None -> e - | Some sc -> Loc.tag @@ CDelimiters (sc,e) + | Some sc -> CAst.make @@ CDelimiters (sc,e) let insert_pat_delimiters ?loc p = function | None -> p - | Some sc -> Loc.tag ?loc @@ CPatDelimiters (sc,p) + | Some sc -> CAst.make ?loc @@ CPatDelimiters (sc,p) let insert_pat_alias ?loc p = function | Anonymous -> p - | Name id -> Loc.tag ?loc @@ CPatAlias (p,id) + | Name id -> CAst.make ?loc @@ CPatAlias (p,id) (**********************************************************************) (* conversion of references *) @@ -178,7 +178,7 @@ let extern_reference ?loc vars l = !my_extern_reference ?loc vars l let add_patt_for_params ind l = if !Flags.in_debugger then l else - Util.List.addn (Inductiveops.inductive_nparamdecls ind) (Loc.tag @@ CPatAtom None) l + Util.List.addn (Inductiveops.inductive_nparamdecls ind) (CAst.make @@ CPatAtom None) l let add_cpatt_for_params ind l = if !Flags.in_debugger then l else @@ -190,7 +190,7 @@ let drop_implicits_in_patt cst nb_expl args = let rec impls_fit l = function |[],t -> Some (List.rev_append l t) |_,[] -> None - |h::t,(_loc, CPatAtom None)::tt when is_status_implicit h -> impls_fit l (t,tt) + |h::t, { CAst.v = CPatAtom None }::tt when is_status_implicit h -> impls_fit l (t,tt) |h::_,_ when is_status_implicit h -> None |_::t,hh::tt -> impls_fit (hh::l) (t,tt) in let rec aux = function @@ -236,8 +236,8 @@ let expand_curly_brackets loc mknot ntn l = (* side effect *) mknot (loc,!ntn',l) -let destPrim = function _loc, CPrim t -> Some t | _ -> None -let destPatPrim = function _loc, CPatPrim t -> Some t | _ -> None +let destPrim = function { CAst.v = CPrim t } -> Some t | _ -> None +let destPatPrim = function { CAst.v = CPatPrim t } -> Some t | _ -> None let make_notation_gen loc ntn mknot mkprim destprim l = if has_curly_brackets ntn @@ -259,21 +259,21 @@ 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 - Loc.tag ?loc @@ CNotation (ntn,subst) + CAst.make ?loc @@ CNotation (ntn,subst) else make_notation_gen loc ntn - (fun (loc,ntn,l) -> Loc.tag ?loc @@ CNotation (ntn,(l,[],[]))) - (fun (loc,p) -> Loc.tag ?loc @@ CPrim p) + (fun (loc,ntn,l) -> CAst.make ?loc @@ CNotation (ntn,(l,[],[]))) + (fun (loc,p) -> CAst.make ?loc @@ CPrim p) destPrim terms let make_pat_notation ?loc ntn (terms,termlists as subst) args = - if not (List.is_empty termlists) then (Loc.tag ?loc @@ CPatNotation (ntn,subst,args)) else + if not (List.is_empty termlists) then (CAst.make ?loc @@ CPatNotation (ntn,subst,args)) else make_notation_gen loc ntn - (fun (loc,ntn,l) -> Loc.tag ?loc @@ CPatNotation (ntn,(l,[]),args)) - (fun (loc,p) -> Loc.tag ?loc @@ CPatPrim p) + (fun (loc,ntn,l) -> CAst.make ?loc @@ CPatNotation (ntn,(l,[]),args)) + (fun (loc,p) -> CAst.make ?loc @@ CPatPrim p) destPatPrim terms -let mkPat ?loc qid l = Loc.tag ?loc @@ +let mkPat ?loc qid l = CAst.make ?loc @@ (* Normally irrelevant test with v8 syntax, but let's do it anyway *) if List.is_empty l then CPatAtom (Some qid) else CPatCstr (qid,None,l) @@ -295,7 +295,7 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = when !Flags.in_debugger||Inductiveops.constructor_has_local_defs cstrsp -> let c = extern_reference ?loc Id.Set.empty (ConstructRef cstrsp) in let args = List.map (extern_cases_pattern_in_scope scopes vars) args in - Loc.tag ?loc @@ CPatCstr (c, Some (add_patt_for_params (fst cstrsp) args), []) + CAst.make ?loc @@ CPatCstr (c, Some (add_patt_for_params (fst cstrsp) args), []) | _ -> try if !Flags.raw_print || !print_no_symbol then raise No_match; @@ -304,7 +304,7 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = | None -> raise No_match | Some key -> let loc = cases_pattern_loc pat in - insert_pat_alias ?loc (insert_pat_delimiters ?loc (Loc.tag ?loc @@ CPatPrim p) key) na + insert_pat_alias ?loc (insert_pat_delimiters ?loc (CAst.make ?loc @@ CPatPrim p) key) na with No_match -> try if !Flags.raw_print || !print_no_symbol then raise No_match; @@ -312,8 +312,8 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = (uninterp_cases_pattern_notations pat) with No_match -> match pat with - | loc, PatVar (Name id) -> Loc.tag ?loc @@ CPatAtom (Some (Ident (loc,id))) - | loc, PatVar (Anonymous) -> Loc.tag ?loc @@ CPatAtom None + | loc, PatVar (Name id) -> CAst.make ?loc @@ CPatAtom (Some (Ident (loc,id))) + | loc, PatVar (Anonymous) -> CAst.make ?loc @@ CPatAtom None | loc, PatCstr(cstrsp,args,na) -> let args = List.map (extern_cases_pattern_in_scope scopes vars) args in let p = @@ -327,24 +327,29 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = | Some c :: q -> match args with | [] -> raise No_match - | (_loc, CPatAtom(None)) :: tail -> ip q tail acc + + + + + + | { CAst.v = CPatAtom None } :: tail -> ip q tail acc (* we don't want to have 'x = _' in our patterns *) | head :: tail -> ip q tail ((extern_reference ?loc Id.Set.empty (ConstRef c), head) :: acc) in - Loc.tag ?loc @@ CPatRecord(List.rev (ip projs args [])) + CAst.make ?loc @@ CPatRecord(List.rev (ip projs args [])) with Not_found | No_match | Exit -> let c = extern_reference ?loc Id.Set.empty (ConstructRef cstrsp) in if !Topconstr.asymmetric_patterns then if pattern_printable_in_both_syntax cstrsp - then Loc.tag ?loc @@ CPatCstr (c, None, args) - else Loc.tag ?loc @@ CPatCstr (c, Some (add_patt_for_params (fst cstrsp) args), []) + then CAst.make ?loc @@ CPatCstr (c, None, args) + else CAst.make ?loc @@ CPatCstr (c, Some (add_patt_for_params (fst cstrsp) args), []) else let full_args = add_patt_for_params (fst cstrsp) args in match drop_implicits_in_patt (ConstructRef cstrsp) 0 full_args with - | Some true_args -> Loc.tag ?loc @@ CPatCstr (c, None, true_args) - | None -> Loc.tag ?loc @@ CPatCstr (c, Some full_args, []) + | Some true_args -> CAst.make ?loc @@ CPatCstr (c, None, true_args) + | None -> CAst.make ?loc @@ CPatCstr (c, Some full_args, []) in insert_pat_alias ?loc p na and apply_notation_to_pattern ?loc gr ((subst,substlist),(nb_to_drop,more_args)) (tmp_scope, scopes as allscopes) vars = @@ -401,8 +406,8 @@ and extern_notation_pattern (tmp_scope,scopes as allscopes) vars (loc, t) = func let p = apply_notation_to_pattern ?loc (ConstructRef cstr) (match_notation_constr_cases_pattern (loc, t) pat) allscopes vars keyrule in insert_pat_alias ?loc p na - | PatVar Anonymous -> Loc.tag ?loc @@ CPatAtom None - | PatVar (Name id) -> Loc.tag ?loc @@ CPatAtom (Some (Ident (loc,id))) + | PatVar Anonymous -> CAst.make ?loc @@ CPatAtom None + | PatVar (Name id) -> CAst.make ?loc @@ CPatAtom (Some (Ident (loc,id))) with No_match -> extern_notation_pattern allscopes vars (loc, t) rules @@ -422,7 +427,7 @@ let extern_ind_pattern_in_scope (scopes:local_scopes) vars ind args = if !Flags.in_debugger||Inductiveops.inductive_has_local_defs ind then let c = extern_reference vars (IndRef ind) in let args = List.map (extern_cases_pattern_in_scope scopes vars) args in - Loc.tag @@ CPatCstr (c, Some (add_patt_for_params ind args), []) + CAst.make @@ CPatCstr (c, Some (add_patt_for_params ind args), []) else try if !Flags.raw_print || !print_no_symbol then raise No_match; @@ -430,7 +435,7 @@ let extern_ind_pattern_in_scope (scopes:local_scopes) vars ind args = match availability_of_prim_token p sc scopes with | None -> raise No_match | Some key -> - insert_pat_delimiters (Loc.tag @@ CPatPrim p) key + insert_pat_delimiters (CAst.make @@ CPatPrim p) key with No_match -> try if !Flags.raw_print || !print_no_symbol then raise No_match; @@ -440,8 +445,8 @@ let extern_ind_pattern_in_scope (scopes:local_scopes) vars ind args = let c = extern_reference vars (IndRef ind) in let args = List.map (extern_cases_pattern_in_scope scopes vars) args in match drop_implicits_in_patt (IndRef ind) 0 args with - |Some true_args -> Loc.tag @@ CPatCstr (c, None, true_args) - |None -> Loc.tag @@ CPatCstr (c, Some args, []) + |Some true_args -> CAst.make @@ CPatCstr (c, None, true_args) + |None -> CAst.make @@ CPatCstr (c, Some args, []) let extern_cases_pattern vars p = extern_cases_pattern_in_scope (None,[]) vars p @@ -466,7 +471,7 @@ let is_projection nargs = function let is_hole = function CHole _ | CEvar _ -> true | _ -> false let is_significant_implicit a = - not (is_hole (snd a)) + not (is_hole (a.CAst.v)) let is_needed_for_correct_partial_application tail imp = List.is_empty tail && not (maximal_insertion_of imp) @@ -515,11 +520,11 @@ let explicitize inctx impl (cf,f) args = CApp ((ip,f),args1@args2) | None -> let args = exprec 1 (args,impl) in - if List.is_empty args then snd f else CApp ((None, f), args) + if List.is_empty args then f.CAst.v else CApp ((None, f), args) in try expl () with Expl -> - let f',us = match f with (_loc, CRef (f,us)) -> f,us | _ -> assert false in + let f',us = match f with { CAst.v = CRef (f,us) } -> f,us | _ -> assert false in let ip = if !print_projections then ip else None in CAppExpl ((ip, f', us), List.map Lazy.force args) @@ -546,7 +551,7 @@ let extern_app inctx impl (cf,f) us args = let args = List.map Lazy.force args in CAppExpl ((is_projection (List.length args) cf,f,us), args) else - explicitize inctx impl (cf, Loc.tag @@ CRef (f,us)) args + explicitize inctx impl (cf, CAst.make @@ CRef (f,us)) args let rec fill_arg_scopes args subscopes scopes = match args, subscopes with | [], _ -> [] @@ -600,7 +605,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 (Loc.tag ?loc:(loc_of_glob_constr r) @@ CPrim n) key) + | Some key -> Some (insert_delimiters (CAst.make ?loc:(loc_of_glob_constr r) @@ CPrim n) key) with No_match -> None @@ -608,7 +613,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 ({ CAst.v = CDelimiters _}) | None) | _, Some n -> n | _ -> raise No_match (**********************************************************************) @@ -644,7 +649,7 @@ let rec extern inctx scopes vars r = let r'' = flatten_application r' in if !Flags.raw_print || !print_no_symbol then raise No_match; extern_notation scopes vars r'' (uninterp_notations r'') - with No_match -> Loc.map_with_loc (fun ?loc -> function + with No_match -> CAst.map_from_loc (fun ?loc -> function | GRef (ref,us) -> extern_global (select_stronger_impargs (implicits_of_global ref)) (extern_reference ?loc vars ref) (extern_universes us) @@ -824,7 +829,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, (loc, CProdN ([nal,Default bk',ty],c)) + | Name id, { CAst.loc ; v = 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 @@ -834,7 +839,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 - | loc, CLambdaN ([nal,Default bk',ty],c) + | { CAst.loc; v = 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 @@ -943,12 +948,12 @@ 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 - Loc.tag ?loc @@ if List.is_empty l then a else CApp ((None, Loc.tag a),l) in + CAst.make ?loc @@ if List.is_empty l then a else CApp ((None, CAst.make a),l) in if List.is_empty args then e else let args = fill_arg_scopes args argsscopes scopes in let args = extern_args (extern true) vars args in - Loc.tag ?loc @@ explicitize false argsimpls (None,e) args + CAst.make ?loc @@ explicitize false argsimpls (None,e) args with No_match -> extern_notation allscopes vars t rules diff --git a/interp/constrintern.ml b/interp/constrintern.ml index a672771b1..541b52972 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 | [] -> [] - | (_loc, CNotation ("{ _ }",([a],[],[]))) :: l -> + | { CAst.v = CNotation ("{ _ }",([a],[],[])) } :: l -> ntn' := expand_notation_string !ntn' n; contract_squash n (a::l) | a :: l -> @@ -230,7 +230,7 @@ let contract_pat_notation ntn (l,ll) = let ntn' = ref ntn in let rec contract_squash n = function | [] -> [] - | (_, CPatNotation ("{ _ }",([a],[]),[])) :: l -> + | { CAst.v = CPatNotation ("{ _ }",([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 + | { CAst.v = CApp ((_, { CAst.v = CRef (Ident (loc,id),_) } ), _) } -> id | _ -> default_non_dependent_ident in Implicit_quantifiers.make_fresh ids' (Global.env ()) id in Name name @@ -430,7 +430,7 @@ let intern_assumption intern lvar env nal bk ty = let env, b = intern_generalized_binder intern_type lvar env (List.hd nal) b b' t ty in env, b -let rec free_vars_of_pat il (loc, pt) = match pt with +let rec free_vars_of_pat il pt = match CAst.(pt.v) with | CPatCstr (c, l1, l2) -> let il = List.fold_left free_vars_of_pat il (Option.default [] l1) in List.fold_left free_vars_of_pat il l2 @@ -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 = Loc.tag ?loc @@ CHole (None, Misctypes.IntroAnonymous, None) in + let ty = CAst.make ?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) @@ -480,7 +480,7 @@ let intern_local_binder_aux ?(global_level=false) intern lvar (env,bl) = functio let tyc = match ty with | Some ty -> ty - | None -> Loc.tag ?loc @@ CHole(None,Misctypes.IntroAnonymous,None) + | None -> CAst.make ?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 @@ -596,16 +596,16 @@ let rec subordinate_letins letins = function letins,[] let terms_of_binders bl = - let rec term_of_pat pt = Loc.map_with_loc (fun ?loc -> function + let rec term_of_pat pt = CAst.map_from_loc (fun ?loc -> function | PatVar (Name id) -> CRef (Ident (loc,id), None) | PatVar (Anonymous) -> error "Cannot turn \"_\" into a term." | PatCstr (c,l,_) -> let r = Qualid (loc,qualid_of_path (path_of_global (ConstructRef c))) in - let hole = Loc.tag ?loc @@ CHole (None,Misctypes.IntroAnonymous,None) in + let hole = CAst.make ?loc @@ CHole (None,Misctypes.IntroAnonymous,None) in let params = List.make (Inductiveops.inductive_nparams (fst c)) hole in CAppExpl ((None,r,None),params @ List.map term_of_pat l)) pt in let rec extract_variables = function - | (loc, GLocalAssum (Name id,_,_))::l -> (Loc.tag ?loc @@ CRef (Ident (loc,id), None)) :: extract_variables l + | (loc, GLocalAssum (Name id,_,_))::l -> (CAst.make ?loc @@ CRef (Ident (loc,id), None)) :: extract_variables l | (loc, GLocalDef (Name id,_,_,_))::l -> extract_variables l | (loc, GLocalDef (Anonymous,_,_,_))::l | (loc, GLocalAssum (Anonymous,_,_))::l -> error "Cannot turn \"_\" into a term." @@ -760,7 +760,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 -> Loc.tag ?loc @@ CRef (Ident (loc,id),None), Some (loc,ExplByName id)) expl_impls in + (fun id -> CAst.make ?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 @@ -993,10 +993,10 @@ let add_implicits_check_length fail nargs nargs_with_letin impls_st len_pl1 pl2 else Int.equal args_len nargs_with_letin || (fst (fail (nargs - List.length impl_list + i)))) ,l) |imp::q as il,[] -> if is_status_implicit imp && maximal_insertion_of imp - then let (b,out) = aux i (q,[]) in (b,(Loc.tag @@ RCPatAtom(None))::out) + then let (b,out) = aux i (q,[]) in (b,(CAst.make @@ RCPatAtom(None))::out) else fail (remaining_args (len_pl1+i) il) |imp::q,(hh::tt as l) -> if is_status_implicit imp - then let (b,out) = aux i (q,l) in (b,(Loc.tag @@ RCPatAtom(None))::out) + then let (b,out) = aux i (q,l) in (b,(CAst.make @@ RCPatAtom(None))::out) else let (b,out) = aux (succ i) (q,tt) in (b,hh::out) in aux 0 (impl_list,pl2) @@ -1199,14 +1199,15 @@ let alias_of als = match als.alias_ids with *) -let rec subst_pat_iterator y t (loc, p) = match p with - | RCPatAtom id -> - begin match id with Some x when Id.equal x y -> t | _ -> Loc.tag ?loc p end +let rec subst_pat_iterator y t = CAst.map (function + | RCPatAtom id as p -> + begin match id with Some x when Id.equal x y -> t.CAst.v | _ -> p end | RCPatCstr (id,l1,l2) -> - Loc.tag ?loc @@ RCPatCstr (id, List.map (subst_pat_iterator y t) l1, - List.map (subst_pat_iterator y t) l2) - | RCPatAlias (p,a) -> Loc.tag ?loc @@ RCPatAlias (subst_pat_iterator y t p,a) - | RCPatOr pl -> Loc.tag ?loc @@ RCPatOr (List.map (subst_pat_iterator y t) pl) + RCPatCstr (id, List.map (subst_pat_iterator y t) l1, + List.map (subst_pat_iterator y t) l2) + | RCPatAlias (p,a) -> RCPatAlias (subst_pat_iterator y t p,a) + | RCPatOr pl -> RCPatOr (List.map (subst_pat_iterator y t) pl) + ) let drop_notations_pattern looked_for = (* At toplevel, Constructors and Inductives are accepted, in recursive calls @@ -1255,26 +1256,29 @@ let drop_notations_pattern looked_for = let (_,argscs) = find_remaining_scopes [] pats g in Some (g,[],List.map2 (fun x -> in_pat false (x,snd scopes)) argscs pats) with Not_found -> None - and in_pat top scopes (loc, pt) = match pt with - | CPatAlias (p, id) -> Loc.tag ?loc @@ RCPatAlias (in_pat top scopes p, id) + and in_pat top scopes pt = + let open CAst in + let loc = pt.loc in + match pt.v with + | CPatAlias (p, id) -> CAst.make ?loc @@ RCPatAlias (in_pat top scopes p, id) | CPatRecord l -> let sorted_fields = - sort_fields ~complete:false loc l (fun _idx -> (loc, CPatAtom None)) in + sort_fields ~complete:false loc l (fun _idx -> CAst.make ?loc @@ CPatAtom None) in begin match sorted_fields with - | None -> Loc.tag ?loc @@ RCPatAtom None + | None -> CAst.make ?loc @@ RCPatAtom None | Some (n, head, pl) -> let pl = if !asymmetric_patterns then pl else - let pars = List.make n (loc, CPatAtom None) in + let pars = List.make n (CAst.make ?loc @@ CPatAtom None) in List.rev_append pars pl in match drop_syndef top scopes head pl with - |Some (a,b,c) -> (loc, RCPatCstr(a, b, c)) - |None -> raise (InternalizationError (loc,NotAConstructor head)) + | Some (a,b,c) -> CAst.make ?loc @@ RCPatCstr(a, b, c) + | None -> raise (InternalizationError (loc,NotAConstructor head)) end | CPatCstr (head, None, pl) -> begin match drop_syndef top scopes head pl with - | Some (a,b,c) -> Loc.tag ?loc @@ RCPatCstr(a, b, c) + | Some (a,b,c) -> CAst.make ?loc @@ RCPatCstr(a, b, c) | None -> raise (InternalizationError (loc,NotAConstructor head)) end | CPatCstr (r, Some expl_pl, pl) -> @@ -1283,13 +1287,13 @@ let drop_notations_pattern looked_for = raise (InternalizationError (loc,NotAConstructor r)) in if expl_pl == [] then (* Convention: (@r) deactivates all further implicit arguments and scopes *) - Loc.tag ?loc @@ RCPatCstr (g, List.map (in_pat false scopes) pl, []) + CAst.make ?loc @@ RCPatCstr (g, List.map (in_pat false scopes) pl, []) else (* Convention: (@r expl_pl) deactivates implicit arguments in expl_pl and in pl *) (* but not scopes in expl_pl *) let (argscs1,_) = find_remaining_scopes expl_pl pl g in - Loc.tag ?loc @@ RCPatCstr (g, List.map2 (in_pat_sc scopes) argscs1 expl_pl @ List.map (in_pat false scopes) pl, []) - | CPatNotation ("- _",([_loc,CPatPrim(Numeral p)],[]),[]) + CAst.make ?loc @@ RCPatCstr (g, List.map2 (in_pat_sc scopes) argscs1 expl_pl @ List.map (in_pat false scopes) pl, []) + | CPatNotation ("- _",([{ CAst.v = CPatPrim(Numeral p) }],[]),[]) when Bigint.is_strictly_pos p -> fst (Notation.interp_prim_token_cases_pattern_expr ?loc (ensure_kind false loc) (Numeral (Bigint.neg p)) scopes) | CPatNotation ("( _ )",([a],[]),[]) -> @@ -1308,11 +1312,11 @@ let drop_notations_pattern looked_for = | CPatAtom Some id -> begin match drop_syndef top scopes id [] with - | Some (a,b,c) -> Loc.tag ?loc @@ RCPatCstr (a, b, c) - | None -> Loc.tag ?loc @@ RCPatAtom (Some (find_pattern_variable id)) + | Some (a,b,c) -> CAst.make ?loc @@ RCPatCstr (a, b, c) + | None -> CAst.make ?loc @@ RCPatAtom (Some (find_pattern_variable id)) end - | CPatAtom None -> Loc.tag ?loc @@ RCPatAtom None - | CPatOr pl -> Loc.tag ?loc @@ RCPatOr (List.map (in_pat top scopes) pl) + | CPatAtom None -> CAst.make ?loc @@ RCPatAtom None + | CPatOr pl -> CAst.make ?loc @@ RCPatOr (List.map (in_pat top scopes) pl) | CPatCast _ -> assert false and in_pat_sc scopes x = in_pat false (x,snd scopes) @@ -1326,17 +1330,17 @@ let drop_notations_pattern looked_for = let (a,(scopt,subscopes)) = Id.Map.find id subst in in_pat top (scopt,subscopes@snd scopes) a with Not_found -> - if Id.equal id ldots_var then Loc.tag ?loc @@ RCPatAtom (Some id) else + if Id.equal id ldots_var then CAst.make ?loc @@ RCPatAtom (Some id) else anomaly (str "Unbound pattern notation variable: " ++ Id.print id) end | NRef g -> ensure_kind top loc g; let (_,argscs) = find_remaining_scopes [] args g in - Loc.tag ?loc @@ RCPatCstr (g, [], List.map2 (in_pat_sc scopes) argscs args) + CAst.make ?loc @@ RCPatCstr (g, [], List.map2 (in_pat_sc scopes) argscs args) | NApp (NRef g,pl) -> ensure_kind top loc g; let (argscs1,argscs2) = find_remaining_scopes pl args g in - Loc.tag ?loc @@ RCPatCstr (g, + CAst.make ?loc @@ RCPatCstr (g, List.map2 (fun x -> in_not false loc (x,snd scopes) fullsubst []) argscs1 pl @ List.map (in_pat false scopes) args, []) | NList (x,y,iter,terminator,lassoc) -> @@ -1355,7 +1359,7 @@ let drop_notations_pattern looked_for = anomaly (Pp.str "Inconsistent substitution of recursive notation")) | NHole _ -> let () = assert (List.is_empty args) in - Loc.tag ?loc @@ RCPatAtom None + CAst.make ?loc @@ RCPatAtom None | t -> error_invalid_pattern_notation ?loc () in in_pat true @@ -1366,11 +1370,12 @@ let rec intern_pat genv aliases pat = let pl' = List.map (fun (asubst,pl) -> (asubst, Loc.tag ?loc @@ PatCstr (c,chop_params_pattern loc (fst c) pl with_letin,alias_of aliases))) pll in ids',pl' in - match pat with - | loc, RCPatAlias (p, id) -> + let loc = CAst.(pat.loc) in + match CAst.(pat.v) with + | RCPatAlias (p, id) -> let aliases' = merge_aliases aliases id in intern_pat genv aliases' p - | loc, RCPatCstr (head, expl_pl, pl) -> + | RCPatCstr (head, expl_pl, pl) -> if !asymmetric_patterns then let len = if List.is_empty expl_pl then Some (List.length pl) else None in let c,idslpl1 = find_constructor loc len head in @@ -1382,13 +1387,13 @@ let rec intern_pat genv aliases pat = let with_letin, pl2 = add_implicits_check_constructor_length genv loc c (List.length idslpl1 + List.length expl_pl) pl in intern_cstr_with_all_args loc c with_letin idslpl1 (expl_pl@pl2) - | loc, RCPatAtom (Some id) -> + | RCPatAtom (Some id) -> let aliases = merge_aliases aliases id in (aliases.alias_ids,[aliases.alias_map, Loc.tag ?loc @@ PatVar (alias_of aliases)]) - | loc, RCPatAtom (None) -> + | RCPatAtom (None) -> let { alias_ids = ids; alias_map = asubst; } = aliases in (ids, [asubst, Loc.tag ?loc @@ PatVar (alias_of aliases)]) - | loc, RCPatOr pl -> + | RCPatOr pl -> assert (not (List.is_empty pl)); let pl' = List.map (intern_pat genv aliases) pl in let (idsl,pl') = List.split pl' in @@ -1406,9 +1411,9 @@ let rec intern_pat genv aliases pat = of lambdas in the encoding of match in constr. We put this check here and not in the parser because it would require to duplicate the levels of the [pattern] rule. *) -let rec check_no_patcast (loc, pt) = match pt with +let rec check_no_patcast pt = match CAst.(pt.v) with | CPatCast (_,_) -> - CErrors.user_err ?loc ~hdr:"check_no_patcast" + CErrors.user_err ?loc:pt.CAst.loc ~hdr:"check_no_patcast" (Pp.strbrk "Casts are not supported here.") | CPatDelimiters(_,p) | CPatAlias(p,_) -> check_no_patcast p @@ -1444,8 +1449,9 @@ let intern_ind_pattern genv scopes pat = drop_notations_pattern (function (IndRef _ | ConstructRef _) -> () | _ -> raise Not_found) scopes pat with InternalizationError(loc,NotAConstructor _) -> error_bad_inductive_type ?loc in - match no_not with - | loc, RCPatCstr (head, expl_pl, pl) -> + let loc = no_not.CAst.loc in + match no_not.CAst.v with + | RCPatCstr (head, expl_pl, pl) -> let c = (function IndRef ind -> ind | _ -> error_bad_inductive_type ?loc) head in let with_letin, pl2 = add_implicits_check_ind_length genv loc c (List.length expl_pl) pl in @@ -1455,7 +1461,7 @@ let intern_ind_pattern genv scopes pat = match product_of_cases_patterns [] (List.rev_append idslpl1 idslpl2) with | _,[_,pl] -> (c,chop_params_pattern loc c pl with_letin) | _ -> error_bad_inductive_type ?loc) - | x -> error_bad_inductive_type ?loc:(raw_cases_pattern_expr_loc x) + | x -> error_bad_inductive_type ?loc (**********************************************************************) (* Utilities for application *) @@ -1521,7 +1527,7 @@ let extract_explicit_arg imps args = (* Main loop *) let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = - let rec intern env = Loc.with_loc (fun ?loc -> function + let rec intern env = CAst.with_loc_val (fun ?loc -> function | CRef (ref,us) -> let (c,imp,subscopes,l),_ = intern_applied_reference intern env (Environ.named_context globalenv) @@ -1602,20 +1608,20 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = | CProdN ([],c2) -> intern_type env c2 | CProdN ((nal,bk,ty)::bll,c2) -> - iterate_prod ?loc env bk ty (Loc.tag ?loc @@ CProdN (bll, c2)) nal + iterate_prod ?loc env bk ty (CAst.make ?loc @@ CProdN (bll, c2)) nal | CLambdaN ([],c2) -> intern env c2 | CLambdaN ((nal,bk,ty)::bll,c2) -> - iterate_lam loc (reset_tmp_scope env) bk ty (Loc.tag ?loc @@ CLambdaN (bll, c2)) nal + iterate_lam loc (reset_tmp_scope env) bk ty (CAst.make ?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 Loc.tag ?loc @@ GLetIn (snd na, inc1, int, intern (push_name_env ntnvars (impls_term_list inc1) env na) c2) - | CNotation ("- _",([_, CPrim (Numeral p)],[],[])) + | CNotation ("- _",([{ CAst.v = CPrim (Numeral p) }],[],[])) when Bigint.is_strictly_pos p -> - intern env (Loc.tag ?loc @@ CPrim (Numeral (Bigint.neg p))) + intern env (CAst.make ?loc @@ CPrim (Numeral (Bigint.neg p))) | CNotation ("( _ )",([a],[],[])) -> intern env a | CNotation (ntn,args) -> intern_notation intern env ntnvars loc ntn args @@ -1639,20 +1645,20 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = | CApp ((isproj,f), args) -> let f,args = match f with (* Compact notations like "t.(f args') args" *) - | _loc, CApp ((Some _,f), args') when not (Option.has_some isproj) -> + | { CAst.v = 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 - | _loc, CRef (ref,us) -> + match f.CAst.v with + | CRef (ref,us) -> intern_applied_reference intern env (Environ.named_context globalenv) lvar us args ref - | _loc, CNotation (ntn,([],[],[])) -> + | 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 + | _ -> (intern env f,[],[],[]), args in apply_impargs c env impargs args_scopes (merge_impargs l args) loc @@ -1660,15 +1666,15 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let st = Evar_kinds.Define (not (Program.get_proofs_transparency ())) in let fields = sort_fields ~complete:true loc fs - (fun _idx -> Loc.tag ?loc @@ CHole (Some (Evar_kinds.QuestionMark st), - Misctypes.IntroAnonymous, None)) + (fun _idx -> CAst.make ?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 (Loc.tag ?loc @@ CHole (None, Misctypes.IntroAnonymous, None)) in - let app = Loc.tag ?loc @@ CAppExpl ((None, constrname,None), List.rev_append pars args) in + let pars = List.make n (CAst.make ?loc @@ CHole (None, Misctypes.IntroAnonymous, None)) in + let app = CAst.make ?loc @@ CAppExpl ((None, constrname,None), List.rev_append pars args) in intern env app end | CCases (sty, rtnpo, tms, eqns) -> @@ -1910,6 +1916,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = in aux 1 l subscopes eargs rargs and apply_impargs c env imp subscopes l loc = + let l : (Constrexpr.constr_expr * Constrexpr.explicitation Loc.located option) list = l in 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 diff --git a/interp/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml index dd04e2030..52a6c450b 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 (loc, c) = match c with + let rec aux bdvars l c = match CAst.(c.v) 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 (loc, c) - | c -> Topconstr.fold_constr_expr_with_binders (fun a l -> Id.Set.add a l) aux bdvars l (loc, c) + | CNotation ("{ _ : _ | _ }", ({ CAst.v = 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 + | _ -> Topconstr.fold_constr_expr_with_binders (fun a l -> Id.Set.add a l) aux bdvars l c in aux bound l c let ids_of_names l = @@ -252,18 +252,22 @@ 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 - (Loc.tag @@ CRef (Ident (Loc.tag id'),None), Id.Set.add id' avoid) + (CAst.make @@ CRef (Ident (Loc.tag id'),None), Id.Set.add id' avoid) -let destClassApp (loc, cl) = - match cl with - | CApp ((None, (_loc, CRef (ref, inst))), l) -> Loc.tag ?loc (ref, List.map fst l, inst) +let destClassApp cl = + let open CAst in + let loc = cl.loc in + match cl.v with + | CApp ((None, { v = CRef (ref, inst) }), l) -> Loc.tag ?loc (ref, List.map fst l, inst) | CAppExpl ((None, ref, inst), l) -> Loc.tag ?loc (ref, l, inst) | CRef (ref, inst) -> Loc.tag ?loc:(loc_of_reference ref) (ref, [], inst) | _ -> raise Not_found -let destClassAppExpl (loc, cl) = - match cl with - | CApp ((None, (_loc, CRef (ref, inst))), l) -> Loc.tag ?loc (ref, l, inst) +let destClassAppExpl cl = + let open CAst in + let loc = cl.loc in + match cl.v with + | CApp ((None, { v = CRef (ref, inst) } ), l) -> Loc.tag ?loc (ref, l, inst) | CRef (ref, inst) -> Loc.tag ?loc:(loc_of_reference ref) (ref, [], inst) | _ -> raise Not_found @@ -296,7 +300,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 - Loc.tag ?loc @@ CAppExpl ((None, id, inst), args), avoid + CAst.make ?loc @@ CAppExpl ((None, id, inst), args), avoid in c, avoid let implicits_of_glob_constr ?(with_products=true) l = diff --git a/interp/notation.ml b/interp/notation.ml index 150be040f..03dffa6ee 100644 --- a/interp/notation.ml +++ b/interp/notation.ml @@ -471,7 +471,7 @@ let interp_prim_token = (** [rcp_of_glob] : from [glob_constr] to [raw_cases_pattern_expr] *) -let rec rcp_of_glob looked_for gt = Loc.map (function +let rec rcp_of_glob looked_for gt = CAst.map_from_loc (fun ?loc -> function | GVar id -> RCPatAtom (Some id) | GHole (_,_,_) -> RCPatAtom None | GRef (g,_) -> looked_for g; RCPatCstr (g,[],[]) diff --git a/interp/topconstr.ml b/interp/topconstr.ml index a74e64172..eb89b2ef2 100644 --- a/interp/topconstr.ml +++ b/interp/topconstr.ml @@ -43,7 +43,7 @@ let is_constructor id = (Nametab.locate_extended (qualid_of_ident id))) with Not_found -> false -let rec cases_pattern_fold_names f a pt = match snd pt with +let rec cases_pattern_fold_names f a pt = match CAst.(pt.v) with | CPatRecord l -> List.fold_left (fun acc (r, cp) -> cases_pattern_fold_names f acc cp) a l | CPatAlias (pat,id) -> f id a @@ -58,7 +58,7 @@ let rec cases_pattern_fold_names f a pt = match snd pt with | CPatDelimiters (_,pat) -> cases_pattern_fold_names f a pat | CPatAtom (Some (Ident (_,id))) when not (is_constructor id) -> f id a | CPatPrim _ | CPatAtom _ -> a - | CPatCast ((loc,_),_) -> + | CPatCast ({CAst.loc},_) -> CErrors.user_err ?loc ~hdr:"cases_pattern_fold_names" (Pp.strbrk "Casts are not supported here.") @@ -103,7 +103,7 @@ 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 = Loc.with_loc (fun ?loc -> function +let fold_constr_expr_with_binders g f n acc = CAst.with_val (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 @@ -115,7 +115,7 @@ let fold_constr_expr_with_binders g f n acc = Loc.with_loc (fun ?loc -> function (* 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 (Loc.tag @@ CHole (None,IntroAnonymous,None)) bl) acc bll + List.fold_left (fun acc bl -> fold_local_binders g f n acc (CAst.make @@ 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 _ -> @@ -146,7 +146,7 @@ let fold_constr_expr_with_binders g f n acc = Loc.with_loc (fun ?loc -> function let free_vars_of_constr_expr c = let rec aux bdvars l = function - | _loc, CRef (Ident (_,id),_) -> if Id.List.mem id bdvars then l else Id.Set.add id l + | { CAst.v = 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 @@ -210,7 +210,7 @@ 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 = Loc.map (function +let map_constr_expr_with_binders g f e = CAst.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) @@ -263,8 +263,8 @@ let map_constr_expr_with_binders g f e = Loc.map (function (* Used in constrintern *) let rec replace_vars_constr_expr l = function - | 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) + | { CAst.loc; v = CRef (Ident (loc_id,id),us) } as x -> + (try CAst.make ?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 23223173e..4f1e9d8e6 100644 --- a/intf/constrexpr.mli +++ b/intf/constrexpr.mli @@ -42,7 +42,7 @@ type raw_cases_pattern_expr_r = (** [CPatCstr (_, c, l1, l2)] represents ((@c l1) l2) *) | RCPatAtom of Id.t option | RCPatOr of raw_cases_pattern_expr list -and raw_cases_pattern_expr = raw_cases_pattern_expr_r Loc.located +and raw_cases_pattern_expr = raw_cases_pattern_expr_r CAst.ast type instance_expr = Misctypes.glob_level list @@ -61,7 +61,7 @@ type cases_pattern_expr_r = | CPatRecord of (reference * cases_pattern_expr) list | CPatDelimiters of string * cases_pattern_expr | CPatCast of cases_pattern_expr * constr_expr -and cases_pattern_expr = cases_pattern_expr_r Loc.located +and cases_pattern_expr = cases_pattern_expr_r CAst.ast and cases_pattern_notation_substitution = cases_pattern_expr list * (** for constr subterms *) @@ -98,7 +98,7 @@ and constr_expr_r = | CGeneralization of binding_kind * abstraction_kind option * constr_expr | CPrim of prim_token | CDelimiters of string * constr_expr -and constr_expr = constr_expr_r Loc.located +and constr_expr = constr_expr_r CAst.ast and case_expr = constr_expr (* expression that is being matched *) * Name.t Loc.located option (* as-clause *) diff --git a/lib/cAst.ml b/lib/cAst.ml new file mode 100644 index 000000000..5916c9ec1 --- /dev/null +++ b/lib/cAst.ml @@ -0,0 +1,24 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* 'a -> 'a ast + +val map : ('a -> 'b) -> 'a ast -> 'b ast +val map_with_loc : (?loc:Loc.t -> 'a -> 'b) -> 'a ast -> 'b ast +val map_from_loc : (?loc:Loc.t -> 'a -> 'b) -> 'a Loc.located -> 'b ast + +val with_val : ('a -> 'b) -> 'a ast -> 'b +val with_loc_val : (?loc:Loc.t -> 'a -> 'b) -> 'a ast -> 'b diff --git a/lib/clib.mllib b/lib/clib.mllib index c73ae9b90..9eb479fcc 100644 --- a/lib/clib.mllib +++ b/lib/clib.mllib @@ -18,6 +18,7 @@ IStream Flags Control Loc +CAst CList CString Deque diff --git a/parsing/egramcoq.ml b/parsing/egramcoq.ml index 49bf267db..712d10a96 100644 --- a/parsing/egramcoq.ml +++ b/parsing/egramcoq.ml @@ -313,11 +313,11 @@ 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) = Loc.tag ?loc @@ match na with +let constr_expr_of_name (loc,na) = CAst.make ?loc @@ match na with | Anonymous -> CHole (None,Misctypes.IntroAnonymous,None) | Name id -> CRef (Ident (Loc.tag ?loc id), None) -let cases_pattern_expr_of_name (loc,na) = Loc.tag ?loc @@ match na with +let cases_pattern_expr_of_name (loc,na) = CAst.make ?loc @@ match na with | Anonymous -> CPatAtom None | Name id -> CPatAtom (Some (Ident (Loc.tag ?loc id))) @@ -342,13 +342,13 @@ match e with | TTBinderListF _ -> { subst with binders = (List.flatten v, false) :: subst.binders } | TTBigint -> begin match forpat with - | ForConstr -> push_constr subst (Loc.tag @@ CPrim (Numeral v)) - | ForPattern -> push_constr subst (Loc.tag @@ CPatPrim (Numeral v)) + | ForConstr -> push_constr subst (CAst.make @@ CPrim (Numeral v)) + | ForPattern -> push_constr subst (CAst.make @@ CPatPrim (Numeral v)) end | TTReference -> begin match forpat with - | ForConstr -> push_constr subst (Loc.tag @@ CRef (v, None)) - | ForPattern -> push_constr subst (Loc.tag @@ CPatAtom (Some v)) + | ForConstr -> push_constr subst (CAst.make @@ CRef (v, None)) + | ForPattern -> push_constr subst (CAst.make @@ CPatAtom (Some v)) end | TTConstrList _ -> { subst with constrlists = v :: subst.constrlists } @@ -431,12 +431,12 @@ 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 - Loc.tag ~loc @@ CNotation (notation , env) + CAst.make ~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 let env = (env.constrs, env.constrlists) in - Loc.tag ~loc @@ CPatNotation (notation, env, []) + CAst.make ~loc @@ CPatNotation (notation, env, []) let extend_constr state forpat ng = let n = ng.notgram_level in diff --git a/parsing/g_constr.ml4 b/parsing/g_constr.ml4 index 45585d9ce..a44aa5400 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_opt (constr_loc c) (constr_loc ty) - in Loc.tag ?loc @@ CCast(c, CastConv ty) + in CAst.make ?loc @@ CCast(c, CastConv ty) let binder_of_name expl (loc,na) = CLocalAssum ([loc, na], Default expl, - Loc.tag ?loc @@ CHole (Some (Evar_kinds.BinderType na), IntroAnonymous, None)) + CAst.make ?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 -> Loc.tag @@ CHole (None, IntroAnonymous, None) in + | None -> CAst.make @@ 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 -> Loc.tag @@ CHole (None, IntroAnonymous, None) in + | None -> CAst.make @@ 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 - Loc.tag ~loc @@ CFix(id,fb) + CAst.make ~loc @@ CFix(id,fb) else let fb = List.map mk_cofixb dcls in - Loc.tag ~loc @@ CCoFix(id,fb) + CAst.make ~loc @@ CCoFix(id,fb) let mk_single_fix (loc,kw,dcl) = let (id,_,_,_,_) = dcl in mk_fix(loc,kw,id,[dcl]) @@ -120,7 +120,7 @@ let name_colon = | _ -> err ()) | _ -> err ()) -let aliasvar = function loc, CPatAlias (_, id) -> Some (loc,Name id) | _ -> None +let aliasvar = function { CAst.loc = loc; CAst.v = CPatAlias (_, id) } -> Some (loc,Name id) | _ -> None GEXTEND Gram GLOBAL: binder_constr lconstr constr operconstr universe_level sort global @@ -159,62 +159,62 @@ GEXTEND Gram ; constr: [ [ c = operconstr LEVEL "8" -> c - | "@"; f=global; i = instance -> Loc.tag ~loc:!@loc @@ CAppExpl((None,f,i),[]) ] ] + | "@"; f=global; i = instance -> CAst.make ~loc:!@loc @@ CAppExpl((None,f,i),[]) ] ] ; operconstr: [ "200" RIGHTA [ c = binder_constr -> c ] | "100" RIGHTA [ c1 = operconstr; "<:"; c2 = binder_constr -> - Loc.tag ~loc:(!@loc) @@ CCast(c1, CastVM c2) + CAst.make ~loc:(!@loc) @@ CCast(c1, CastVM c2) | c1 = operconstr; "<:"; c2 = SELF -> - Loc.tag ~loc:(!@loc) @@ CCast(c1, CastVM c2) + CAst.make ~loc:(!@loc) @@ CCast(c1, CastVM c2) | c1 = operconstr; "<<:"; c2 = binder_constr -> - Loc.tag ~loc:(!@loc) @@ CCast(c1, CastNative c2) + CAst.make ~loc:(!@loc) @@ CCast(c1, CastNative c2) | c1 = operconstr; "<<:"; c2 = SELF -> - Loc.tag ~loc:(!@loc) @@ CCast(c1, CastNative c2) + CAst.make ~loc:(!@loc) @@ CCast(c1, CastNative c2) | c1 = operconstr; ":";c2 = binder_constr -> - Loc.tag ~loc:(!@loc) @@ CCast(c1, CastConv c2) + CAst.make ~loc:(!@loc) @@ CCast(c1, CastConv c2) | c1 = operconstr; ":"; c2 = SELF -> - Loc.tag ~loc:(!@loc) @@ CCast(c1, CastConv c2) + CAst.make ~loc:(!@loc) @@ CCast(c1, CastConv c2) | c1 = operconstr; ":>" -> - Loc.tag ~loc:(!@loc) @@ CCast(c1, CastCoerce) ] + CAst.make ~loc:(!@loc) @@ CCast(c1, CastCoerce) ] | "99" RIGHTA [ ] | "90" RIGHTA [ ] | "10" LEFTA - [ 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) + [ f=operconstr; args=LIST1 appl_arg -> CAst.make ~loc:(!@loc) @@ CApp((None,f),args) + | "@"; f=global; i = instance; args=LIST0 NEXT -> CAst.make ~loc:!@loc @@ CAppExpl((None,f,i),args) | "@"; (locid,id) = pattern_identref; args=LIST1 identref -> - 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) ] + let args = List.map (fun x -> CAst.make @@ CRef (Ident x,None), None) args in + CAst.make ~loc:(!@loc) @@ CApp((None, CAst.make ?loc:locid @@ CPatVar id),args) ] | "9" [ ".."; c = operconstr LEVEL "0"; ".." -> - Loc.tag ~loc:!@loc @@ CAppExpl ((None, Ident (Loc.tag ~loc:!@loc ldots_var),None),[c]) ] + CAst.make ~loc:!@loc @@ CAppExpl ((None, Ident (Loc.tag ~loc:!@loc ldots_var),None),[c]) ] | "8" [ ] | "1" LEFTA [ c=operconstr; ".("; f=global; args=LIST0 appl_arg; ")" -> - Loc.tag ~loc:(!@loc) @@ CApp((Some (List.length args+1), Loc.tag @@ CRef (f,None)),args@[c,None]) + CAst.make ~loc:(!@loc) @@ CApp((Some (List.length args+1), CAst.make @@ CRef (f,None)),args@[c,None]) | c=operconstr; ".("; "@"; f=global; args=LIST0 (operconstr LEVEL "9"); ")" -> - 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) ] + CAst.make ~loc:(!@loc) @@ CAppExpl((Some (List.length args+1),f,None),args@[c]) + | c=operconstr; "%"; key=IDENT -> CAst.make ~loc:(!@loc) @@ CDelimiters (key,c) ] | "0" [ c=atomic_constr -> c | c=match_constr -> c | "("; c = operconstr LEVEL "200"; ")" -> - (match snd c with + (match c.CAst.v with CPrim (Numeral z) when Bigint.is_pos_or_zero z -> - Loc.tag ~loc:(!@loc) @@ CNotation("( _ )",([c],[],[])) + CAst.make ~loc:(!@loc) @@ CNotation("( _ )",([c],[],[])) | _ -> c) | "{|"; c = record_declaration; "|}" -> c | "`{"; c = operconstr LEVEL "200"; "}" -> - Loc.tag ~loc:(!@loc) @@ CGeneralization (Implicit, None, c) + CAst.make ~loc:(!@loc) @@ CGeneralization (Implicit, None, c) | "`("; c = operconstr LEVEL "200"; ")" -> - Loc.tag ~loc:(!@loc) @@ CGeneralization (Explicit, None, c) + CAst.make ~loc:(!@loc) @@ CGeneralization (Explicit, None, c) ] ] ; record_declaration: - [ [ fs = record_fields -> Loc.tag ~loc:(!@loc) @@ CRecord fs ] ] + [ [ fs = record_fields -> CAst.make ~loc:(!@loc) @@ CRecord fs ] ] ; record_fields: @@ -236,40 +236,40 @@ 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(c, CastConv t)) -> (Loc.tag ?loc:(constr_loc t) @@ Some t), c (* Tolerance, see G_vernac.def_body *) + | (_,None), { CAst.v = CCast(c, CastConv t) } -> (Loc.tag ?loc:(constr_loc t) @@ Some t), c (* Tolerance, see G_vernac.def_body *) | _, _ -> ty, c1 in - Loc.tag ~loc:!@loc @@ CLetIn(id,mkCLambdaN ?loc:(constr_loc c1) bl c1, + CAst.make ~loc:!@loc @@ CLetIn(id,mkCLambdaN ?loc:(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 snd fixp with + let (li,id) = match fixp.CAst.v with CFix(id,_) -> id | CCoFix(id,_) -> id | _ -> assert false in - Loc.tag ~loc:!@loc @@ CLetIn((li,Name id),fixp,None,c) + CAst.make ~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" -> - Loc.tag ~loc:!@loc @@ CLetTuple (lb,po,c1,c2) + CAst.make ~loc:!@loc @@ CLetTuple (lb,po,c1,c2) | "let"; "'"; p=pattern; ":="; c1 = operconstr LEVEL "200"; "in"; c2 = operconstr LEVEL "200" -> - Loc.tag ~loc:!@loc @@ + CAst.make ~loc:!@loc @@ CCases (LetPatternStyle, None, [c1, None, None], [Loc.tag ~loc:!@loc ([(Loc.tag ~loc:!@loc [p])], c2)]) | "let"; "'"; p=pattern; ":="; c1 = operconstr LEVEL "200"; rt = case_type; "in"; c2 = operconstr LEVEL "200" -> - Loc.tag ~loc:!@loc @@ + CAst.make ~loc:!@loc @@ CCases (LetPatternStyle, Some rt, [c1, aliasvar p, None], [Loc.tag ~loc:!@loc ([(Loc.tag ~loc:!@loc [p])], c2)]) | "let"; "'"; p=pattern; "in"; t = pattern LEVEL "200"; ":="; c1 = operconstr LEVEL "200"; rt = case_type; "in"; c2 = operconstr LEVEL "200" -> - Loc.tag ~loc:!@loc @@ + CAst.make ~loc:!@loc @@ CCases (LetPatternStyle, Some rt, [c1, aliasvar p, Some t], [Loc.tag ~loc:!@loc ([(Loc.tag ~loc:!@loc [p])], c2)]) | "if"; c=operconstr LEVEL "200"; po = return_type; "then"; b1=operconstr LEVEL "200"; "else"; b2=operconstr LEVEL "200" -> - Loc.tag ~loc:(!@loc) @@ CIf (c, po, b1, b2) + CAst.make ~loc:(!@loc) @@ CIf (c, po, b1, b2) | c=fix_constr -> c ] ] ; appl_arg: @@ -278,14 +278,14 @@ GEXTEND Gram | c=operconstr LEVEL "9" -> (c,None) ] ] ; atomic_constr: - [ [ 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) ] ] + [ [ g=global; i=instance -> CAst.make ~loc:!@loc @@ CRef (g,i) + | s=sort -> CAst.make ~loc:!@loc @@ CSort s + | n=INT -> CAst.make ~loc:!@loc @@ CPrim (Numeral (Bigint.of_string n)) + | s=string -> CAst.make ~loc:!@loc @@ CPrim (String s) + | "_" -> CAst.make ~loc:!@loc @@ CHole (None, IntroAnonymous, None) + | "?"; "["; id=ident; "]" -> CAst.make ~loc:!@loc @@ CHole (None, IntroIdentifier id, None) + | "?"; "["; id=pattern_ident; "]" -> CAst.make ~loc:!@loc @@ CHole (None, IntroFresh id, None) + | id=pattern_ident; inst = evar_instance -> CAst.make ~loc:!@loc @@ CEvar(id,inst) ] ] ; inst: [ [ id = ident; ":="; c = lconstr -> (id,c) ] ] @@ -326,7 +326,7 @@ GEXTEND Gram ; match_constr: [ [ "match"; ci=LIST1 case_item SEP ","; ty=OPT case_type; "with"; - br=branches; "end" -> Loc.tag ~loc:!@loc @@ CCases(RegularStyle,ty,ci,br) ] ] + br=branches; "end" -> CAst.make ~loc:!@loc @@ CCases(RegularStyle,ty,ci,br) ] ] ; case_item: [ [ c=operconstr LEVEL "100"; @@ -368,46 +368,46 @@ GEXTEND Gram pattern: [ "200" RIGHTA [ ] | "100" RIGHTA - [ p = pattern; "|"; pl = LIST1 pattern SEP "|" -> Loc.tag ~loc:!@loc @@ CPatOr (p::pl) ] + [ p = pattern; "|"; pl = LIST1 pattern SEP "|" -> CAst.make ~loc:!@loc @@ CPatOr (p::pl) ] | "99" RIGHTA [ ] | "11" LEFTA [ p = pattern; "as"; id = ident -> - Loc.tag ~loc:!@loc @@ CPatAlias (p, id) ] + CAst.make ~loc:!@loc @@ CPatAlias (p, id) ] | "10" RIGHTA [ p = pattern; lp = LIST1 NEXT -> - (match p with - | _, CPatAtom (Some r) -> Loc.tag ~loc:!@loc @@ CPatCstr (r, None, lp) - | loc, CPatCstr (r, None, l2) -> + (let open CAst in match p with + | { v = CPatAtom (Some r) } -> CAst.make ~loc:!@loc @@ CPatCstr (r, None, lp) + | { v = CPatCstr (r, None, l2); loc } -> CErrors.user_err ?loc ~hdr:"compound_pattern" (Pp.str "Nested applications not supported.") - | _, CPatCstr (r, l1, l2) -> Loc.tag ~loc:!@loc @@ CPatCstr (r, l1 , l2@lp) - | _, CPatNotation (n, s, l) -> Loc.tag ~loc:!@loc @@ CPatNotation (n , s, l@lp) + | { v = CPatCstr (r, l1, l2) } -> CAst.make ~loc:!@loc @@ CPatCstr (r, l1 , l2@lp) + | { v = CPatNotation (n, s, l) } -> CAst.make ~loc:!@loc @@ CPatNotation (n , s, l@lp) | _ -> CErrors.user_err ?loc:(cases_pattern_expr_loc p) ~hdr:"compound_pattern" (Pp.str "Such pattern cannot have arguments.")) |"@"; r = Prim.reference; lp = LIST0 NEXT -> - Loc.tag ~loc:!@loc @@ CPatCstr (r, Some lp, []) ] + CAst.make ~loc:!@loc @@ CPatCstr (r, Some lp, []) ] | "1" LEFTA - [ c = pattern; "%"; key=IDENT -> Loc.tag ~loc:!@loc @@ CPatDelimiters (key,c) ] + [ c = pattern; "%"; key=IDENT -> CAst.make ~loc:!@loc @@ CPatDelimiters (key,c) ] | "0" - [ r = Prim.reference -> Loc.tag ~loc:!@loc @@ CPatAtom (Some r) - | "{|"; pat = record_patterns; "|}" -> Loc.tag ~loc:!@loc @@ CPatRecord pat - | "_" -> Loc.tag ~loc:!@loc @@ CPatAtom None + [ r = Prim.reference -> CAst.make ~loc:!@loc @@ CPatAtom (Some r) + | "{|"; pat = record_patterns; "|}" -> CAst.make ~loc:!@loc @@ CPatRecord pat + | "_" -> CAst.make ~loc:!@loc @@ CPatAtom None | "("; p = pattern LEVEL "200"; ")" -> - (match p with - | _, CPatPrim (Numeral z) when Bigint.is_pos_or_zero z -> - Loc.tag ~loc:!@loc @@ CPatNotation("( _ )",([p],[]),[]) + (match p.CAst.v with + | CPatPrim (Numeral z) when Bigint.is_pos_or_zero z -> + CAst.make ~loc:!@loc @@ CPatNotation("( _ )",([p],[]),[]) | _ -> p) | "("; p = pattern LEVEL "200"; ":"; ty = lconstr; ")" -> let p = match p with - | _, CPatPrim (Numeral z) when Bigint.is_pos_or_zero z -> - Loc.tag ~loc:!@loc @@ CPatNotation("( _ )",([p],[]),[]) + | { CAst.v = CPatPrim (Numeral z) } when Bigint.is_pos_or_zero z -> + CAst.make ~loc:!@loc @@ CPatNotation("( _ )",([p],[]),[]) | _ -> p in - Loc.tag ~loc:!@loc @@ CPatCast (p, ty) - | n = INT -> Loc.tag ~loc:!@loc @@ CPatPrim (Numeral (Bigint.of_string n)) - | s = string -> Loc.tag ~loc:!@loc @@ CPatPrim (String s) ] ] + CAst.make ~loc:!@loc @@ CPatCast (p, ty) + | n = INT -> CAst.make ~loc:!@loc @@ CPatPrim (Numeral (Bigint.of_string n)) + | s = string -> CAst.make ~loc:!@loc @@ CPatPrim (String s) ] ] ; impl_ident_tail: [ [ "}" -> binder_of_name Implicit @@ -415,7 +415,7 @@ GEXTEND Gram (fun na -> CLocalAssum (na::nal,Default Implicit,c)) | nal=LIST1 name; "}" -> (fun na -> CLocalAssum (na::nal,Default Implicit, - Loc.tag ?loc:(Loc.merge_opt (fst na) (Some !@loc)) @@ CHole (Some (Evar_kinds.BinderType (snd na)), IntroAnonymous, None))) + CAst.make ?loc:(Loc.merge_opt (fst na) (Some !@loc)) @@ CHole (Some (Evar_kinds.BinderType (snd na)), IntroAnonymous, None))) | ":"; c=lconstr; "}" -> (fun na -> CLocalAssum ([na],Default Implicit,c)) ] ] @@ -449,7 +449,7 @@ GEXTEND Gram binders_of_names (id::idl) @ bl | id1 = name; ".."; id2 = name -> [CLocalAssum ([id1;(Loc.tag ~loc:!@loc (Name ldots_var));id2], - Default Explicit, Loc.tag ~loc:!@loc @@ CHole (None, IntroAnonymous, None))] + Default Explicit, CAst.make ~loc:!@loc @@ CHole (None, IntroAnonymous, None))] | bl = closed_binder; bl' = binders -> bl@bl' ] ] @@ -458,7 +458,7 @@ GEXTEND Gram [ [ l = LIST0 binder -> List.flatten l ] ] ; binder: - [ [ id = name -> [CLocalAssum ([id],Default Explicit, Loc.tag ~loc:!@loc @@ CHole (None, IntroAnonymous, None))] + [ [ id = name -> [CLocalAssum ([id],Default Explicit, CAst.make ~loc:!@loc @@ CHole (None, IntroAnonymous, None))] | bl = closed_binder -> bl ] ] ; closed_binder: @@ -467,27 +467,27 @@ GEXTEND Gram | "("; id=name; ":"; c=lconstr; ")" -> [CLocalAssum ([id],Default Explicit,c)] | "("; id=name; ":="; c=lconstr; ")" -> - (match c with - | (_, CCast(c, CastConv t)) -> [CLocalDef (id,c,Some t)] + (match c.CAst.v with + | 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, Loc.tag ~loc:!@loc @@ CHole (None, IntroAnonymous, None))] + [CLocalAssum ([id],Default Implicit, CAst.make ~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, Loc.tag ~loc:!@loc @@ CHole (None, IntroAnonymous, None))) (id::idl) + List.map (fun id -> CLocalAssum ([id],Default Implicit, CAst.make ~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 "," ; "}" -> List.map (fun (n, b, t) -> CLocalAssum ([n], Generalized (Implicit, Implicit, b), t)) tc | "'"; p = pattern LEVEL "0" -> let (p, ty) = - match p with - | _, CPatCast (p, ty) -> (p, Some ty) + match p.CAst.v with + | CPatCast (p, ty) -> (p, Some ty) | _ -> (p, None) in [CLocalPattern (Loc.tag ~loc:!@loc (p, ty))] diff --git a/parsing/g_proofs.ml4 b/parsing/g_proofs.ml4 index 488995201..ee0f06fe0 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 -> Loc.tag ~loc:!@loc @@ CCast(c,CastConv t) ] ] + | ":"; t = lconstr; ":="; c = lconstr -> CAst.make ~loc:!@loc @@ CCast(c,CastConv t) ] ] ; mode: [ [ l = LIST1 [ "+" -> ModeInput diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4 index c8101875d..27f879154 100644 --- a/parsing/g_vernac.ml4 +++ b/parsing/g_vernac.ml4 @@ -233,14 +233,14 @@ GEXTEND Gram DefineBody ([], red, c, None) else (match c with - | _, CCast(c, CastConv t) -> DefineBody (bl, red, c, Some t) + | { CAst.v = 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 = Loc.tag ~loc:!@loc @@ CCast (c, CastConv t) in + let c = CAst.make ~loc:!@loc @@ CCast (c, CastConv t) in (([],mkCLambdaN ~loc:!@loc bl c), None) else ((bl, c), Some t) in @@ -305,7 +305,7 @@ GEXTEND Gram ; type_cstr: [ [ ":"; c=lconstr -> c - | -> Loc.tag ~loc:!@loc @@ CHole (None, Misctypes.IntroAnonymous, None) ] ] + | -> CAst.make ~loc:!@loc @@ CHole (None, Misctypes.IntroAnonymous, None) ] ] ; (* Inductive schemes *) scheme: @@ -354,14 +354,14 @@ GEXTEND Gram t = lconstr; ":="; b = lconstr -> fun id -> (oc,DefExpr (id,mkCLambdaN ~loc:!@loc l b,Some (mkCProdN ~loc:!@loc l t))) | l = binders; ":="; b = lconstr -> fun id -> - match snd b with + match b.CAst.v with | CCast(b', (CastConv t|CastVM t|CastNative t)) -> (None,DefExpr(id,mkCLambdaN ~loc:!@loc l b',Some (mkCProdN ~loc:!@loc l t))) | _ -> (None,DefExpr(id,mkCLambdaN ~loc:!@loc l b,None)) ] ] ; record_binder: - [ [ id = name -> (None,AssumExpr(id, Loc.tag ~loc:!@loc @@ CHole (None, Misctypes.IntroAnonymous, None))) + [ [ id = name -> (None,AssumExpr(id, CAst.make ~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:!@loc l c)) | -> - fun l id -> (false,(id,mkCProdN ~loc:!@loc l (Loc.tag ~loc:!@loc @@ CHole (None, Misctypes.IntroAnonymous, None)))) ] + fun l id -> (false,(id,mkCProdN ~loc:!@loc l (CAst.make ~loc:!@loc @@ CHole (None, Misctypes.IntroAnonymous, None)))) ] -> t l ]] ; diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml index 48ab3dd57..938fe5237 100644 --- a/plugins/funind/glob_term_to_relation.ml +++ b/plugins/funind/glob_term_to_relation.ml @@ -1246,15 +1246,16 @@ let compute_params_name relnames (args : (Name.t * Glob_term.glob_constr * glob_ in List.rev !l -let rec rebuild_return_type (loc, rt) = - match rt with +let rec rebuild_return_type rt = + let loc = rt.CAst.loc in + match rt.CAst.v with | Constrexpr.CProdN(n,t') -> - Loc.tag ?loc @@ Constrexpr.CProdN(n,rebuild_return_type t') + CAst.make ?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.tag Anonymous], - Constrexpr.Default Decl_kinds.Explicit,Loc.tag ?loc rt], - Loc.tag @@ Constrexpr.CSort(GType [])) + CAst.make ?loc @@ Constrexpr.CLetIn(na,v,t,rebuild_return_type t') + | _ -> CAst.make ?loc @@ Constrexpr.CProdN([[Loc.tag Anonymous], + Constrexpr.Default Decl_kinds.Explicit, rt], + CAst.make @@ Constrexpr.CSort(GType [])) let do_build_inductive evd (funconstants: Term.pconstant list) (funsargs: (Name.t * glob_constr * glob_constr option) list list) @@ -1305,11 +1306,11 @@ let do_build_inductive (fun (n,t,typ) acc -> match typ with | Some typ -> - Loc.tag @@ Constrexpr.CLetIn((Loc.tag n),with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t, + CAst.make @@ Constrexpr.CLetIn((Loc.tag 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 -> - Loc.tag @@ Constrexpr.CProdN + CAst.make @@ Constrexpr.CProdN ([[(Loc.tag n)],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t], acc ) @@ -1372,11 +1373,11 @@ let do_build_inductive (fun (n,t,typ) acc -> match typ with | Some typ -> - Loc.tag @@ Constrexpr.CLetIn((Loc.tag n),with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t, + CAst.make @@ Constrexpr.CLetIn((Loc.tag 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 -> - Loc.tag @@ Constrexpr.CProdN + CAst.make @@ Constrexpr.CProdN ([[(Loc.tag 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 3dc16626c..f4e9aa372 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -469,7 +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 = - Loc.tag @@ Constrexpr.CAppExpl( + CAst.make @@ Constrexpr.CAppExpl( (None,(Ident (Loc.tag fname)),None) , (List.map (function @@ -480,7 +480,7 @@ let register_wf ?(is_mes=false) fname rec_impls wf_rel_expr wf_arg using_lemmas ) ) in - Loc.tag @@ Constrexpr.CApp ((None,Constrexpr_ops.mkRefC (Qualid (Loc.tag (qualid_of_string "Logic.eq")))), + CAst.make @@ Constrexpr.CApp ((None,Constrexpr_ops.mkRefC (Qualid (Loc.tag (qualid_of_string "Logic.eq")))), [(f_app_args,None);(body,None)]) in let eq = Constrexpr_ops.mkCProdN args unbounded_eq in @@ -588,12 +588,12 @@ 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',(_loc, Constrexpr.CLetIn(_,nat,ty,typ')) -> + | (Constrexpr.CLocalDef(na,_,_))::bl',{ CAst.v = 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, snd typ with + match nal, typ.CAst.v 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') -> @@ -606,7 +606,7 @@ 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 Loc.tag @@ if List.is_empty new_nal' + else CAst.make @@ if List.is_empty new_nal' then CProdN(rest,typ') else CProdN(((new_nal',bk',nal't)::rest),typ')) else @@ -614,7 +614,7 @@ let rec rebuild_bl (aux,assoc) bl typ = 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') (Loc.tag @@ CProdN(rest,typ')) + bk bl' non_captured_nal (lnal - lnal') (CAst.make @@ CProdN(rest,typ')) | _ -> assert false let rebuild_bl (aux,assoc) bl typ = rebuild_bl (aux,assoc) bl typ @@ -725,7 +725,7 @@ let do_generate_principle pconstants on_error register_built interactive_proof in () -let rec add_args id new_args = Loc.map (function +let rec add_args id new_args = CAst.map (function | CRef (r,_) as b -> begin match r with | Libnames.Ident(loc,fname) when Id.equal fname id -> @@ -794,7 +794,7 @@ 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 snd t with + match t.CAst.v 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 @@ -813,7 +813,7 @@ let rec chop_n_arrow n t = then aux (n - nal_l) nal_ta' else - let new_t' = Loc.tag @@ + let new_t' = CAst.make @@ Constrexpr.CProdN( ((snd (List.chop n nal)),k,t'')::nal_ta',t') in @@ -829,7 +829,7 @@ 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 snd b with + match b.CAst.v with | Constrexpr.CLambdaN ((nal_ta), b') -> begin let n = @@ -869,7 +869,7 @@ let make_graph (f_ref:global_reference) = in let (nal_tas,b,t) = get_args extern_body extern_type in let expr_list = - match snd b with + match b.CAst.v with | Constrexpr.CFix(l_id,fixexprl) -> let l = List.map @@ -882,7 +882,7 @@ let make_graph (f_ref:global_reference) = | Constrexpr.CLocalDef (na,_,_)-> [] | Constrexpr.CLocalAssum (nal,_,_) -> List.map - (fun (loc,n) -> Loc.tag ?loc @@ + (fun (loc,n) -> CAst.make ?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 f2654b5de..5b51a213a 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 = Loc.tag @@ CRef (Libnames.Ident (Loc.tag id), None) in + let ans = CAst.make @@ CRef (Libnames.Ident (Loc.tag 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 - Loc.tag @@ CProdN ([[(Loc.tag nm)],Constrexpr_ops.default_binder_kind,typ] , acc) , newenv) + CAst.make @@ CProdN ([[(Loc.tag 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 e20beae96..5fc22cb4a 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 _loc, CRef (r,None) -> Reference r | c -> ConstrMayEval (ConstrTerm c)) + | c = Constr.constr -> (match c with { CAst.v = 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 - | _loc, CCast (t, (CastConv ty | CastVM ty | CastNative ty)) -> Term t, Some (Term ty) + | { CAst.v = CCast (t, (CastConv ty | CastVM ty | CastNative ty)) } -> Term t, Some (Term ty) | _ -> mpv, None) | _ -> mpv, None - in Def (na, t, Option.default (Term (Loc.tag @@ CHole (None, IntroAnonymous, None))) ty) + in Def (na, t, Option.default (Term (CAst.make @@ 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 - Loc.tag ~loc:!@loc @@ CHole (None, IntroAnonymous, Some arg) ] ] + CAst.make ~loc:!@loc @@ CHole (None, IntroAnonymous, Some arg) ] ] ; END diff --git a/plugins/ltac/g_tactic.ml4 b/plugins/ltac/g_tactic.ml4 index 8aaad0566..60deb443a 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, Loc.tag ~loc @@ CProdN(bl,ty)) + (id,n, CAst.make ~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,Loc.tag ~loc @@ CProdN(bl,ty)) + (id,CAst.make ~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,(Loc.tag @@ CPrim (Numeral (Bigint.of_int n)), + (clear,(CAst.make @@ 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,(Loc.tag @@ CRef (Ident id,None),NoBindings))) + TacCase (with_evar,(clear,(CAst.make @@ 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 -> - Loc.tag ?loc @@ CLambdaN ([idl,bk,t],mkCLambdaN_simple_loc ?loc:(Loc.merge_opt loc1 loc) bll c) + CAst.make ?loc @@ CLambdaN ([idl,bk,t],mkCLambdaN_simple_loc ?loc:(Loc.merge_opt 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, Loc.tag ~loc:!@loc @@ CHole (Some (Evar_kinds.BinderType (snd na)), IntroAnonymous, None)) + [ [ na=name -> ([na],Default Explicit, CAst.make ~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 bdafbdc78..58473d7dd 100644 --- a/plugins/ltac/pptactic.ml +++ b/plugins/ltac/pptactic.ml @@ -340,7 +340,7 @@ type 'a extra_genarg_printer = let strip_prod_binders_expr n ty = let rec strip_ty acc n ty = - match snd ty with + match ty.CAst.v with Constrexpr.CProdN(bll,a) -> let nb = List.fold_left (fun i (nal,_,_) -> i + List.length nal) 0 bll in diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml index 19c2eaf0a..2ef435b6b 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 = Loc.tag @@ CAppExpl ((None,(Libnames.Ident (Loc.tag @@ Id.of_string s)),None),l) +let mkappc s l = CAst.make @@ CAppExpl ((None,(Libnames.Ident (Loc.tag @@ Id.of_string s)),None),l) let declare_an_instance n s args = (((Loc.tag @@ Name n),None), Explicit, - Loc.tag @@ CAppExpl ((None, Qualid (Loc.tag @@ qualid_of_string s),None), + CAst.make @@ CAppExpl ((None, Qualid (Loc.tag @@ 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, Loc.tag @@ CRecord (fields))) + binders instance (Some (true, CAst.make @@ 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.tag @@ Id.of_string "Equivalence_Symmetric"), lemma2); (Ident (Loc.tag @@ Id.of_string "Equivalence_Transitive"), lemma3)]) -let cHole = Loc.tag @@ CHole (None, Misctypes.IntroAnonymous, None) +let cHole = CAst.make @@ 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 @@ -2012,13 +2012,13 @@ let add_morphism glob binders m s n = let instance_id = add_suffix n "_Proper" in let instance = (((Loc.tag @@ Name instance_id),None), Explicit, - Loc.tag @@ CAppExpl ( + CAst.make @@ CAppExpl ( (None, Qualid (Loc.tag @@ 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, Loc.tag @@ CRecord [])) + (Some (true, CAst.make @@ 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 44ea3ff1d..566dd8ed7 100644 --- a/plugins/ltac/tacintern.ml +++ b/plugins/ltac/tacintern.ml @@ -108,13 +108,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 -> - (Loc.tag @@ GVar id), Some (Loc.tag @@ CRef (r,None)) + (Loc.tag @@ GVar id), Some (CAst.make @@ CRef (r,None)) | Ident (_,id) as r when find_var id ist -> - (Loc.tag @@ GVar id), if strict then None else Some (Loc.tag @@ CRef (r,None)) + (Loc.tag @@ GVar id), if strict then None else Some (CAst.make @@ CRef (r,None)) | r -> let loc,_ as lqid = qualid_of_reference r in Loc.tag @@ GRef (locate_global_with_alias lqid,None), - if strict then None else Some (Loc.tag @@ CRef (r,None)) + if strict then None else Some (CAst.make @@ CRef (r,None)) let intern_move_location ist = function | MoveAfter id -> MoveAfter (intern_hyp ist id) @@ -271,7 +271,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 (Loc.tag @@ CRef (Ident (Loc.tag id), None)) with + match intern_constr ist (CAst.make @@ CRef (Ident (Loc.tag id), None)) with | (loc, GVar id), _ -> clear,ElimOnIdent (loc,id) | c -> clear,ElimOnConstr (c,NoBindings) else @@ -361,7 +361,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 { CAst.v = 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 4d5b84455..449027b52 100644 --- a/plugins/ltac/tacinterp.ml +++ b/plugins/ltac/tacinterp.ml @@ -1072,7 +1072,7 @@ let interp_destruction_arg ist gl arg = if Tactics.is_quantified_hypothesis id gl then keep,ElimOnIdent (loc,id) else - let c = (Loc.tag ?loc @@ GVar id,Some (Loc.tag @@ CRef (Ident (loc,id),None))) in + let c = (Loc.tag ?loc @@ GVar id,Some (CAst.make @@ 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 0bc3f502c..c11c9f83b 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -148,15 +148,15 @@ let add_genarg tag pr = (** Constructors for cast type *) let dC t = CastConv t (** Constructors for constr_expr *) -let isCVar = function _loc, CRef (Ident _, _) -> true | _ -> false -let destCVar = function _loc, CRef (Ident (_, id), _) -> id | _ -> +let isCVar = function { CAst.v = CRef (Ident _, _) } -> true | _ -> false +let destCVar = function { CAst.v = CRef (Ident (_, id), _) } -> id | _ -> CErrors.anomaly (str"not a CRef") -let mkCHole ~loc = Loc.tag ?loc @@ CHole (None, IntroAnonymous, None) -let mkCLambda ?loc name ty t = Loc.tag ?loc @@ +let mkCHole ~loc = CAst.make ?loc @@ CHole (None, IntroAnonymous, None) +let mkCLambda ?loc name ty t = CAst.make ?loc @@ CLambdaN ([[Loc.tag ?loc name], Default Explicit, ty], t) -let mkCLetIn ?loc name bo t = Loc.tag ?loc @@ +let mkCLetIn ?loc name bo t = CAst.make ?loc @@ CLetIn ((Loc.tag ?loc name), bo, None, t) -let mkCCast ?loc t ty = Loc.tag ?loc @@ CCast (t, dC ty) +let mkCCast ?loc t ty = CAst.make ?loc @@ CCast (t, dC ty) (** Constructors for rawconstr *) let mkRHole = Loc.tag @@ GHole (InternalHole, IntroAnonymous, None) let mkRApp f args = if args = [] then f else Loc.tag @@ GApp (f, args) @@ -951,7 +951,7 @@ 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 snd (Loc.to_pair t) with + match t.CAst.v with | CNotation("( _ in _ )", ([t1; t2], [], [])) -> (try match glob t1, glob t2 with | (r1, None), (r2, None) -> encode k "In" [r1;r2] @@ -1019,9 +1019,9 @@ let pr_rpattern = pr_pattern type pattern = Evd.evar_map * (constr, constr) ssrpattern -let id_of_cpattern = function - | _,(_,Some (_loc, CRef (Ident (_, x), _))) -> Some x - | _,(_,Some (_loc, CAppExpl ((_, Ident (_, x), _), []))) -> Some x +let id_of_cpattern = let open CAst in function + | _,(_,Some { v = CRef (Ident (_, x), _) } ) -> Some x + | _,(_,Some { v = CAppExpl ((_, Ident (_, x), _), []) } ) -> Some x | _,((_, GRef (VarRef x, _)) ,None) -> Some x | _ -> None let id_of_Cterm t = match id_of_cpattern t with @@ -1377,7 +1377,7 @@ let pf_fill_occ_term gl occ t = let cpattern_of_id id = ' ', (Loc.tag @@ GRef (VarRef id, None), None) let is_wildcard : cpattern -> bool = function - | _,(_,Some (_, CHole _)| (_, GHole _),None) -> true + | _,(_,Some { CAst.v = CHole _ } | (_, GHole _),None) -> true | _ -> false (* "ssrpattern" *) diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml index 900bf2daf..1f3593a71 100644 --- a/printing/ppconstr.ml +++ b/printing/ppconstr.ml @@ -217,7 +217,7 @@ let tag_var = tag Tag.variable | t -> cut () ++ str ":" ++ pr t let pr_opt_type_spc pr = function - | _loc, CHole (_,Misctypes.IntroAnonymous,_) -> mt () + | { CAst.v = CHole (_,Misctypes.IntroAnonymous,_) } -> mt () | t -> str " :" ++ pr_sep_com (fun()->brk(1,2)) (pr ltop) t let pr_lident (loc,id) = @@ -251,8 +251,8 @@ let tag_var = tag Tag.variable let lpator = 100 let lpatrec = 0 - let rec pr_patt sep inh (loc, p) = - let (strm,prec) = match p with + let rec pr_patt sep inh p = + let (strm,prec) = match CAst.(p.v) with | CPatRecord l -> let pp (c, p) = pr_reference c ++ spc() ++ str ":=" ++ pr_patt spc (lpatrec, Any) p @@ -300,7 +300,7 @@ let tag_var = tag Tag.variable | CPatCast _ -> assert false in - let loc = cases_pattern_expr_loc (loc, p) in + let loc = p.CAst.loc in pr_with_comments ?loc (sep() ++ if prec_less prec inh then strm else surround strm) @@ -353,7 +353,7 @@ let tag_var = tag Tag.variable end | Default b -> match t with - | _loc, CHole (_,Misctypes.IntroAnonymous,_) -> + | { CAst.v = CHole (_,Misctypes.IntroAnonymous,_) } -> let s = prlist_with_sep spc pr_lname nal in hov 1 (surround_implicit b s) | _ -> @@ -391,42 +391,42 @@ let tag_var = tag Tag.variable if is_open then pr_delimited_binders pr_com_at sep pr_c else pr_undelimited_binders sep pr_c - let rec extract_prod_binders = function + let rec extract_prod_binders = let open CAst in function (* | 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*) - | _loc, CProdN ([],c) -> + | { v = CProdN ([],c) } -> extract_prod_binders c - | loc, CProdN ([[_,Name id],bk,t], - (_loc', CCases (LetPatternStyle,None, [(_loc'', CRef (Ident (_,id'),None)),None,None],[(_,([_,[p]],b))]))) + | { loc; v = CProdN ([[_,Name id],bk,t], + { v = CCases (LetPatternStyle,None, [{ v = 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 - | loc, CProdN ((nal,bk,t)::bl,c) -> - let bl,c = extract_prod_binders (loc, CProdN(bl,c)) in + | { loc; v = CProdN ((nal,bk,t)::bl,c) } -> + let bl,c = extract_prod_binders (CAst.make ?loc @@ CProdN(bl,c)) in CLocalAssum (nal,bk,t) :: bl, c | c -> [], c - let rec extract_lam_binders (loc, ce) = match ce with + let rec extract_lam_binders ce = let open CAst in match ce.v 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 ([],c) -> extract_lam_binders c | CLambdaN ([[_,Name id],bk,t], - (_loc, CCases (LetPatternStyle,None, [(_loc', CRef (Ident (_,id'),None)),None,None],[(_,([_,[p]],b))]))) + { v = CCases (LetPatternStyle,None, [{ v = 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 + CLocalPattern (ce.loc,(p,None)) :: bl, c | CLambdaN ((nal,bk,t)::bl,c) -> - let bl,c = extract_lam_binders (loc, CLambdaN(bl,c)) in + let bl,c = extract_lam_binders (CAst.make ?loc:ce.loc @@ CLambdaN(bl,c)) in CLocalAssum (nal,bk,t) :: bl, c - | c -> [], Loc.tag ?loc c + | _ -> [], ce - let split_lambda = Loc.with_loc (fun ?loc -> function + let split_lambda = CAst.with_loc_val (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)) + | CLambdaN (([na],bk,t)::bl,c) -> (na,t, CAst.make ?loc @@ CLambdaN(bl,c)) + | CLambdaN ((na::nal,bk,t)::bl,c) -> (na,t, CAst.make ?loc @@ CLambdaN((nal,bk,t)::bl,c)) | _ -> anomaly (Pp.str "ill-formed fixpoint body") ) @@ -437,11 +437,11 @@ let tag_var = tag Tag.variable | (_,Name id), (_,Anonymous) -> (na,t,c) | _ -> (na',t,c) - let split_product na' = Loc.with_loc (fun ?loc -> function + let split_product na' = CAst.with_loc_val (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],bk,t)::bl,c) -> rename na na' t (CAst.make ?loc @@ CProdN(bl,c)) | CProdN ((na::nal,bk,t)::bl,c) -> - rename na na' t (Loc.tag ?loc @@ CProdN((nal,bk,t)::bl,c)) + rename na na' t (CAst.make ?loc @@ CProdN((nal,bk,t)::bl,c)) | _ -> anomaly (Pp.str "ill-formed fixpoint body") ) @@ -509,7 +509,7 @@ let tag_var = tag Tag.variable let pr_case_type pr po = match po with - | None | Some (_, CHole (_,Misctypes.IntroAnonymous,_)) -> mt() + | None | Some { CAst.v = CHole (_,Misctypes.IntroAnonymous,_) } -> mt() | Some p -> spc() ++ hov 2 (keyword "return" ++ pr_sep_com spc (pr lsimpleconstr) p) @@ -546,7 +546,7 @@ let tag_var = tag Tag.variable let pr_fun_sep = spc () ++ str "=>" let pr_dangling_with_for sep pr inherited a = - match snd a with + match a.CAst.v with | (CFix (_,[_])|CCoFix(_,[_])) -> pr sep (latom,E) a | _ -> @@ -554,7 +554,7 @@ let tag_var = tag Tag.variable let pr pr sep inherited a = let return (cmds, prec) = (tag_constr_expr a cmds, prec) in - let (strm, prec) = match snd @@ Loc.to_pair a with + let (strm, prec) = match CAst.(a.v) with | CRef (r, us) -> return (pr_cref r us, latom) | CFix (id,fix) -> @@ -589,8 +589,8 @@ let tag_var = tag Tag.variable pr_fun_sep ++ pr spc ltop a), llambda ) - | CLetIn ((_,Name x), ((_loc, CFix((_,x'),[_])) - | (_loc, CCoFix((_,x'),[_])) as fx), t, b) + | CLetIn ((_,Name x), ({ CAst.v = CFix((_,x'),[_])} + | { CAst.v = CCoFix((_,x'),[_]) } as fx), t, b) when Id.equal x x' -> return ( hv 0 ( @@ -619,7 +619,7 @@ let tag_var = tag Tag.variable else return (p, lproj) | CAppExpl ((None,Ident (_,var),us),[t]) - | CApp ((_,(_, CRef(Ident(_,var),us))),[t,None]) + | CApp ((_, {CAst.v = 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 ".."), @@ -755,7 +755,7 @@ let tag_var = tag Tag.variable let pr prec c = pr prec (transf (Global.env()) c) let pr_simpleconstr = function - | _lock, CAppExpl ((None,f,us),[]) -> str "@" ++ pr_cref f us + | { CAst.v = 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 5b6c5580a..8928d83b0 100644 --- a/printing/ppvernac.ml +++ b/printing/ppvernac.ml @@ -252,7 +252,7 @@ open Decl_kinds prlist_strict (pr_module_vardecls pr_c) l let pr_type_option pr_c = function - | _loc, CHole (k, Misctypes.IntroAnonymous, _) -> mt() + | { CAst.v = CHole (k, Misctypes.IntroAnonymous, _) } -> mt() | _ as c -> brk(0,2) ++ str" :" ++ pr_c c let pr_decl_notation prc ((loc,ntn),c,scopt) = @@ -883,7 +883,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, (_loc, CRecord l)) -> spc () ++ str":=" ++ spc () ++ str"{" ++ pr_record_body l ++ str "}" + | Some (true, { CAst.v = 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 fb300dbc1..f9a3b01b6 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 = Loc.tag @@ CHole (None, Misctypes.IntroAnonymous, None) in + let t = CAst.make @@ 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 Loc.tag @@ CGeneralization (Implicit, Some AbsPi, tclass) + if generalize then CAst.make @@ 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, (_loc, CRecord fs)) -> + | Some (true, { CAst.v = 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 -> - ((Loc.tag @@ CHole (None(* Some Evar_kinds.GoalEvar *), Misctypes.IntroAnonymous, None)) :: props), rest + ((CAst.make @@ 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 82d7b19d7..12df344c2 100644 --- a/vernac/command.ml +++ b/vernac/command.ml @@ -53,7 +53,7 @@ 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 = Loc.map_with_loc (fun ?loc -> function +let rec complete_conclusion a cs = CAst.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, _, _) -> @@ -62,7 +62,7 @@ let rec complete_conclusion a cs = Loc.map_with_loc (fun ?loc -> function user_err ?loc (strbrk"Cannot infer the non constant arguments of the conclusion of " ++ pr_id cs ++ str "."); - let args = List.map (fun id -> Loc.tag ?loc @@ CRef(Ident(loc,id),None)) params in + let args = List.map (fun id -> CAst.make ?loc @@ CRef(Ident(loc,id),None)) params in CAppExpl ((None,Ident(loc,name),None),List.rev args) | c -> c ) @@ -683,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) (Loc.tag @@ CSort (GType [])) ar; + ind_arity = Option.cata (fun x -> x) (CAst.make @@ CSort (GType [])) ar; ind_lc = List.map (fun (_,((_,id),t)) -> (id,t)) lc }) indl @@ -1354,7 +1354,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 (Loc.tag @@ CRef (lt_ref,None)) r) m ntn + (Option.default (CAst.make @@ 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 feacfe915..83896992c 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 = Loc.tag @@ CRef (Ident (Loc.tag @@ Id.of_string x),None) +let inject_var x = CAst.make @@ CRef (Ident (Loc.tag @@ 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 - | [], (_loc, CRef (ref,_)) -> intern_reference ref + | [], { CAst.v = 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 8bd583b81..43a24e167 100644 --- a/vernac/record.ml +++ b/vernac/record.ml @@ -94,7 +94,7 @@ 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 -> Loc.tag ?loc:(fst n) @@ CHole (None, Misctypes.IntroAnonymous, None)) + | None -> CAst.make ?loc:(fst n) @@ CHole (None, Misctypes.IntroAnonymous, None)) let binders_of_decls = List.map binder_of_decl @@ -121,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 snd t with - | CSort (Misctypes.GType []) -> true | _ -> false in + match t with + | { CAst.v = 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 -- cgit v1.2.3 From ee2197096fe75a63b4d92cb3a1bb05122c5c625b Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Sun, 9 Apr 2017 03:35:20 +0200 Subject: [location] [ast] Port module AST to CAst --- interp/constrextern.ml | 83 +++++++------- interp/constrintern.ml | 167 +++++++++++++-------------- interp/implicit_quantifiers.ml | 4 +- interp/modintern.ml | 8 +- interp/notation.ml | 18 +-- interp/notation_ops.ml | 195 ++++++++++++++++---------------- intf/constrexpr.mli | 2 +- intf/glob_term.mli | 6 +- parsing/g_vernac.ml4 | 10 +- plugins/funind/glob_term_to_relation.ml | 84 +++++++------- plugins/funind/glob_termops.ml | 146 ++++++++++++------------ plugins/funind/indfun.ml | 2 +- plugins/funind/indfun_common.ml | 4 +- plugins/funind/merge.ml | 42 +++---- plugins/funind/recdef.ml | 10 +- plugins/ltac/extratactics.ml4 | 8 +- plugins/ltac/g_rewrite.ml4 | 2 +- plugins/ltac/pptactic.ml | 2 +- plugins/ltac/tacintern.ml | 10 +- plugins/ltac/tacinterp.ml | 4 +- plugins/setoid_ring/newring.ml | 2 +- plugins/ssrmatching/ssrmatching.ml4 | 33 +++--- plugins/syntax/ascii_syntax.ml | 12 +- plugins/syntax/nat_syntax.ml | 12 +- plugins/syntax/numbers_syntax.ml | 64 +++++------ plugins/syntax/r_syntax.ml | 32 +++--- plugins/syntax/string_syntax.ml | 12 +- plugins/syntax/z_syntax.ml | 59 +++++----- pretyping/cases.ml | 79 ++++++------- pretyping/coercion.ml | 4 +- pretyping/detyping.ml | 110 +++++++++--------- pretyping/detyping.mli | 8 -- pretyping/glob_ops.ml | 52 ++++----- pretyping/patternops.ml | 20 ++-- pretyping/pretyping.ml | 7 +- printing/ppvernac.ml | 10 +- tactics/hipattern.ml | 14 +-- vernac/command.ml | 2 +- 38 files changed, 673 insertions(+), 666 deletions(-) (limited to 'intf') diff --git a/interp/constrextern.ml b/interp/constrextern.ml index e8a5b5265..692a0872b 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -17,6 +17,7 @@ open Termops open Libnames open Globnames open Impargs +open CAst open Constrexpr open Constrexpr_ops open Notation_ops @@ -182,7 +183,7 @@ let add_patt_for_params ind l = let add_cpatt_for_params ind l = if !Flags.in_debugger then l else - Util.List.addn (Inductiveops.inductive_nparamdecls ind) (Loc.tag @@ PatVar Anonymous) l + Util.List.addn (Inductiveops.inductive_nparamdecls ind) (CAst.make @@ PatVar Anonymous) l let drop_implicits_in_patt cst nb_expl args = let impl_st = (implicits_of_global cst) in @@ -291,7 +292,7 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = (* pboutill: There are letins in pat which is incompatible with notations and not explicit application. *) match pat with - | loc, PatCstr(cstrsp,args,na) + | { loc; v = PatCstr(cstrsp,args,na) } when !Flags.in_debugger||Inductiveops.constructor_has_local_defs cstrsp -> let c = extern_reference ?loc Id.Set.empty (ConstructRef cstrsp) in let args = List.map (extern_cases_pattern_in_scope scopes vars) args in @@ -311,10 +312,10 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = extern_notation_pattern scopes vars pat (uninterp_cases_pattern_notations pat) with No_match -> - match pat with - | loc, PatVar (Name id) -> CAst.make ?loc @@ CPatAtom (Some (Ident (loc,id))) - | loc, PatVar (Anonymous) -> CAst.make ?loc @@ CPatAtom None - | loc, PatCstr(cstrsp,args,na) -> + CAst.map_with_loc (fun ?loc -> function + | PatVar (Name id) -> CPatAtom (Some (Ident (loc,id))) + | PatVar (Anonymous) -> CPatAtom None + | PatCstr(cstrsp,args,na) -> let args = List.map (extern_cases_pattern_in_scope scopes vars) args in let p = try @@ -337,20 +338,21 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = | head :: tail -> ip q tail ((extern_reference ?loc Id.Set.empty (ConstRef c), head) :: acc) in - CAst.make ?loc @@ CPatRecord(List.rev (ip projs args [])) + CPatRecord(List.rev (ip projs args [])) with Not_found | No_match | Exit -> let c = extern_reference ?loc Id.Set.empty (ConstructRef cstrsp) in if !Topconstr.asymmetric_patterns then if pattern_printable_in_both_syntax cstrsp - then CAst.make ?loc @@ CPatCstr (c, None, args) - else CAst.make ?loc @@ CPatCstr (c, Some (add_patt_for_params (fst cstrsp) args), []) + then CPatCstr (c, None, args) + else CPatCstr (c, Some (add_patt_for_params (fst cstrsp) args), []) else let full_args = add_patt_for_params (fst cstrsp) args in match drop_implicits_in_patt (ConstructRef cstrsp) 0 full_args with - | Some true_args -> CAst.make ?loc @@ CPatCstr (c, None, true_args) - | None -> CAst.make ?loc @@ CPatCstr (c, Some full_args, []) - in insert_pat_alias ?loc p na + | Some true_args -> CPatCstr (c, None, true_args) + | None -> CPatCstr (c, Some full_args, []) + in (insert_pat_alias ?loc (CAst.make ?loc p) na).v + ) pat and apply_notation_to_pattern ?loc gr ((subst,substlist),(nb_to_drop,more_args)) (tmp_scope, scopes as allscopes) vars = function @@ -396,20 +398,21 @@ and apply_notation_to_pattern ?loc gr ((subst,substlist),(nb_to_drop,more_args)) in assert (List.is_empty substlist); mkPat ?loc qid (List.rev_append l1 l2') -and extern_notation_pattern (tmp_scope,scopes as allscopes) vars (loc, t) = function +and extern_notation_pattern (tmp_scope,scopes as allscopes) vars t = function | [] -> raise No_match | (keyrule,pat,n as _rule)::rules -> try if List.mem keyrule !print_non_active_notations then raise No_match; - match t with + let loc = t.loc in + match t.v with | PatCstr (cstr,_,na) -> let p = apply_notation_to_pattern ?loc (ConstructRef cstr) - (match_notation_constr_cases_pattern (loc, t) pat) allscopes vars keyrule in + (match_notation_constr_cases_pattern t pat) allscopes vars keyrule in insert_pat_alias ?loc p na | PatVar Anonymous -> CAst.make ?loc @@ CPatAtom None | PatVar (Name id) -> CAst.make ?loc @@ CPatAtom (Some (Ident (loc,id))) with - No_match -> extern_notation_pattern allscopes vars (loc, t) rules + No_match -> extern_notation_pattern allscopes vars t rules let rec extern_notation_ind_pattern allscopes vars ind args = function | [] -> raise No_match @@ -565,7 +568,7 @@ let extern_args extern env args = List.map map args let match_coercion_app = function - | (loc, GApp ((_, GRef (r,_)),args)) -> Some (loc, r, 0, args) + | {loc; v = GApp ({ v = GRef (r,_) },args)} -> Some (loc, r, 0, args) | _ -> None let rec remove_coercions inctx c = @@ -587,13 +590,13 @@ let rec remove_coercions inctx c = been confused with ordinary application or would have need a surrounding context and the coercion to funclass would have been made explicit to match *) - if List.is_empty l then a' else Loc.tag ?loc @@ GApp (a',l) + if List.is_empty l then a' else CAst.make ?loc @@ GApp (a',l) | _ -> c with Not_found -> c) | _ -> c let rec flatten_application = function - | (loc, GApp ((_, GApp(a,l')),l)) -> flatten_application (Loc.tag ?loc @@ GApp (a,l'@l)) + | {loc; v = GApp ({ v = GApp(a,l')},l)} -> flatten_application (CAst.make ?loc @@ GApp (a,l'@l)) | a -> a (**********************************************************************) @@ -621,10 +624,10 @@ let extern_optimal_prim_token scopes r r' = let extended_glob_local_binder_of_decl loc = function | (p,bk,None,t) -> GLocalAssum (p,bk,t) - | (p,bk,Some x,(_,GHole ( _, Misctypes.IntroAnonymous, None))) -> GLocalDef (p,bk,x,None) + | (p,bk,Some x, { v = GHole ( _, Misctypes.IntroAnonymous, None) } ) -> GLocalDef (p,bk,x,None) | (p,bk,Some x,t) -> GLocalDef (p,bk,x,Some t) -let extended_glob_local_binder_of_decl ?loc u = Loc.tag ?loc (extended_glob_local_binder_of_decl loc u) +let extended_glob_local_binder_of_decl ?loc u = CAst.make ?loc (extended_glob_local_binder_of_decl loc u) (**********************************************************************) (* mapping glob_constr to constr_expr *) @@ -649,7 +652,7 @@ let rec extern inctx scopes vars r = let r'' = flatten_application r' in if !Flags.raw_print || !print_no_symbol then raise No_match; extern_notation scopes vars r'' (uninterp_notations r'') - with No_match -> CAst.map_from_loc (fun ?loc -> function + with No_match -> CAst.map_with_loc (fun ?loc -> function | GRef (ref,us) -> extern_global (select_stronger_impargs (implicits_of_global ref)) (extern_reference ?loc vars ref) (extern_universes us) @@ -667,7 +670,7 @@ let rec extern inctx scopes vars r = | GApp (f,args) -> (match f with - | (rloc, GRef (ref,us)) -> + | {loc = rloc; v = GRef (ref,us) } -> let subscopes = find_arguments_scope ref in let args = fill_arg_scopes args subscopes (snd scopes) in begin @@ -743,7 +746,7 @@ let rec extern inctx scopes vars r = let rtntypopt' = Option.map (extern_typ scopes vars') rtntypopt in let tml = List.map (fun (tm,(na,x)) -> let na' = match na,tm with - | Anonymous, (_, GVar id) -> + | Anonymous, { v = GVar id } -> begin match rtntypopt with | None -> None | Some ntn -> @@ -752,12 +755,12 @@ let rec extern inctx scopes vars r = else None end | Anonymous, _ -> None - | Name id, (_, GVar id') when Id.equal id id' -> None + | Name id, { v = GVar id' } when Id.equal id id' -> None | Name _, _ -> Some (Loc.tag na) in (sub_extern false scopes vars tm, na', Option.map (fun (loc,(ind,nal)) -> - let args = List.map (fun x -> Loc.tag @@ PatVar x) nal in + let args = List.map (fun x -> CAst.make @@ PatVar x) nal in let fullargs = add_cpatt_for_params ind args in extern_ind_pattern_in_scope scopes vars ind fullargs ) x)) @@ -848,14 +851,14 @@ and factorize_lambda inctx scopes vars na bk aty c = and extern_local_binder scopes vars = function [] -> ([],[],[]) - | (_, GLocalDef (na,bk,bd,ty))::l -> + | { v = GLocalDef (na,bk,bd,ty)}::l -> let (assums,ids,l) = extern_local_binder scopes (name_fold Id.Set.add na vars) l in (assums,na::ids, CLocalDef((Loc.tag na), extern false scopes vars bd, Option.map (extern false scopes vars) ty) :: l) - | (_, GLocalAssum (na,bk,ty))::l -> + | { v = GLocalAssum (na,bk,ty)}::l -> let ty = extern_typ scopes vars ty in (match extern_local_binder scopes (name_fold Id.Set.add na vars) l with (assums,ids,CLocalAssum(nal,k,ty')::l) @@ -868,7 +871,7 @@ and extern_local_binder scopes vars = function (na::assums,na::ids, CLocalAssum([(Loc.tag na)],Default bk,ty) :: l)) - | (_, GLocalPattern ((p,_),_,bk,ty))::l -> + | { v = GLocalPattern ((p,_),_,bk,ty)}::l -> let ty = if !Flags.raw_print then Some (extern_typ scopes vars ty) else None in let p = extern_cases_pattern vars p in @@ -886,12 +889,12 @@ and extern_notation (tmp_scope,scopes as allscopes) vars t = function try if List.mem keyrule !print_non_active_notations then raise No_match; (* Adjusts to the number of arguments expected by the notation *) - let (t,args,argsscopes,argsimpls) = match snd t,n with + let (t,args,argsscopes,argsimpls) = match t.v ,n with | GApp (f,args), Some n when List.length args >= n -> let args1, args2 = List.chop n args in let subscopes, impls = - match snd f with + match f.v with | GRef (ref,us) -> let subscopes = try List.skipn n (find_arguments_scope ref) @@ -904,15 +907,15 @@ and extern_notation (tmp_scope,scopes as allscopes) vars t = function subscopes,impls | _ -> [], [] in - (if Int.equal n 0 then f else Loc.tag @@ GApp (f,args1)), + (if Int.equal n 0 then f else CAst.make @@ GApp (f,args1)), args2, subscopes, impls - | GApp ((_, GRef (ref,us) as f),args), None -> + | GApp ({ v = GRef (ref,us) } as f, args), None -> let subscopes = find_arguments_scope ref in let impls = select_impargs_size (List.length args) (implicits_of_global ref) in f, args, subscopes, impls - | GRef (ref,us), Some 0 -> Loc.tag @@ GApp (t,[]), [], [], [] + | GRef (ref,us), Some 0 -> CAst.make @@ GApp (t,[]), [], [], [] | _, None -> t, [], [], [] | _ -> raise No_match in (* Try matching ... *) @@ -1014,9 +1017,9 @@ let extern_closed_glob ?lax goal_concl_style env sigma t = let any_any_branch = (* | _ => _ *) - Loc.tag ([],[Loc.tag @@ PatVar Anonymous], Loc.tag @@ GHole (Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None)) + Loc.tag ([],[CAst.make @@ PatVar Anonymous], CAst.make @@ GHole (Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None)) -let rec glob_of_pat env sigma pat = Loc.tag @@ match pat with +let rec glob_of_pat env sigma pat = CAst.make @@ match pat with | PRef ref -> GRef (ref,None) | PVar id -> GVar id | PEvar (evk,l) -> @@ -1036,12 +1039,12 @@ let rec glob_of_pat env sigma pat = Loc.tag @@ match pat with GVar id | PMeta None -> GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous,None) | PMeta (Some n) -> GPatVar (false,n) - | PProj (p,c) -> GApp (Loc.tag @@ GRef (ConstRef (Projection.constant p),None), + | PProj (p,c) -> GApp (CAst.make @@ GRef (ConstRef (Projection.constant p),None), [glob_of_pat env sigma c]) | PApp (f,args) -> GApp (glob_of_pat env sigma f,Array.map_to_list (glob_of_pat env sigma) args) | PSoApp (n,args) -> - GApp (Loc.tag @@ GPatVar (true,n), + GApp (CAst.make @@ GPatVar (true,n), List.map (glob_of_pat env sigma) args) | PProd (na,t,c) -> GProd (na,Explicit,glob_of_pat env sigma t,glob_of_pat (na::env) sigma c) @@ -1073,8 +1076,8 @@ let rec glob_of_pat env sigma pat = Loc.tag @@ match pat with | _ -> anomaly (Pp.str "PCase with non-trivial predicate but unknown inductive") in GCases (RegularStyle,rtn,[glob_of_pat env sigma tm,indnames],mat) - | PFix f -> Loc.obj @@ Detyping.detype_names false [] env (Global.env()) sigma (EConstr.of_constr (mkFix f)) (** FIXME bad env *) - | PCoFix c -> Loc.obj @@ Detyping.detype_names false [] env (Global.env()) sigma (EConstr.of_constr (mkCoFix c)) + | PFix f -> (Detyping.detype_names false [] env (Global.env()) sigma (EConstr.of_constr (mkFix f))).v (** FIXME bad env *) + | PCoFix c -> (Detyping.detype_names false [] env (Global.env()) sigma (EConstr.of_constr (mkCoFix c))).v | PSort s -> GSort s let extern_constr_pattern env sigma pat = diff --git a/interp/constrintern.ml b/interp/constrintern.ml index 541b52972..bd7c05e6f 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -15,6 +15,7 @@ open Namegen open Libnames open Globnames open Impargs +open CAst open Glob_term open Glob_ops open Patternops @@ -304,12 +305,12 @@ let reset_tmp_scope env = {env with tmp_scope = None} let rec it_mkGProd ?loc env body = match env with - (loc2, (na, bk, t)) :: tl -> it_mkGProd ?loc:loc2 tl (Loc.tag ?loc:(Loc.merge_opt loc loc2) @@ GProd (na, bk, t, body)) + (loc2, (na, bk, t)) :: tl -> it_mkGProd ?loc:loc2 tl (CAst.make ?loc:(Loc.merge_opt loc loc2) @@ GProd (na, bk, t, body)) | [] -> body let rec it_mkGLambda ?loc env body = match env with - (loc2, (na, bk, t)) :: tl -> it_mkGLambda ?loc:loc2 tl (Loc.tag ?loc:(Loc.merge_opt loc loc2) @@ GLambda (na, bk, t, body)) + (loc2, (na, bk, t)) :: tl -> it_mkGLambda ?loc:loc2 tl (CAst.make ?loc:(Loc.merge_opt loc loc2) @@ GLambda (na, bk, t, body)) | [] -> body (**********************************************************************) @@ -322,14 +323,14 @@ let build_impls = function let impls_type_list ?(args = []) = let rec aux acc = function - |_, GProd (na,bk,_,c) -> aux ((build_impls bk na)::acc) c - |_ -> (Variable,[],List.append args (List.rev acc),[]) + | { v = GProd (na,bk,_,c) } -> aux ((build_impls bk na)::acc) c + | _ -> (Variable,[],List.append args (List.rev acc),[]) in aux [] let impls_term_list ?(args = []) = let rec aux acc = function - |_, GLambda (na,bk,_,c) -> aux ((build_impls bk na)::acc) c - |_, GRec (fix_kind, nas, args, tys, bds) -> + | { v = GLambda (na,bk,_,c) } -> aux ((build_impls bk na)::acc) c + | { v = GRec (fix_kind, nas, args, tys, bds) } -> let nb = match fix_kind with |GFix (_, n) -> n | GCoFix n -> n in let acc' = List.fold_left (fun a (na, bk, _, _) -> (build_impls bk na)::a) acc args.(nb) in aux acc' bds.(nb) @@ -346,12 +347,12 @@ let rec check_capture ty = function () let locate_if_hole ?loc na = function - | _, GHole (_,naming,arg) -> + | { v = GHole (_,naming,arg) } -> (try match na with | Name id -> glob_constr_of_notation_constr ?loc (Reserve.find_reserved_type id) | Anonymous -> raise Not_found - with Not_found -> Loc.tag ?loc @@ GHole (Evar_kinds.BinderType na, naming, arg)) + with Not_found -> CAst.make ?loc @@ GHole (Evar_kinds.BinderType na, naming, arg)) | x -> x let reset_hidden_inductive_implicit_test env = @@ -397,7 +398,7 @@ let intern_generalized_binder ?(global_level=false) intern_type lvar env fvs in let bl = List.map (fun (loc, id) -> - (loc, (Name id, b, Loc.tag ?loc @@ GHole (Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None)))) + (loc, (Name id, b, CAst.make ?loc @@ GHole (Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None)))) fvs in let na = match na with @@ -454,11 +455,11 @@ let intern_local_pattern intern lvar env p = env) env (free_vars_of_pat [] p) -let glob_local_binder_of_extended = Loc.with_loc (fun ?loc -> function +let glob_local_binder_of_extended = CAst.with_loc_val (fun ?loc -> function | GLocalAssum (na,bk,t) -> (na,bk,None,t) | GLocalDef (na,bk,c,Some t) -> (na,bk,Some c,t) | GLocalDef (na,bk,c,None) -> - let t = Loc.tag ?loc @@ GHole(Evar_kinds.BinderType na,Misctypes.IntroAnonymous,None) in + let t = CAst.make ?loc @@ GHole(Evar_kinds.BinderType na,Misctypes.IntroAnonymous,None) in (na,bk,Some c,t) | GLocalPattern (_,_,_,_) -> Loc.raise ?loc (Stream.Error "pattern with quote not allowed here.") @@ -469,13 +470,13 @@ let intern_cases_pattern_fwd = ref (fun _ -> failwith "intern_cases_pattern_fwd" let intern_local_binder_aux ?(global_level=false) intern lvar (env,bl) = function | CLocalAssum(nal,bk,ty) -> let env, bl' = intern_assumption intern lvar env nal bk ty in - let bl' = List.map (fun (loc,(na,c,t)) -> Loc.tag ?loc @@ GLocalAssum (na,c,t)) bl' in + let bl' = List.map (fun (loc,(na,c,t)) -> CAst.make ?loc @@ GLocalAssum (na,c,t)) bl' in env, bl' @ bl | CLocalDef((loc,na as locna),def,ty) -> let term = intern env def in let ty = Option.map (intern env) ty in (push_name_env lvar (impls_term_list term) env locna, - (Loc.tag ?loc @@ GLocalDef (na,Explicit,term,ty)) :: bl) + (CAst.make ?loc @@ GLocalDef (na,Explicit,term,ty)) :: bl) | CLocalPattern (loc,(p,ty)) -> let tyc = match ty with @@ -495,7 +496,7 @@ let intern_local_binder_aux ?(global_level=false) intern lvar (env,bl) = functio let bk = Default Explicit in let _, bl' = intern_assumption intern lvar env [na] bk tyc in let _,(_,bk,t) = List.hd bl' in - (env, (Loc.tag ?loc @@ GLocalPattern((cp,il),id,bk,t)) :: bl) + (env, (CAst.make ?loc @@ GLocalPattern((cp,il),id,bk,t)) :: bl) let intern_generalization intern env lvar loc bk ak c = let c = intern {env with unb = true} c in @@ -518,12 +519,12 @@ let intern_generalization intern env lvar loc bk ak c = in if pi then (fun (loc', id) acc -> - Loc.tag ?loc:(Loc.merge_opt loc' loc) @@ - GProd (Name id, bk, Loc.tag ?loc:loc' @@ GHole (Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None), acc)) + CAst.make ?loc:(Loc.merge_opt loc' loc) @@ + GProd (Name id, bk, CAst.make ?loc:loc' @@ GHole (Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None), acc)) else (fun (loc', id) acc -> - Loc.tag ?loc:(Loc.merge_opt loc' loc) @@ - GLambda (Name id, bk, Loc.tag ?loc:loc' @@ GHole (Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None), acc)) + CAst.make ?loc:(Loc.merge_opt loc' loc) @@ + GLambda (Name id, bk, CAst.make ?loc:loc' @@ GHole (Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None), acc)) in List.fold_right (fun (loc, id as lid) (env, acc) -> let env' = push_name_env lvar (Variable,[],[],[]) env (loc, Name id) in @@ -576,27 +577,27 @@ let make_letins = (fun a c -> match a with | loc, LPLetIn (na,b,t) -> - Loc.tag ?loc @@ GLetIn(na,b,t,c) + CAst.make ?loc @@ GLetIn(na,b,t,c) | loc, LPCases ((cp,il),id) -> - let tt = (Loc.tag ?loc @@ GVar id, (Name id,None)) in - Loc.tag ?loc @@ GCases(Misctypes.LetPatternStyle,None,[tt],[(loc,(il,[cp],c))])) + let tt = (CAst.make ?loc @@ GVar id, (Name id,None)) in + CAst.make ?loc @@ GCases(Misctypes.LetPatternStyle,None,[tt],[(loc,(il,[cp],c))])) let rec subordinate_letins letins = function (* binders come in reverse order; the non-let are returned in reverse order together *) (* with the subordinated let-in in writing order *) - | (loc, GLocalDef (na,_,b,t))::l -> + | { loc; v = GLocalDef (na,_,b,t) }::l -> subordinate_letins ((Loc.tag ?loc @@ LPLetIn (na,b,t))::letins) l - | (loc, GLocalAssum (na,bk,t))::l -> + | { loc; v = GLocalAssum (na,bk,t)}::l -> let letins',rest = subordinate_letins [] l in letins',((loc,(na,bk,t)),letins)::rest - | (loc, GLocalPattern (u,id,bk,t)) :: l -> + | { loc; v = GLocalPattern (u,id,bk,t)} :: l -> subordinate_letins ((Loc.tag ?loc @@ LPCases (u,id))::letins) - ([Loc.tag ?loc @@ GLocalAssum (Name id,bk,t)] @ l) + ([CAst.make ?loc @@ GLocalAssum (Name id,bk,t)] @ l) | [] -> letins,[] let terms_of_binders bl = - let rec term_of_pat pt = CAst.map_from_loc (fun ?loc -> function + let rec term_of_pat pt = CAst.map_with_loc (fun ?loc -> function | PatVar (Name id) -> CRef (Ident (loc,id), None) | PatVar (Anonymous) -> error "Cannot turn \"_\" into a term." | PatCstr (c,l,_) -> @@ -605,11 +606,11 @@ let terms_of_binders bl = let params = List.make (Inductiveops.inductive_nparams (fst c)) hole in CAppExpl ((None,r,None),params @ List.map term_of_pat l)) pt in let rec extract_variables = function - | (loc, GLocalAssum (Name id,_,_))::l -> (CAst.make ?loc @@ CRef (Ident (loc,id), None)) :: extract_variables l - | (loc, GLocalDef (Name id,_,_,_))::l -> extract_variables l - | (loc, GLocalDef (Anonymous,_,_,_))::l - | (loc, GLocalAssum (Anonymous,_,_))::l -> error "Cannot turn \"_\" into a term." - | (loc, GLocalPattern ((u,_),_,_,_)) :: l -> term_of_pat u :: extract_variables l + | {loc; v = GLocalAssum (Name id,_,_)}::l -> (CAst.make ?loc @@ CRef (Ident (loc,id), None)) :: extract_variables l + | {loc; v = GLocalDef (Name id,_,_,_)}::l -> extract_variables l + | {loc; v = GLocalDef (Anonymous,_,_,_)}::l + | {loc; v = GLocalAssum (Anonymous,_,_)}::l -> error "Cannot turn \"_\" into a term." + | {loc; v = GLocalPattern ((u,_),_,_,_)}::l -> term_of_pat u :: extract_variables l | [] -> [] in extract_variables bl @@ -665,7 +666,7 @@ let instantiate_notation_constr loc intern ntnvars subst infos c = let bindings = Id.Map.map mk_env terms in Some (Genintern.generic_substitute_notation bindings arg) in - Loc.tag ?loc @@ GHole (knd, naming, arg) + CAst.make ?loc @@ GHole (knd, naming, arg) | NBinderList (x,y,iter,terminator) -> (try (* All elements of the list are in scopes (scopt,subscopes) *) @@ -683,22 +684,22 @@ let instantiate_notation_constr loc intern ntnvars subst infos c = let a,letins = snd (Option.get binderopt) in let e = make_letins letins (aux subst' infos c') in let (loc,(na,bk,t)) = a in - Loc.tag ?loc @@ GProd (na,bk,t,e) + CAst.make ?loc @@ GProd (na,bk,t,e) | NLambda (Name id,NHole _,c') when option_mem_assoc id binderopt -> let a,letins = snd (Option.get binderopt) in let (loc,(na,bk,t)) = a in - Loc.tag ?loc @@ GLambda (na,bk,t,make_letins letins (aux subst' infos c')) + CAst.make ?loc @@ GLambda (na,bk,t,make_letins letins (aux subst' infos c')) (* Two special cases to keep binder name synchronous with BinderType *) | NProd (na,NHole(Evar_kinds.BinderType na',naming,arg),c') when Name.equal na na' -> let subinfos,na = traverse_binder subst avoid subinfos na in - let ty = Loc.tag ?loc @@ GHole (Evar_kinds.BinderType na,naming,arg) in - Loc.tag ?loc @@ GProd (na,Explicit,ty,aux subst' subinfos c') + let ty = CAst.make ?loc @@ GHole (Evar_kinds.BinderType na,naming,arg) in + CAst.make ?loc @@ GProd (na,Explicit,ty,aux subst' subinfos c') | NLambda (na,NHole(Evar_kinds.BinderType na',naming,arg),c') when Name.equal na na' -> let subinfos,na = traverse_binder subst avoid subinfos na in - let ty = Loc.tag ?loc @@ GHole (Evar_kinds.BinderType na,naming,arg) in - Loc.tag ?loc @@ GLambda (na,Explicit,ty,aux subst' subinfos c') + let ty = CAst.make ?loc @@ GHole (Evar_kinds.BinderType na,naming,arg) in + CAst.make ?loc @@ GLambda (na,Explicit,ty,aux subst' subinfos c') | t -> glob_constr_of_notation_constr_with_binders ?loc (traverse_binder subst avoid) (aux subst') subinfos t @@ -710,7 +711,7 @@ let instantiate_notation_constr loc intern ntnvars subst infos c = intern {env with tmp_scope = scopt; scopes = subscopes @ env.scopes} a with Not_found -> - Loc.tag ?loc ( + CAst.make ?loc ( try GVar (Id.Map.find id renaming) with Not_found -> @@ -750,7 +751,7 @@ let string_of_ty = function | Variable -> "var" let gvar (loc, id) us = match us with -| None -> Loc.tag ?loc @@ GVar id +| None -> CAst.make ?loc @@ GVar id | Some _ -> user_err ?loc (str "Variable " ++ pr_id id ++ str " cannot have a universe instance") @@ -792,18 +793,18 @@ let intern_var genv (ltacvars,ntnvars) namedctx loc id us = let impls = implicits_of_global ref in let scopes = find_arguments_scope ref in Dumpglob.dump_reference ?loc "<>" (string_of_qualid (Decls.variable_secpath id)) "var"; - Loc.tag ?loc @@ GRef (ref, us), impls, scopes, [] + CAst.make ?loc @@ GRef (ref, us), impls, scopes, [] with e when CErrors.noncritical e -> (* [id] a goal variable *) gvar (loc,id) us, [], [], [] let find_appl_head_data c = - match Loc.obj c with + match c.v with | GRef (ref,_) -> let impls = implicits_of_global ref in let scopes = find_arguments_scope ref in c, impls, scopes, [] - | GApp ((_, GRef (ref,_)),l) + | GApp ({ v = GRef (ref,_) },l) when l != [] && Flags.version_strictly_greater Flags.V8_2 -> let n = List.length l in let impls = implicits_of_global ref in @@ -842,7 +843,7 @@ let intern_reference ref = (* Is it a global reference or a syntactic definition? *) let intern_qualid loc qid intern env lvar us args = match intern_extended_global_of_qualid (loc,qid) with - | TrueGlobal ref -> (Loc.tag ?loc @@ GRef (ref, us)), true, args + | TrueGlobal ref -> (CAst.make ?loc @@ GRef (ref, us)), true, args | SynDef sp -> let (ids,c) = Syntax_def.search_syntactic_definition sp in let nids = List.length ids in @@ -856,9 +857,9 @@ let intern_qualid loc qid intern env lvar us args = let c = instantiate_notation_constr loc intern lvar subst infos c in let c = match us, c with | None, _ -> c - | Some _, (loc, GRef (ref, None)) -> Loc.tag ?loc @@ GRef (ref, us) - | Some _, (loc, GApp ((loc', GRef (ref, None)), arg)) -> - Loc.tag ?loc @@ GApp (Loc.tag ?loc:loc' @@ GRef (ref, us), arg) + | Some _, { loc; v = GRef (ref, None) } -> CAst.make ?loc @@ GRef (ref, us) + | Some _, { loc; v = GApp ({ loc = loc' ; v = GRef (ref, None) }, arg) } -> + CAst.make ?loc @@ GApp (CAst.make ?loc:loc' @@ GRef (ref, us), arg) | Some _, _ -> user_err ?loc (str "Notation " ++ pr_qualid qid ++ str " cannot have a universe instance, its expanded head @@ -869,7 +870,7 @@ let intern_qualid loc qid intern env lvar us args = (* Rule out section vars since these should have been found by intern_var *) let intern_non_secvar_qualid loc qid intern env lvar us args = match intern_qualid loc qid intern env lvar us args with - | (_, GRef (VarRef _, _)),_,_ -> raise Not_found + | { v = GRef (VarRef _, _) },_,_ -> raise Not_found | r -> r let intern_applied_reference intern env namedctx (_, ntnvars as lvar) us args = function @@ -1021,8 +1022,8 @@ let chop_params_pattern loc ind args with_letin = else Inductiveops.inductive_nparams ind in assert (nparams <= List.length args); let params,args = List.chop nparams args in - List.iter (function _, PatVar Anonymous -> () - | loc', PatVar _ | loc', PatCstr(_,_,_) -> error_parameter_not_implicit ?loc:loc') params; + List.iter (function { v = PatVar Anonymous } -> () + | { loc; v = PatVar _ } | { loc; v = PatCstr(_,_,_) } -> error_parameter_not_implicit ?loc) params; args let find_constructor loc add_params ref = @@ -1042,7 +1043,7 @@ let find_constructor loc add_params ref = then Inductiveops.inductive_nparamdecls ind else Inductiveops.inductive_nparams ind in - List.make nb ([], [(Id.Map.empty, Loc.tag @@ PatVar Anonymous)]) + List.make nb ([], [(Id.Map.empty, CAst.make @@ PatVar Anonymous)]) | None -> [] let find_pattern_variable = function @@ -1368,7 +1369,7 @@ let rec intern_pat genv aliases pat = let idslpl2 = List.map (intern_pat genv empty_alias) pl2 in let (ids',pll) = product_of_cases_patterns aliases.alias_ids (idslpl1@idslpl2) in let pl' = List.map (fun (asubst,pl) -> - (asubst, Loc.tag ?loc @@ PatCstr (c,chop_params_pattern loc (fst c) pl with_letin,alias_of aliases))) pll in + (asubst, CAst.make ?loc @@ PatCstr (c,chop_params_pattern loc (fst c) pl with_letin,alias_of aliases))) pll in ids',pl' in let loc = CAst.(pat.loc) in match CAst.(pat.v) with @@ -1389,10 +1390,10 @@ let rec intern_pat genv aliases pat = intern_cstr_with_all_args loc c with_letin idslpl1 (expl_pl@pl2) | RCPatAtom (Some id) -> let aliases = merge_aliases aliases id in - (aliases.alias_ids,[aliases.alias_map, Loc.tag ?loc @@ PatVar (alias_of aliases)]) + (aliases.alias_ids,[aliases.alias_map, CAst.make ?loc @@ PatVar (alias_of aliases)]) | RCPatAtom (None) -> let { alias_ids = ids; alias_map = asubst; } = aliases in - (ids, [asubst, Loc.tag ?loc @@ PatVar (alias_of aliases)]) + (ids, [asubst, CAst.make ?loc @@ PatVar (alias_of aliases)]) | RCPatOr pl -> assert (not (List.is_empty pl)); let pl' = List.map (intern_pat genv aliases) pl in @@ -1482,8 +1483,8 @@ let get_implicit_name n imps = Some (Impargs.name_of_implicit (List.nth imps (n-1))) let set_hole_implicit i b = function - | (loc, GRef (r,_)) | (_, GApp ((loc, (GRef (r,_))),_)) -> Loc.tag ?loc (Evar_kinds.ImplicitArg (r,i,b),Misctypes.IntroAnonymous,None) - | (loc, GVar id) -> Loc.tag ?loc (Evar_kinds.ImplicitArg (VarRef id,i,b),Misctypes.IntroAnonymous,None) + | {loc; v = GRef (r,_) } | { v = GApp ({loc; v = GRef (r,_)},_) } -> Loc.tag ?loc (Evar_kinds.ImplicitArg (r,i,b),Misctypes.IntroAnonymous,None) + | {loc; v = GVar id } -> Loc.tag ?loc (Evar_kinds.ImplicitArg (VarRef id,i,b),Misctypes.IntroAnonymous,None) | _ -> anomaly (Pp.str "Only refs have implicits") let exists_implicit_name id = @@ -1549,7 +1550,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let before, after = split_at_annot bl n in let (env',rbefore) = List.fold_left intern_local_binder (env,[]) before in let ro = f (intern env') in - let n' = Option.map (fun _ -> List.count (function | _, GLocalAssum _ -> true + let n' = Option.map (fun _ -> List.count (function | { v = GLocalAssum _ } -> true | _ -> false (* remove let-ins *)) rbefore) n in n', ro, List.fold_left intern_local_binder (env',rbefore) after @@ -1572,7 +1573,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = push_name_env ntnvars (impls_type_list ~args:fix_args tyi) en (Loc.tag @@ Name name)) 0 env' lf in (a,b,c,intern {env'' with tmp_scope = None} bd)) dl idl_temp in - Loc.tag ?loc @@ + CAst.make ?loc @@ GRec (GFix (Array.map (fun (ro,_,_,_) -> ro) idl,n), Array.of_list lf, @@ -1599,7 +1600,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = push_name_env ntnvars (impls_type_list ~args:cofix_args tyi) en (Loc.tag @@ Name name)) 0 env' lf in (b,c,intern {env'' with tmp_scope = None} bd)) dl idl_tmp in - Loc.tag ?loc @@ + CAst.make ?loc @@ GRec (GCoFix n, Array.of_list lf, Array.map (fun (bl,_,_) -> bl) idl, @@ -1616,7 +1617,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = | CLetIn (na,c1,t,c2) -> let inc1 = intern (reset_tmp_scope env) c1 in let int = Option.map (intern_type env) t in - Loc.tag ?loc @@ + CAst.make ?loc @@ GLetIn (snd na, inc1, int, intern (push_name_env ntnvars (impls_term_list inc1) env na) c2) | CNotation ("- _",([{ CAst.v = CPrim (Numeral p) }],[],[])) @@ -1639,7 +1640,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = lvar us args ref in (* Rem: GApp(_,f,[]) stands for @f *) - Loc.tag ?loc @@ + CAst.make ?loc @@ GApp (f, intern_args env args_scopes (List.map fst args)) | CApp ((isproj,f), args) -> @@ -1696,7 +1697,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let stripped_match_from_in = let rec aux = function | [] -> [] - | (_, (_loc, PatVar _)) :: q -> aux q + | (_, { v = PatVar _}) :: q -> aux q | l -> l in aux match_from_in in let rtnpo = match stripped_match_from_in with @@ -1705,20 +1706,20 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = (* Build a return predicate by expansion of the patterns of the "in" clause *) let thevars, thepats = List.split l in let sub_rtn = (* Some (GSort (Loc.ghost,GType None)) *) None in - let sub_tms = List.map (fun id -> (Loc.tag @@ GVar id),(Name id,None)) thevars (* "match v1,..,vn" *) in + let sub_tms = List.map (fun id -> (CAst.make @@ GVar id),(Name id,None)) thevars (* "match v1,..,vn" *) in let main_sub_eqn = Loc.tag @@ ([],thepats, (* "|p1,..,pn" *) Option.cata (intern_type env') - (Loc.tag ?loc @@ GHole(Evar_kinds.CasesType false,Misctypes.IntroAnonymous,None)) + (CAst.make ?loc @@ GHole(Evar_kinds.CasesType false,Misctypes.IntroAnonymous,None)) rtnpo) (* "=> P" if there were a return predicate P, and "=> _" otherwise *) in let catch_all_sub_eqn = if List.for_all (irrefutable globalenv) thepats then [] else - [Loc.tag @@ ([],List.make (List.length thepats) (Loc.tag @@ PatVar Anonymous), (* "|_,..,_" *) - Loc.tag @@ GHole(Evar_kinds.ImpossibleCase,Misctypes.IntroAnonymous,None))] (* "=> _" *) in - Some (Loc.tag @@ GCases(Term.RegularStyle,sub_rtn,sub_tms,main_sub_eqn::catch_all_sub_eqn)) + [Loc.tag @@ ([],List.make (List.length thepats) (CAst.make @@ PatVar Anonymous), (* "|_,..,_" *) + CAst.make @@ GHole(Evar_kinds.ImpossibleCase,Misctypes.IntroAnonymous,None))] (* "=> _" *) in + Some (CAst.make @@ GCases(Term.RegularStyle,sub_rtn,sub_tms,main_sub_eqn::catch_all_sub_eqn)) in let eqns' = List.map (intern_eqn (List.length tms) env) eqns in - Loc.tag ?loc @@ + CAst.make ?loc @@ GCases (sty, rtnpo, tms, List.flatten eqns') | CLetTuple (nal, (na,po), b, c) -> let env' = reset_tmp_scope env in @@ -1728,7 +1729,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let env'' = push_name_env ntnvars (Variable,[],[],[]) (reset_hidden_inductive_implicit_test env') (Loc.tag na') in intern_type env'' u) po in - Loc.tag ?loc @@ + CAst.make ?loc @@ GLetTuple (List.map snd nal, (na', p'), b', intern (List.fold_left (push_name_env ntnvars (Variable,[],[],[])) (reset_hidden_inductive_implicit_test env) nal) c) | CIf (c, (na,po), b1, b2) -> @@ -1738,7 +1739,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let env'' = push_name_env ntnvars (Variable,[],[],[]) (reset_hidden_inductive_implicit_test env) (Loc.tag na') in intern_type env'' p) po in - Loc.tag ?loc @@ + CAst.make ?loc @@ GIf (c', (na', p'), intern env b1, intern env b2) | CHole (k, naming, solve) -> let k = match k with @@ -1764,28 +1765,28 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let (_, glb) = Genintern.generic_intern ist gen in Some glb in - Loc.tag ?loc @@ + CAst.make ?loc @@ GHole (k, naming, solve) (* Parsing pattern variables *) | CPatVar n when allow_patvar -> - Loc.tag ?loc @@ + CAst.make ?loc @@ GPatVar (true,n) | CEvar (n, []) when allow_patvar -> - Loc.tag ?loc @@ + CAst.make ?loc @@ GPatVar (false,n) (* end *) (* Parsing existential variables *) | CEvar (n, l) -> - Loc.tag ?loc @@ + CAst.make ?loc @@ GEvar (n, List.map (on_snd (intern env)) l) | CPatVar _ -> raise (InternalizationError (loc,IllegalMetavariable)) (* end *) | CSort s -> - Loc.tag ?loc @@ + CAst.make ?loc @@ GSort s | CCast (c1, c2) -> - Loc.tag ?loc @@ + CAst.make ?loc @@ GCast (intern env c1, Miscops.map_cast_type (intern_type env) c2) ) and intern_type env = intern (set_type_scope env) @@ -1824,8 +1825,8 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let tm' = intern env tm in (* the "as" part *) let extra_id,na = match tm', na with - | (loc , GVar id), None when not (Id.Map.mem id (snd lvar)) -> Some id,(loc,Name id) - | (loc, GRef (VarRef id, _)), None -> Some id,(loc,Name id) + | {loc; v = GVar id}, None when not (Id.Map.mem id (snd lvar)) -> Some id,(loc,Name id) + | {loc; v = GRef (VarRef id, _)}, None -> Some id,(loc,Name id) | _, None -> None,(Loc.tag Anonymous) | _, Some (loc,na) -> None,(loc,na) in (* the "in" part *) @@ -1844,14 +1845,14 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let rec canonize_args case_rel_ctxt arg_pats forbidden_names match_acc var_acc = let add_name l = function | _,Anonymous -> l - | loc,(Name y as x) -> (y, Loc.tag ?loc @@ PatVar x) :: l in + | loc,(Name y as x) -> (y, CAst.make ?loc @@ PatVar x) :: l in match case_rel_ctxt,arg_pats with (* LetIn in the rel_context *) | LocalDef _ :: t, l when not with_letin -> canonize_args t l forbidden_names match_acc ((Loc.tag Anonymous)::var_acc) | [],[] -> (add_name match_acc na, var_acc) - | _::t, (loc, PatVar x)::tt -> + | _::t, { loc; v = PatVar x}::tt -> canonize_args t tt forbidden_names (add_name match_acc (loc,x)) ((loc,x)::var_acc) | (LocalAssum (cano_name,ty) | LocalDef (cano_name,_,ty)) :: t, c::tt -> @@ -1897,7 +1898,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = (* with implicit arguments if maximal insertion is set *) [] else - (Loc.map (fun (a,b,c) -> GHole(a,b,c)) + (CAst.map_from_loc (fun ?loc (a,b,c) -> GHole(a,b,c)) (set_hole_implicit (n,get_implicit_name n l) (force_inference_of imp) c) ) :: aux (n+1) impl' subscopes' eargs rargs end @@ -1924,8 +1925,8 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = and smart_gapp f loc = function | [] -> f | l -> match f with - | (loc', GApp (g, args)) -> Loc.tag ?loc:(Loc.merge_opt loc' loc) @@ GApp (g, args@l) - | _ -> Loc.tag ?loc:(Loc.merge_opt (loc_of_glob_constr f) loc) @@ GApp (f, l) + | { loc = loc'; v = GApp (g, args) } -> CAst.make ?loc:(Loc.merge_opt loc' loc) @@ GApp (g, args@l) + | _ -> CAst.make ?loc:(Loc.merge_opt (loc_of_glob_constr f) loc) @@ GApp (f, l) and intern_args env subscopes = function | [] -> [] diff --git a/interp/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml index 52a6c450b..deb567865 100644 --- a/interp/implicit_quantifiers.ml +++ b/interp/implicit_quantifiers.ml @@ -125,7 +125,7 @@ let add_name_to_ids set na = | Name id -> Id.Set.add id set let generalizable_vars_of_glob_constr ?(bound=Id.Set.empty) ?(allowed=Id.Set.empty) = - let rec vars bound vs (loc, t) = match t with + let rec vars bound vs { loc; CAst.v = t } = match t with | GVar id -> if is_freevar bound (Global.env ()) id then if Id.List.mem_assoc_sym id vs then vs @@ -314,7 +314,7 @@ let implicits_of_glob_constr ?(with_products=true) l = (ExplByPos (i, name), (true, true, true)) :: l | _ -> l in - let rec aux i (loc, c) = + let rec aux i { loc; CAst.v = c } = let abs na bk b = add_impl i na bk (aux (succ i) b) in diff --git a/interp/modintern.ml b/interp/modintern.ml index 45e6cd06c..3115c2bcb 100644 --- a/interp/modintern.ml +++ b/interp/modintern.ml @@ -65,14 +65,14 @@ let transl_with_decl env = function let ctx = Evd.evar_context_universe_context ectx in WithDef (fqid,(c,ctx)) -let loc_of_module (l, _) = l +let loc_of_module l = l.CAst.loc (* Invariant : the returned kind is never ModAny, and it is equal to the input kind when this one isn't ModAny. *) -let rec interp_module_ast env kind (loc, m) = match m with +let rec interp_module_ast env kind m = match m.CAst.v with | CMident qid -> - let (mp,kind) = lookup_module_or_modtype kind (loc,qid) in + let (mp,kind) = lookup_module_or_modtype kind (m.CAst.loc,qid) in (MEident mp, kind) | CMapply (me1,me2) -> let me1',kind1 = interp_module_ast env kind me1 in @@ -86,6 +86,6 @@ let rec interp_module_ast env kind (loc, m) = match m with (MEapply (me1',mp2), kind1) | CMwith (me,decl) -> let me,kind = interp_module_ast env kind me in - if kind == Module then error_incorrect_with_in_module loc; + if kind == Module then error_incorrect_with_in_module m.CAst.loc; let decl = transl_with_decl env decl in (MEwith(me,decl), kind) diff --git a/interp/notation.ml b/interp/notation.ml index 03dffa6ee..6b963b8c8 100644 --- a/interp/notation.ml +++ b/interp/notation.ml @@ -264,16 +264,16 @@ let notations_key_table = ref (KeyMap.empty : notation_rule list KeyMap.t) let prim_token_key_table = ref KeyMap.empty let glob_prim_constr_key = function - | _, GApp ((_, GRef (ref,_)),_) | _, GRef (ref,_) -> RefKey (canonical_gr ref) + | { CAst.v = GApp ({ CAst.v = GRef (ref,_) } ,_) } | { CAst.v = GRef (ref,_) } -> RefKey (canonical_gr ref) | _ -> Oth let glob_constr_keys = function - | _, GApp ((_, GRef (ref,_)),_) -> [RefKey (canonical_gr ref); Oth] - | _, GRef (ref,_) -> [RefKey (canonical_gr ref)] + | { CAst.v = GApp ({ CAst.v = GRef (ref,_) },_) } -> [RefKey (canonical_gr ref); Oth] + | { CAst.v = GRef (ref,_) } -> [RefKey (canonical_gr ref)] | _ -> [Oth] let cases_pattern_key = function - | _, PatCstr (ref,_,_) -> RefKey (canonical_gr (ConstructRef ref)) + | { CAst.v = PatCstr (ref,_,_) } -> RefKey (canonical_gr (ConstructRef ref)) | _ -> Oth let notation_constr_key = function (* Rem: NApp(NRef ref,[]) stands for @ref *) @@ -471,14 +471,14 @@ let interp_prim_token = (** [rcp_of_glob] : from [glob_constr] to [raw_cases_pattern_expr] *) -let rec rcp_of_glob looked_for gt = CAst.map_from_loc (fun ?loc -> function +let rec rcp_of_glob looked_for = CAst.map (function | GVar id -> RCPatAtom (Some id) | GHole (_,_,_) -> RCPatAtom None | GRef (g,_) -> looked_for g; RCPatCstr (g,[],[]) - | GApp ((_, GRef (g,_)),l) -> + | GApp ({ CAst.v = GRef (g,_)},l) -> looked_for g; RCPatCstr (g, List.map (rcp_of_glob looked_for) l,[]) | _ -> raise Not_found - ) gt + ) let interp_prim_token_cases_pattern_expr ?loc looked_for p = interp_prim_token_gen (rcp_of_glob looked_for) ?loc p @@ -522,8 +522,8 @@ let uninterp_prim_token_ind_pattern ind args = if not b then raise Notation_ops.No_match; let args' = List.map (fun x -> snd (glob_constr_of_closed_cases_pattern x)) args in - let ref = Loc.tag @@ GRef (ref,None) in - match numpr (Loc.tag @@ GApp (ref,args')) with + let ref = CAst.make @@ GRef (ref,None) in + match numpr (CAst.make @@ GApp (ref,args')) with | None -> raise Notation_ops.No_match | Some n -> (sc,n) with Not_found -> raise Notation_ops.No_match diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml index 328fdd519..dd3043803 100644 --- a/interp/notation_ops.ml +++ b/interp/notation_ops.ml @@ -24,7 +24,7 @@ open Notation_term let on_true_do b f c = if b then (f c; b) else b -let compare_glob_constr f add (_l1, t1) (_l2, t2) = match t1,t2 with +let compare_glob_constr f add t1 t2 = match CAst.(t1.v,t2.v) with | GRef (r1,_), GRef (r2,_) -> eq_gr r1 r2 | GVar v1, GVar v2 -> on_true_do (Id.equal v1 v2) add (Name v1) | GApp (f1,l1), GApp (f2,l2) -> f f1 f2 && List.for_all2eq f l1 l2 @@ -117,43 +117,43 @@ let name_to_ident = function let to_id g e id = let e,na = g e (Name id) in e,name_to_ident na -let rec cases_pattern_fold_map ?loc g e = Loc.with_unloc (function +let rec cases_pattern_fold_map ?loc g e = CAst.with_val (function | PatVar na -> - let e',na' = g e na in e', Loc.tag ?loc @@ PatVar na' + let e',na' = g e na in e', CAst.make ?loc @@ PatVar na' | PatCstr (cstr,patl,na) -> let e',na' = g e na in let e',patl' = List.fold_map (cases_pattern_fold_map ?loc g) e patl in - e', Loc.tag ?loc @@ PatCstr (cstr,patl',na') + e', CAst.make ?loc @@ PatCstr (cstr,patl',na') ) let subst_binder_type_vars l = function | Evar_kinds.BinderType (Name id) -> let id = - try match snd @@ Id.List.assoc id l with GVar id' -> id' | _ -> id + try match Id.List.assoc id l with { CAst.v = GVar id' } -> id' | _ -> id with Not_found -> id in Evar_kinds.BinderType (Name id) | e -> e -let rec subst_glob_vars l gc = Loc.map (function - | GVar id as r -> (try snd @@ Id.List.assoc id l with Not_found -> r) +let rec subst_glob_vars l gc = CAst.map (function + | GVar id as r -> (try (Id.List.assoc id l).CAst.v with Not_found -> r) | GProd (Name id,bk,t,c) -> let id = - try match snd @@ Id.List.assoc id l with GVar id' -> id' | _ -> id + try match Id.List.assoc id l with { CAst.v = GVar id' } -> id' | _ -> id with Not_found -> id in GProd (Name id,bk,subst_glob_vars l t,subst_glob_vars l c) | GLambda (Name id,bk,t,c) -> let id = - try match snd @@ Id.List.assoc id l with GVar id' -> id' | _ -> id + try match Id.List.assoc id l with { CAst.v = GVar id' } -> id' | _ -> id with Not_found -> id in GLambda (Name id,bk,subst_glob_vars l t,subst_glob_vars l c) | GHole (x,naming,arg) -> GHole (subst_binder_type_vars l x,naming,arg) - | _ -> snd @@ map_glob_constr (subst_glob_vars l) gc (* assume: id is not binding *) + | _ -> (map_glob_constr (subst_glob_vars l) gc).CAst.v (* assume: id is not binding *) ) gc let ldots_var = Id.of_string ".." let glob_constr_of_notation_constr_with_binders ?loc g f e nc = - let lt x = Loc.tag ?loc x in lt @@ match nc with + let lt x = CAst.make ?loc x in lt @@ match nc with | NVar id -> GVar id | NApp (a,args) -> GApp (f e a, List.map (f e) args) | NList (x,y,iter,tail,swap) -> @@ -161,13 +161,13 @@ let glob_constr_of_notation_constr_with_binders ?loc g f e nc = let innerl = (ldots_var,t)::(if swap then [] else [x, lt @@ GVar y]) in let inner = lt @@ GApp (lt @@ GVar (ldots_var),[subst_glob_vars innerl it]) in let outerl = (ldots_var,inner)::(if swap then [x, lt @@ GVar y] else []) in - Loc.obj @@ subst_glob_vars outerl it + (subst_glob_vars outerl it).CAst.v | NBinderList (x,y,iter,tail) -> let t = f e tail in let it = f e iter in let innerl = [(ldots_var,t);(x, lt @@ GVar y)] in let inner = lt @@ GApp (lt @@ GVar ldots_var,[subst_glob_vars innerl it]) in let outerl = [(ldots_var,inner)] in - Loc.obj @@ subst_glob_vars outerl it + (subst_glob_vars outerl it).CAst.v | NLambda (na,ty,c) -> let e',na = g e na in GLambda (na,Explicit,f e ty,f e' c) | NProd (na,ty,c) -> @@ -188,7 +188,7 @@ let glob_constr_of_notation_constr_with_binders ?loc g f e nc = let eqnl' = List.map (fun (patl,rhs) -> let ((idl,e),patl) = List.fold_map (cases_pattern_fold_map ?loc fold) ([],e) patl in - lt (idl,patl,f e rhs)) eqnl in + Loc.tag (idl,patl,f e rhs)) eqnl in GCases (sty,Option.map (f e') rtntypopt,tml',eqnl') | NLetTuple (nal,(na,po),b,c) -> let e',nal = List.fold_map g e nal in @@ -221,14 +221,15 @@ let add_name r = function Anonymous -> () | Name id -> add_id r id let split_at_recursive_part c = let sub = ref None in + let open CAst in let rec aux = function - | loc0, GApp ((loc,GVar v),c::l) when Id.equal v ldots_var -> + | { loc = loc0; v = GApp ({ loc; v = GVar v },c::l) } when Id.equal v ldots_var -> (* *) begin match !sub with | None -> let () = sub := Some c in begin match l with - | [] -> Loc.tag ?loc @@ GVar ldots_var - | _ :: _ -> Loc.tag ?loc:loc0 @@ GApp (Loc.tag ?loc @@ GVar ldots_var, l) + | [] -> CAst.make ?loc @@ GVar ldots_var + | _ :: _ -> CAst.make ?loc:loc0 @@ GApp (CAst.make ?loc @@ GVar ldots_var, l) end | Some _ -> (* Not narrowed enough to find only one recursive part *) @@ -239,7 +240,7 @@ let split_at_recursive_part c = match !sub with | None -> (* No recursive pattern found *) raise Not_found | Some c -> - match Loc.obj outer_iterator with + match outer_iterator.v with | GVar v when Id.equal v ldots_var -> (* Not enough context *) raise Not_found | _ -> outer_iterator, c @@ -248,7 +249,7 @@ let subtract_loc loc1 loc2 = let l2 = fst (Option.cata Loc.unloc (0,0) loc2) in Some (Loc.make_loc (l1,l2-1)) -let check_is_hole id = function _, GHole _ -> () | t -> +let check_is_hole id = function { CAst.v = GHole _ } -> () | t -> user_err ?loc:(loc_of_glob_constr t) (strbrk "In recursive notation with binders, " ++ pr_id id ++ strbrk " is expected to come without type.") @@ -260,15 +261,16 @@ type recursive_pattern_kind = | RecursiveBinders of glob_constr * glob_constr let compare_recursive_parts found f f' (iterator,subc) = + let open CAst in let diff = ref None in let terminator = ref None in - let rec aux (l1, c1) (l2, c2) = match c1, c2 with + let rec aux c1 c2 = match c1.v, c2.v with | GVar v, term when Id.equal v ldots_var -> (* We found the pattern *) assert (match !terminator with None -> true | Some _ -> false); - terminator := Some (l2, term); + terminator := Some c2; true - | GApp ((_, GVar v),l1), GApp (term, l2) when Id.equal v ldots_var -> + | GApp ({ v = GVar v },l1), GApp (term, l2) when Id.equal v ldots_var -> (* We found the pattern, but there are extra arguments *) (* (this allows e.g. alternative (recursive) notation of application) *) assert (match !terminator with None -> true | Some _ -> false); @@ -294,7 +296,7 @@ let compare_recursive_parts found f f' (iterator,subc) = | Some _ -> false end | _ -> - compare_glob_constr aux (add_name found) (l1, c1) (l2, c2) in + compare_glob_constr aux (add_name found) c1 c2 in if aux iterator subc then match !diff with | None -> @@ -317,13 +319,13 @@ let compare_recursive_parts found f f' (iterator,subc) = (pi1 !found, (x,y) :: pi2 !found, pi3 !found),x,y,lassoc in let iterator = f' (if lassoc then iterator - else subst_glob_vars [x, Loc.tag @@ GVar y] iterator) in + else subst_glob_vars [x, CAst.make @@ GVar y] iterator) in (* found have been collected by compare_constr *) found := newfound; NList (x,y,iterator,f (Option.get !terminator),lassoc) | Some (x,y,RecursiveBinders (t_x,t_y)) -> let newfound = (pi1 !found, pi2 !found, (x,y) :: pi3 !found) in - let iterator = f' (subst_glob_vars [x, Loc.tag @@ GVar y] iterator) in + let iterator = f' (subst_glob_vars [x, CAst.make @@ GVar y] iterator) in (* found have been collected by compare_constr *) found := newfound; check_is_hole x t_x; @@ -341,15 +343,15 @@ let notation_constr_and_vars_of_glob_constr a = try compare_recursive_parts found aux aux' (split_at_recursive_part c) with Not_found -> found := keepfound; - match snd c with - | GApp ((loc, GVar f),[c]) when Id.equal f ldots_var -> + match c.CAst.v with + | GApp ({ CAst.v = GVar f; loc},[c]) when Id.equal f ldots_var -> (* Fall on the second part of the recursive pattern w/o having found the first part *) user_err ?loc (str "Cannot find where the recursive pattern starts.") | _c -> aux' c - and aux' x = Loc.with_unloc (function + and aux' x = CAst.with_val (function | GVar id -> add_id found id; NVar id | GApp (g,args) -> NApp (aux g, List.map aux args) | GLambda (na,bk,ty,c) -> add_name found na; NLambda (na,aux ty,aux c) @@ -456,15 +458,14 @@ let notation_constr_of_constr avoiding t = } in notation_constr_of_glob_constr nenv t -let rec subst_pat subst (loc, pat) = - match pat with - | PatVar _ -> (loc, pat) +let rec subst_pat subst pat = + match pat.CAst.v with + | PatVar _ -> pat | PatCstr (((kn,i),j),cpl,n) -> let kn' = subst_mind subst kn and cpl' = List.smartmap (subst_pat subst) cpl in - Loc.tag ?loc @@ - if kn' == kn && cpl' == cpl then pat else - PatCstr (((kn',i),j),cpl',n) + if kn' == kn && cpl' == cpl then pat else + CAst.make ?loc:pat.CAst.loc @@ PatCstr (((kn',i),j),cpl',n) let rec subst_notation_constr subst bound raw = match raw with @@ -595,8 +596,8 @@ let abstract_return_type_context pi mklam tml rtno = let abstract_return_type_context_glob_constr = abstract_return_type_context (fun (_,(_,nal)) -> nal) - (fun na c -> Loc.tag @@ - GLambda(na,Explicit,Loc.tag @@ GHole(Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None),c)) + (fun na c -> CAst.make @@ + GLambda(na,Explicit,CAst.make @@ GHole(Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None),c)) let abstract_return_type_context_notation_constr = abstract_return_type_context snd @@ -668,9 +669,9 @@ let add_binding_env alp (terms,onlybinders,termlists,binderlists) var v = let add_bindinglist_env (terms,onlybinders,termlists,binderlists) x bl = (terms,onlybinders,termlists,(x,bl)::binderlists) -let rec pat_binder_of_term t = Loc.map (function +let rec pat_binder_of_term t = CAst.map (function | GVar id -> PatVar (Name id) - | GApp ((_, GRef (ConstructRef cstr,_)), l) -> + | GApp ({ CAst.v = GRef (ConstructRef cstr,_)}, l) -> let nparams = Inductiveops.inductive_nparams (fst cstr) in let _,l = List.chop nparams l in PatCstr (cstr, List.map pat_binder_of_term l, Anonymous) @@ -680,7 +681,7 @@ let rec pat_binder_of_term t = Loc.map (function let bind_term_env alp (terms,onlybinders,termlists,binderlists as sigma) var v = try let v' = Id.List.assoc var terms in - match Loc.obj v, Loc.obj v' with + match CAst.(v.v, v'.v) with | GHole _, _ -> sigma | _, GHole _ -> let sigma = Id.List.remove_assoc var terms,onlybinders,termlists,binderlists in @@ -694,7 +695,7 @@ let bind_termlist_env alp (terms,onlybinders,termlists,binderlists as sigma) var try let vl' = Id.List.assoc var termlists in let unify_term v v' = - match Loc.obj v, Loc.obj v' with + match CAst.(v.v, v'.v) with | GHole _, _ -> v' | _, GHole _ -> v | _, _ -> if glob_constr_eq (alpha_rename (snd alp) v) v' then v' else raise No_match in @@ -710,8 +711,8 @@ let bind_termlist_env alp (terms,onlybinders,termlists,binderlists as sigma) var let bind_term_as_binding_env alp (terms,onlybinders,termlists,binderlists as sigma) var id = try - match Loc.obj @@ Id.List.assoc var terms with - | GVar id' -> + match Id.List.assoc var terms with + | { CAst.v = GVar id' } -> (if not (Id.equal id id') then (fst alp,(id,id')::snd alp) else alp), sigma | _ -> anomaly (str "A term which can be a binder has to be a variable") @@ -719,7 +720,7 @@ let bind_term_as_binding_env alp (terms,onlybinders,termlists,binderlists as sig (* The matching against a term allowing to find the instance has not been found yet *) (* If it will be a different name, we shall unfortunately fail *) (* TODO: look at the consequences for alp *) - alp, add_env alp sigma var (Loc.tag @@ GVar id) + alp, add_env alp sigma var (CAst.make @@ GVar id) let bind_binding_as_term_env alp (terms,onlybinders,termlists,binderlists as sigma) var id = try @@ -746,17 +747,17 @@ let bind_binding_env alp (terms,onlybinders,termlists,binderlists as sigma) var else (fst alp,(id1,id2)::snd alp),sigma with Not_found -> alp, add_binding_env alp sigma var v -let rec map_cases_pattern_name_left f = Loc.map (function +let rec map_cases_pattern_name_left f = CAst.map (function | PatVar na -> PatVar (f na) | PatCstr (c,l,na) -> PatCstr (c,List.map_left (map_cases_pattern_name_left f) l,f na) ) -let rec fold_cases_pattern_eq f x p p' = match p, p' with - | (loc, PatVar na), (_, PatVar na') -> let x,na = f x na na' in x, Loc.tag ?loc @@ PatVar na - | (loc, PatCstr (c,l,na)), (_, PatCstr (c',l',na')) when eq_constructor c c' -> +let rec fold_cases_pattern_eq f x p p' = let open CAst in match p, p' with + | { loc; v = PatVar na}, { v = PatVar na' } -> let x,na = f x na na' in x, CAst.make ?loc @@ PatVar na + | { loc; v = PatCstr (c,l,na)}, { v = PatCstr (c',l',na') } when eq_constructor c c' -> let x,l = fold_cases_pattern_list_eq f x l l' in let x,na = f x na na' in - x, Loc.tag ?loc @@ PatCstr (c,l,na) + x, CAst.make ?loc @@ PatCstr (c,l,na) | _ -> failwith "Not equal" and fold_cases_pattern_list_eq f x pl pl' = match pl, pl' with @@ -767,7 +768,7 @@ and fold_cases_pattern_list_eq f x pl pl' = match pl, pl' with x, p :: pl | _ -> assert false -let rec cases_pattern_eq (_,p1) (_,p2) = match p1, p2 with +let rec cases_pattern_eq p1 p2 = match CAst.(p1.v, p2.v) with | PatVar na1, PatVar na2 -> Name.equal na1 na2 | PatCstr (c1, pl1, na1), PatCstr (c2, pl2, na2) -> eq_constructor c1 c2 && List.equal cases_pattern_eq pl1 pl2 && @@ -788,7 +789,7 @@ let bind_bindinglist_env alp (terms,onlybinders,termlists,binderlists as sigma) let unify_pat alp p p' = try fold_cases_pattern_eq unify_name alp p p' with Failure _ -> raise No_match in let unify_term alp v v' = - match Loc.obj v, Loc.obj v' with + match CAst.(v.v, v'.v) with | GHole _, _ -> v' | _, GHole _ -> v | _, _ -> if glob_constr_eq (alpha_rename (snd alp) v) v' then v else raise No_match in @@ -798,17 +799,18 @@ let bind_bindinglist_env alp (terms,onlybinders,termlists,binderlists as sigma) | (Some _ as x), None | None, (Some _ as x) -> x | None, None -> None in let unify_binding_kind bk bk' = if bk == bk' then bk' else raise No_match in - let unify_binder alp (loc, b) (loc', b') = - match b, b' with + let unify_binder alp b b' = + let loc, loc' = CAst.(b.loc, b'.loc) in + match CAst.(b.v, b'.v) with | GLocalAssum (na,bk,t), GLocalAssum (na',bk',t') -> let alp, na = unify_name alp na na' in - alp, Loc.tag ?loc @@ GLocalAssum (na, unify_binding_kind bk bk', unify_term alp t t') + alp, CAst.make ?loc @@ GLocalAssum (na, unify_binding_kind bk bk', unify_term alp t t') | GLocalDef (na,bk,c,t), GLocalDef (na',bk',c',t') -> let alp, na = unify_name alp na na' in - alp, Loc.tag ?loc @@ GLocalDef (na, unify_binding_kind bk bk', unify_term alp c c', unify_opt_term alp t t') + alp, CAst.make ?loc @@ GLocalDef (na, unify_binding_kind bk bk', unify_term alp c c', unify_opt_term alp t t') | GLocalPattern ((p,ids),id,bk,t), GLocalPattern ((p',_),_,bk',t') -> let alp, p = unify_pat alp p p' in - alp, Loc.tag ?loc @@ GLocalPattern ((p,ids), id, unify_binding_kind bk bk', unify_term alp t t') + alp, CAst.make ?loc @@ GLocalPattern ((p,ids), id, unify_binding_kind bk bk', unify_term alp t t') | _ -> raise No_match in let rec unify alp bl bl' = match bl, bl' with @@ -835,18 +837,18 @@ let bind_bindinglist_as_term_env alp (terms,onlybinders,termlists,binderlists) v let unify_pat p p' = if cases_pattern_eq (map_cases_pattern_name_left (name_app (rename_var (snd alp))) p) p' then p' else raise No_match in - let unify_term_binder c (loc, b') = Loc.tag ?loc @@ + let unify_term_binder c = CAst.(map (fun b' -> match c, b' with - | (_, GVar id), GLocalAssum (na', bk', t') -> + | { v = GVar id}, GLocalAssum (na', bk', t') -> GLocalAssum (unify_id id na', bk', t') | c, GLocalPattern ((p',ids), id, bk', t') -> let p = pat_binder_of_term c in GLocalPattern ((unify_pat p p',ids), id, bk', t') - | _ -> raise No_match in + | _ -> raise No_match )) in let rec unify cl bl' = match cl, bl' with | [], [] -> [] - | c :: cl, (_loc, GLocalDef ( _, _, _, t)) :: bl' -> unify cl bl' + | c :: cl, { CAst.v = GLocalDef ( _, _, _, t) } :: bl' -> unify cl bl' | c :: cl, b' :: bl' -> unify_term_binder c b' :: unify cl bl' | _ -> raise No_match in let bl = unify cl bl' in @@ -887,8 +889,8 @@ let match_names metas (alp,sigma) na1 na2 = match (na1,na2) with | (Anonymous,Anonymous) -> alp,sigma | _ -> raise No_match -let rec match_cases_pattern_binders metas acc (_, pat1) (_, pat2) = - match pat1, pat2 with +let rec match_cases_pattern_binders metas acc pat1 pat2 = + match CAst.(pat1.v, pat2.v) with | PatVar na1, PatVar na2 -> match_names metas acc na1 na2 | PatCstr (c1,patl1,na1), PatCstr (c2,patl2,na2) when eq_constructor c1 c2 && Int.equal (List.length patl1) (List.length patl2) -> @@ -898,22 +900,22 @@ let rec match_cases_pattern_binders metas acc (_, pat1) (_, pat2) = let glue_letin_with_decls = true -let rec match_iterated_binders islambda decls bi = Loc.with_loc (fun ?loc -> function - | GLambda (Name p,bk,t,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b))]))) +let rec match_iterated_binders islambda decls bi = CAst.(with_loc_val (fun ?loc -> function + | GLambda (Name p,bk,t, { v = GCases (LetPatternStyle,None,[({ v = GVar e },_)],[(_,(ids,[cp],b))])}) when islambda && Id.equal p e -> - match_iterated_binders islambda ((Loc.tag ?loc @@ GLocalPattern((cp,ids),p,bk,t))::decls) b + match_iterated_binders islambda ((CAst.make ?loc @@ GLocalPattern((cp,ids),p,bk,t))::decls) b | GLambda (na,bk,t,b) when islambda -> - match_iterated_binders islambda ((Loc.tag ?loc @@ GLocalAssum(na,bk,t))::decls) b - | GProd (Name p,bk,t,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b))]))) + match_iterated_binders islambda ((CAst.make ?loc @@ GLocalAssum(na,bk,t))::decls) b + | GProd (Name p,bk,t, { v = GCases (LetPatternStyle,None,[({ v = GVar e },_)],[(_,(ids,[cp],b))]) } ) when not islambda && Id.equal p e -> - match_iterated_binders islambda ((Loc.tag ?loc @@ GLocalPattern((cp,ids),p,bk,t))::decls) b + match_iterated_binders islambda ((CAst.make ?loc @@ GLocalPattern((cp,ids),p,bk,t))::decls) b | GProd ((Name _ as na),bk,t,b) when not islambda -> - match_iterated_binders islambda ((Loc.tag ?loc @@ GLocalAssum(na,bk,t))::decls) b + match_iterated_binders islambda ((CAst.make ?loc @@ GLocalAssum(na,bk,t))::decls) b | GLetIn (na,c,t,b) when glue_letin_with_decls -> match_iterated_binders islambda - ((Loc.tag ?loc @@ GLocalDef (na,Explicit (*?*), c,t))::decls) b - | b -> (decls, Loc.tag ?loc b) - ) bi + ((CAst.make ?loc @@ GLocalDef (na,Explicit (*?*), c,t))::decls) b + | b -> (decls, CAst.make ?loc b) + )) bi let remove_sigma x (terms,onlybinders,termlists,binderlists) = (Id.List.remove_assoc x terms,onlybinders,termlists,binderlists) @@ -974,12 +976,12 @@ let does_not_come_from_already_eta_expanded_var = (* The following test is then an approximation of what can be done *) (* optimally (whether other looping situations can occur remains to be *) (* checked). *) - function _loc, GVar _ -> false | _ -> true + function { CAst.v = GVar _ } -> false | _ -> true let rec match_ inner u alp metas sigma a1 a2 = - let loc, a1_val = Loc.to_pair a1 in - match a1_val, a2 with - + let open CAst in + let loc = a1.loc in + match a1.v, a2 with (* Matching notation variable *) | r1, NVar id2 when is_term_meta id2 metas -> bind_term_env alp sigma id2 a1 | GVar id1, NVar id2 when is_onlybinding_meta id2 metas -> bind_binding_as_term_env alp sigma id2 id1 @@ -990,29 +992,29 @@ let rec match_ inner u alp metas sigma a1 a2 = match_termlist (match_hd u alp) alp metas sigma a1 x y iter termin lassoc (* "λ p, let 'cp = p in t" -> "λ 'cp, t" *) - | GLambda (Name p,bk,t1,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b1))]))), + | GLambda (Name p,bk,t1, { v = GCases (LetPatternStyle,None,[({ v = GVar e},_)],[(_,(ids,[cp],b1))])}), NBinderList (x,_,NLambda (Name _id2,_,b2),termin) when Id.equal p e -> - let (decls,b) = match_iterated_binders true [Loc.tag ?loc @@ GLocalPattern((cp,ids),p,bk,t1)] b1 in + let (decls,b) = match_iterated_binders true [CAst.make ?loc @@ GLocalPattern((cp,ids),p,bk,t1)] b1 in let alp,sigma = bind_bindinglist_env alp sigma x decls in match_in u alp metas sigma b termin (* Matching recursive notations for binders: ad hoc cases supporting let-in *) | GLambda (na1,bk,t1,b1), NBinderList (x,_,NLambda (Name _id2,_,b2),termin)-> - let (decls,b) = match_iterated_binders true [Loc.tag ?loc @@ GLocalAssum (na1,bk,t1)] b1 in + let (decls,b) = match_iterated_binders true [CAst.make ?loc @@ GLocalAssum (na1,bk,t1)] b1 in (* TODO: address the possibility that termin is a Lambda itself *) let alp,sigma = bind_bindinglist_env alp sigma x decls in match_in u alp metas sigma b termin (* "∀ p, let 'cp = p in t" -> "∀ 'cp, t" *) - | GProd (Name p,bk,t1,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b1))]))), + | GProd (Name p,bk,t1, { v = GCases (LetPatternStyle,None,[({ v = GVar e },_)],[(_,(ids,[cp],b1))]) } ), NBinderList (x,_,NProd (Name _id2,_,b2),(NVar v as termin)) when Id.equal p e -> - let (decls,b) = match_iterated_binders true [Loc.tag ?loc @@ GLocalPattern ((cp,ids),p,bk,t1)] b1 in + let (decls,b) = match_iterated_binders true [CAst.make ?loc @@ GLocalPattern ((cp,ids),p,bk,t1)] b1 in let alp,sigma = bind_bindinglist_env alp sigma x decls in match_in u alp metas sigma b termin | GProd (na1,bk,t1,b1), NBinderList (x,_,NProd (Name _id2,_,b2),termin) when na1 != Anonymous -> - let (decls,b) = match_iterated_binders false [Loc.tag ?loc @@ GLocalAssum (na1,bk,t1)] b1 in + let (decls,b) = match_iterated_binders false [CAst.make ?loc @@ GLocalAssum (na1,bk,t1)] b1 in (* TODO: address the possibility that termin is a Prod itself *) let alp,sigma = bind_bindinglist_env alp sigma x decls in match_in u alp metas sigma b termin @@ -1021,18 +1023,18 @@ let rec match_ inner u alp metas sigma a1 a2 = match_binderlist_with_app (match_hd u) alp metas sigma a1 x y iter termin (* Matching individual binders as part of a recursive pattern *) - | GLambda (Name p,bk,t,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b1))]))), + | GLambda (Name p,bk,t, { v = GCases (LetPatternStyle,None,[({ v = GVar e },_)],[(_,(ids,[cp],b1))])}), NLambda (Name id,_,b2) when is_bindinglist_meta id metas -> - let alp,sigma = bind_bindinglist_env alp sigma id [Loc.tag ?loc @@ GLocalPattern ((cp,ids),p,bk,t)] in + let alp,sigma = bind_bindinglist_env alp sigma id [CAst.make ?loc @@ GLocalPattern ((cp,ids),p,bk,t)] in match_in u alp metas sigma b1 b2 | GLambda (na,bk,t,b1), NLambda (Name id,_,b2) when is_bindinglist_meta id metas -> - let alp,sigma = bind_bindinglist_env alp sigma id [Loc.tag ?loc @@ GLocalAssum (na,bk,t)] in + let alp,sigma = bind_bindinglist_env alp sigma id [CAst.make ?loc @@ GLocalAssum (na,bk,t)] in match_in u alp metas sigma b1 b2 | GProd (na,bk,t,b1), NProd (Name id,_,b2) when is_bindinglist_meta id metas && na != Anonymous -> - let alp,sigma = bind_bindinglist_env alp sigma id [Loc.tag ?loc @@ GLocalAssum (na,bk,t)] in + let alp,sigma = bind_bindinglist_env alp sigma id [CAst.make ?loc @@ GLocalAssum (na,bk,t)] in match_in u alp metas sigma b1 b2 (* Matching compositionally *) @@ -1044,7 +1046,7 @@ let rec match_ inner u alp metas sigma a1 a2 = if n1 < n2 then let l21,l22 = List.chop (n2-n1) l2 in f1,l1, NApp (f2,l21), l22 else if n1 > n2 then - let l11,l12 = List.chop (n1-n2) l1 in Loc.tag ?loc @@ GApp (f1,l11),l12, f2,l2 + let l11,l12 = List.chop (n1-n2) l1 in CAst.make ?loc @@ GApp (f1,l11),l12, f2,l2 else f1,l1, f2, l2 in let may_use_eta = does_not_come_from_already_eta_expanded_var f1 in List.fold_left2 (match_ may_use_eta u alp metas) @@ -1117,17 +1119,17 @@ let rec match_ inner u alp metas sigma a1 a2 = let avoid = free_glob_vars a1 @ (* as in Namegen: *) glob_visible_short_qualid a1 in let id' = Namegen.next_ident_away id avoid in - let t1 = Loc.tag @@ GHole(Evar_kinds.BinderType (Name id'),Misctypes.IntroAnonymous,None) in + let t1 = CAst.make @@ GHole(Evar_kinds.BinderType (Name id'),Misctypes.IntroAnonymous,None) in let sigma = match t2 with | NHole _ -> sigma | NVar id2 -> bind_term_env alp sigma id2 t1 | _ -> assert false in let (alp,sigma) = if is_bindinglist_meta id metas then - bind_bindinglist_env alp sigma id [Loc.tag @@ GLocalAssum (Name id',Explicit,t1)] + bind_bindinglist_env alp sigma id [CAst.make @@ GLocalAssum (Name id',Explicit,t1)] else match_names metas (alp,sigma) (Name id') na in - match_in u alp metas sigma (mkGApp a1 (Loc.tag @@ GVar id')) b2 + match_in u alp metas sigma (mkGApp a1 (CAst.make @@ GVar id')) b2 | (GRec _ | GEvar _), _ | _,_ -> raise No_match @@ -1148,7 +1150,7 @@ and match_equations u alp metas sigma (_,(_,patl1,rhs1)) (patl2,rhs2) = (alp,sigma) patl1 patl2 in match_in u alp metas sigma rhs1 rhs2 -let term_of_binder bi = Loc.tag @@ match bi with +let term_of_binder bi = CAst.make @@ match bi with | Name id -> GVar id | Anonymous -> GHole (Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None) @@ -1165,7 +1167,7 @@ let match_notation_constr u c (metas,pat) = with Not_found -> (* Happens for binders bound to Anonymous *) (* Find a better way to propagate Anonymous... *) - Loc.tag @@GVar x in + CAst.make @@GVar x in List.fold_right (fun (x,(scl,typ)) (terms',termlists',binders') -> match typ with | NtnTypeConstr -> @@ -1184,7 +1186,7 @@ let match_notation_constr u c (metas,pat) = let add_patterns_for_params ind l = let mib,_ = Global.lookup_inductive ind in let nparams = mib.Declarations.mind_nparams in - Util.List.addn nparams (Loc.tag @@ PatVar Anonymous) l + Util.List.addn nparams (CAst.make @@ PatVar Anonymous) l let bind_env_cases_pattern (terms,x,termlists,y as sigma) var v = try @@ -1208,9 +1210,10 @@ let match_cases_pattern_list match_fun metas sigma rest x y iter termin lassoc = let l,(terms,onlybinders,termlists,binderlists as sigma) = aux sigma [] rest in (terms,onlybinders,(x,if lassoc then l else List.rev l)::termlists, binderlists) -let rec match_cases_pattern metas (terms,(),termlists,() as sigma) (loc, a1) a2 = - match a1, a2 with - | r1, NVar id2 when Id.List.mem_assoc id2 metas -> (bind_env_cases_pattern sigma id2 (loc, r1)),(0,[]) +let rec match_cases_pattern metas (terms,(),termlists,() as sigma) a1 a2 = + let open CAst in + match a1.v, a2 with + | r1, NVar id2 when Id.List.mem_assoc id2 metas -> (bind_env_cases_pattern sigma id2 a1),(0,[]) | PatVar Anonymous, NHole _ -> sigma,(0,[]) | PatCstr ((ind,_ as r1),largs,_), NRef (ConstructRef r2) when eq_constructor r1 r2 -> sigma,(0,add_patterns_for_params (fst r1) largs) @@ -1226,7 +1229,7 @@ let rec match_cases_pattern metas (terms,(),termlists,() as sigma) (loc, a1) a2 (List.fold_left2 (match_cases_pattern_no_more_args metas) sigma l1' l2),(le2,more_args) | r1, NList (x,y,iter,termin,lassoc) -> (match_cases_pattern_list (match_cases_pattern_no_more_args) - metas (terms,(),termlists,()) (loc, r1) x y iter termin lassoc),(0,[]) + metas (terms,(),termlists,()) a1 x y iter termin lassoc),(0,[]) | _ -> raise No_match and match_cases_pattern_no_more_args metas sigma a1 a2 = diff --git a/intf/constrexpr.mli b/intf/constrexpr.mli index 4f1e9d8e6..5366a302e 100644 --- a/intf/constrexpr.mli +++ b/intf/constrexpr.mli @@ -149,4 +149,4 @@ type module_ast_r = | CMident of qualid | CMapply of module_ast * module_ast | CMwith of module_ast * with_declaration_ast -and module_ast = module_ast_r Loc.located +and module_ast = module_ast_r CAst.ast diff --git a/intf/glob_term.mli b/intf/glob_term.mli index 5aa5bdeeb..aefccd518 100644 --- a/intf/glob_term.mli +++ b/intf/glob_term.mli @@ -28,7 +28,7 @@ type cases_pattern_r = | PatVar of Name.t | PatCstr of constructor * cases_pattern list * Name.t (** [PatCstr(p,C,l,x)] = "|'C' 'l' as 'x'" *) -and cases_pattern = cases_pattern_r Loc.located +and cases_pattern = cases_pattern_r CAst.ast (** Representation of an internalized (or in other words globalized) term. *) type glob_constr_r = @@ -53,7 +53,7 @@ type glob_constr_r = | GSort of glob_sort | GHole of Evar_kinds.t * intro_pattern_naming_expr * Genarg.glob_generic_argument option | GCast of glob_constr * glob_constr cast_type -and glob_constr = glob_constr_r Loc.located +and glob_constr = glob_constr_r CAst.ast and glob_decl = Name.t * binding_kind * glob_constr option * glob_constr @@ -83,7 +83,7 @@ type extended_glob_local_binder_r = | GLocalAssum of Name.t * binding_kind * glob_constr | GLocalDef of Name.t * binding_kind * glob_constr * glob_constr option | GLocalPattern of (cases_pattern * Id.t list) * Id.t * binding_kind * glob_constr -and extended_glob_local_binder = extended_glob_local_binder_r Loc.located +and extended_glob_local_binder = extended_glob_local_binder_r CAst.ast (** A globalised term together with a closure representing the value of its free variables. Intended for use when these variables are taken diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4 index 27f879154..7239bc23b 100644 --- a/parsing/g_vernac.ml4 +++ b/parsing/g_vernac.ml4 @@ -511,11 +511,11 @@ GEXTEND Gram (* Module expressions *) module_expr: [ [ me = module_expr_atom -> me - | me1 = module_expr; me2 = module_expr_atom -> Loc.tag ~loc:!@loc @@ CMapply (me1,me2) + | me1 = module_expr; me2 = module_expr_atom -> CAst.make ~loc:!@loc @@ CMapply (me1,me2) ] ] ; module_expr_atom: - [ [ qid = qualid -> Loc.tag ~loc:!@loc @@ CMident (snd qid) | "("; me = module_expr; ")" -> me ] ] + [ [ qid = qualid -> CAst.make ~loc:!@loc @@ CMident (snd qid) | "("; me = module_expr; ")" -> me ] ] ; with_declaration: [ [ "Definition"; fqid = fullyqualid; ":="; c = Constr.lconstr -> @@ -525,12 +525,12 @@ GEXTEND Gram ] ] ; module_type: - [ [ qid = qualid -> Loc.tag ~loc:!@loc @@ CMident (snd qid) + [ [ qid = qualid -> CAst.make ~loc:!@loc @@ CMident (snd qid) | "("; mt = module_type; ")" -> mt | mty = module_type; me = module_expr_atom -> - Loc.tag ~loc:!@loc @@ CMapply (mty,me) + CAst.make ~loc:!@loc @@ CMapply (mty,me) | mty = module_type; "with"; decl = with_declaration -> - Loc.tag ~loc:!@loc @@ CMwith (mty,decl) + CAst.make ~loc:!@loc @@ CMwith (mty,decl) ] ] ; (* Proof using *) diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml index 938fe5237..7f2b21e79 100644 --- a/plugins/funind/glob_term_to_relation.ml +++ b/plugins/funind/glob_term_to_relation.ml @@ -345,10 +345,10 @@ let raw_push_named (na,raw_value,raw_typ) env = let add_pat_variables pat typ env : Environ.env = - let rec add_pat_variables env (loc, pat) typ : Environ.env = + let rec add_pat_variables env pat typ : Environ.env = observe (str "new rel env := " ++ Printer.pr_rel_context_of env (Evd.from_env env)); - match pat with + match pat.CAst.v with | PatVar na -> Environ.push_rel (RelDecl.LocalAssum (na,typ)) env | PatCstr(c,patl,na) -> let Inductiveops.IndType(indf,indargs) = @@ -398,7 +398,7 @@ let add_pat_variables pat typ env : Environ.env = -let rec pattern_to_term_and_type env typ = Loc.with_unloc (function +let rec pattern_to_term_and_type env typ = CAst.with_val (function | PatVar Anonymous -> assert false | PatVar (Name id) -> mkGVar id @@ -466,11 +466,12 @@ let rec pattern_to_term_and_type env typ = Loc.with_unloc (function let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = observe (str " Entering : " ++ Printer.pr_glob_constr rt); - match rt with - | _, GRef _ | _, GVar _ | _, GEvar _ | _, GPatVar _ | _, GSort _ | _, GHole _ -> + let open CAst in + match rt.v with + | GRef _ | GVar _ | GEvar _ | GPatVar _ | GSort _ | GHole _ -> (* do nothing (except changing type of course) *) mk_result [] rt avoid - | _, GApp(_,_) -> + | GApp(_,_) -> let f,args = glob_decompose_app rt in let args_res : (glob_constr list) build_entry_return = List.fold_right (* create the arguments lists of constructors and combine them *) @@ -482,14 +483,14 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = (mk_result [] [] avoid) in begin - match Loc.obj f with + match f.v with | GLambda _ -> let rec aux t l = match l with | [] -> t - | u::l -> Loc.tag @@ - match t with - | loc, GLambda(na,_,nat,b) -> + | u::l -> CAst.make @@ + match t.v with + | GLambda(na,_,nat,b) -> GLetIn(na,u,None,aux b l) | _ -> GApp(t,l) @@ -550,7 +551,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = let new_b = replace_var_by_term id - (Loc.tag @@ GVar id) + (CAst.make @@ GVar id) b in (Name new_id,new_b,new_avoid) @@ -579,7 +580,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = | GProd _ -> error "Cannot apply a type" end (* end of the application treatement *) - | _, GLambda(n,_,t,b) -> + | GLambda(n,_,t,b) -> (* we first compute the list of constructor corresponding to the body of the function, then the one corresponding to the type @@ -594,7 +595,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = let new_env = raw_push_named (new_n,None,t) env in let b_res = build_entry_lc new_env funnames avoid b in combine_results (combine_lam new_n) t_res b_res - | _, GProd(n,_,t,b) -> + | GProd(n,_,t,b) -> (* we first compute the list of constructor corresponding to the body of the function, then the one corresponding to the type @@ -604,13 +605,13 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = let new_env = raw_push_named (n,None,t) env in let b_res = build_entry_lc new_env funnames avoid b in combine_results (combine_prod n) t_res b_res - | loc, GLetIn(n,v,typ,b) -> + | GLetIn(n,v,typ,b) -> (* we first compute the list of constructor corresponding to the body of the function, then the one corresponding to the value [t] and combine the two result *) - let v = match typ with None -> v | Some t -> Loc.tag ?loc @@ GCast (v,CastConv t) in + let v = match typ with None -> v | Some t -> CAst.make ?loc:rt.loc @@ GCast (v,CastConv t) in let v_res = build_entry_lc env funnames avoid v in let v_as_constr,ctx = Pretyping.understand env (Evd.from_env env) v in let v_type = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr v_as_constr) in @@ -622,13 +623,13 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = in let b_res = build_entry_lc new_env funnames avoid b in combine_results (combine_letin n) v_res b_res - | _, GCases(_,_,el,brl) -> + | GCases(_,_,el,brl) -> (* we create the discrimination function and treat the case itself *) let make_discr = make_discr_match brl in build_entry_lc_from_case env funnames make_discr el brl avoid - | _, GIf(b,(na,e_option),lhs,rhs) -> + | GIf(b,(na,e_option),lhs,rhs) -> let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in let b_typ = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr b_as_constr) in let (ind,_) = @@ -651,7 +652,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = in (* Pp.msgnl (str "new case := " ++ Printer.pr_glob_constr match_expr); *) build_entry_lc env funnames avoid match_expr - | _, GLetTuple(nal,_,b,e) -> + | GLetTuple(nal,_,b,e) -> begin let nal_as_glob_constr = List.map @@ -677,8 +678,8 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = build_entry_lc env funnames avoid match_expr end - | _, GRec _ -> error "Not handled GRec" - | _, GCast(b,_) -> + | GRec _ -> error "Not handled GRec" + | GCast(b,_) -> build_entry_lc env funnames avoid b and build_entry_lc_from_case env funname make_discr (el:tomatch_tuples) @@ -860,8 +861,8 @@ let is_res id = -let same_raw_term (_,rt1) (_,rt2) = - match rt1,rt2 with +let same_raw_term rt1 rt2 = + match CAst.(rt1.v, rt2.v) with | GRef(r1,_), GRef (r2,_) -> Globnames.eq_gr r1 r2 | GHole _, GHole _ -> true | _ -> false @@ -894,16 +895,17 @@ exception Continue let rec rebuild_cons env nb_args relname args crossed_types depth rt = observe (str "rebuilding : " ++ pr_glob_constr rt); let open Context.Rel.Declaration in - match rt with - | _, GProd(n,k,t,b) -> + let open CAst in + match rt.v with + | GProd(n,k,t,b) -> let not_free_in_t id = not (is_free_in id t) in let new_crossed_types = t::crossed_types in begin match t with - | _, GApp(((_, GVar(res_id)) as res_rt),args') when is_res res_id -> + | { v = GApp(({ v = GVar res_id } as res_rt),args') } when is_res res_id -> begin match args' with - | (_, (GVar this_relname))::args' -> + | { v = GVar this_relname }::args' -> (*i The next call to mk_rel_id is valid since we are constructing the graph Ensures by: obvious @@ -925,7 +927,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = | _ -> (* the first args is the name of the function! *) assert false end - | loc1, GApp((loc2, GRef(eq_as_ref,_)),[ty;(loc3, GVar id);rt]) + | { loc = loc1; v = GApp({ loc = loc2; v = GRef(eq_as_ref,_) },[ty; { loc = loc3; v = GVar id};rt]) } when Globnames.eq_gr eq_as_ref (Lazy.force Coqlib.coq_eq_ref) && n == Anonymous -> begin @@ -962,8 +964,8 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = let params,arg' = ((Util.List.chop nparam args')) in - let rt_typ = Loc.tag @@ - GApp(Loc.tag @@ GRef (Globnames.IndRef (fst ind),None), + let rt_typ = CAst.make @@ + GApp(CAst.make @@ GRef (Globnames.IndRef (fst ind),None), (List.map (fun p -> Detyping.detype false [] env (Evd.from_env env) @@ -973,7 +975,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = (mkGHole ())))) in let eq' = - Loc.tag ?loc:loc1 @@ GApp(Loc.tag ?loc:loc2 @@GRef(jmeq,None),[ty;Loc.tag ?loc:loc3 @@ GVar id;rt_typ;rt]) + CAst.make ?loc:loc1 @@ GApp(CAst.make ?loc:loc2 @@GRef(jmeq,None),[ty;CAst.make ?loc:loc3 @@ GVar id;rt_typ;rt]) in observe (str "computing new type for jmeq : " ++ pr_glob_constr eq'); let eq'_as_constr,ctx = Pretyping.understand env (Evd.from_env env) eq' in @@ -1042,7 +1044,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = mkGProd(n,t,new_b),id_to_exclude else new_b, Id.Set.add id id_to_exclude *) - | loc1, GApp((loc2, GRef(eq_as_ref,_)),[ty;rt1;rt2]) + | { loc = loc1; v = GApp({ loc = loc2; v = GRef(eq_as_ref,_) },[ty;rt1;rt2]) } when Globnames.eq_gr eq_as_ref (Lazy.force Coqlib.coq_eq_ref) && n == Anonymous -> begin @@ -1093,7 +1095,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = (Id.Set.filter not_free_in_t id_to_exclude) | _ -> mkGProd(n,t,new_b),Id.Set.filter not_free_in_t id_to_exclude end - | _, GLambda(n,k,t,b) -> + | GLambda(n,k,t,b) -> begin let not_free_in_t id = not (is_free_in id t) in let new_crossed_types = t :: crossed_types in @@ -1112,14 +1114,14 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = then new_b, Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude) else - Loc.tag @@ GProd(n,k,t,new_b),Id.Set.filter not_free_in_t id_to_exclude + CAst.make @@ GProd(n,k,t,new_b),Id.Set.filter not_free_in_t id_to_exclude | _ -> anomaly (Pp.str "Should not have an anonymous function here") (* We have renamed all the anonymous functions during alpha_renaming phase *) end - | loc, GLetIn(n,v,t,b) -> + | GLetIn(n,v,t,b) -> begin - let t = match t with None -> v | Some t -> Loc.tag ?loc @@ GCast (v,CastConv t) in + let t = match t with None -> v | Some t -> CAst.make ?loc:rt.loc @@ GCast (v,CastConv t) in let not_free_in_t id = not (is_free_in id t) in let evd = (Evd.from_env env) in let t',ctx = Pretyping.understand env evd t in @@ -1135,10 +1137,10 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = match n with | Name id when Id.Set.mem id id_to_exclude && depth >= nb_args -> new_b,Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude) - | _ -> Loc.tag @@ GLetIn(n,t,None,new_b), (* HOPING IT WOULD WORK *) + | _ -> CAst.make @@ GLetIn(n,t,None,new_b), (* HOPING IT WOULD WORK *) Id.Set.filter not_free_in_t id_to_exclude end - | _, GLetTuple(nal,(na,rto),t,b) -> + | GLetTuple(nal,(na,rto),t,b) -> assert (Option.is_empty rto); begin let not_free_in_t id = not (is_free_in id t) in @@ -1161,7 +1163,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = (* | Name id when Id.Set.mem id id_to_exclude -> *) (* new_b,Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude) *) (* | _ -> *) - Loc.tag @@ GLetTuple(nal,(na,None),t,new_b), + CAst.make @@ GLetTuple(nal,(na,None),t,new_b), Id.Set.filter not_free_in_t (Id.Set.union id_to_exclude id_to_exclude') end @@ -1187,9 +1189,9 @@ let rebuild_cons env nb_args relname args crossed_types rt = TODO: Find a valid way to deal with implicit arguments here! *) -let rec compute_cst_params relnames params gt = Loc.with_unloc (function +let rec compute_cst_params relnames params gt = CAst.with_val (function | GRef _ | GVar _ | GEvar _ | GPatVar _ -> params - | GApp((_, GVar relname'),rtl) when Id.Set.mem relname' relnames -> + | GApp({ CAst.v = GVar relname' },rtl) when Id.Set.mem relname' relnames -> compute_cst_params_from_app [] (params,rtl) | GApp(f,args) -> List.fold_left (compute_cst_params relnames) params (f::args) @@ -1211,7 +1213,7 @@ let rec compute_cst_params relnames params gt = Loc.with_unloc (function and compute_cst_params_from_app acc (params,rtl) = match params,rtl with | _::_,[] -> assert false (* the rel has at least nargs + 1 arguments ! *) - | ((Name id,_,None) as param)::params',(_, GVar id')::rtl' + | ((Name id,_,None) as param)::params', { CAst.v = GVar id' }::rtl' when Id.compare id id' == 0 -> compute_cst_params_from_app (param::acc) (params',rtl') | _ -> List.rev acc diff --git a/plugins/funind/glob_termops.ml b/plugins/funind/glob_termops.ml index 66b9897d0..5abcb100f 100644 --- a/plugins/funind/glob_termops.ml +++ b/plugins/funind/glob_termops.ml @@ -10,16 +10,16 @@ open Misctypes Some basic functions to rebuild glob_constr In each of them the location is Loc.ghost *) -let mkGRef ref = Loc.tag @@ GRef(ref,None) -let mkGVar id = Loc.tag @@ GVar(id) -let mkGApp(rt,rtl) = Loc.tag @@ GApp(rt,rtl) -let mkGLambda(n,t,b) = Loc.tag @@ GLambda(n,Explicit,t,b) -let mkGProd(n,t,b) = Loc.tag @@ GProd(n,Explicit,t,b) -let mkGLetIn(n,b,t,c) = Loc.tag @@ GLetIn(n,b,t,c) -let mkGCases(rto,l,brl) = Loc.tag @@ GCases(Term.RegularStyle,rto,l,brl) -let mkGSort s = Loc.tag @@ GSort(s) -let mkGHole () = Loc.tag @@ GHole(Evar_kinds.BinderType Anonymous,Misctypes.IntroAnonymous,None) -let mkGCast(b,t) = Loc.tag @@ GCast(b,CastConv t) +let mkGRef ref = CAst.make @@ GRef(ref,None) +let mkGVar id = CAst.make @@ GVar(id) +let mkGApp(rt,rtl) = CAst.make @@ GApp(rt,rtl) +let mkGLambda(n,t,b) = CAst.make @@ GLambda(n,Explicit,t,b) +let mkGProd(n,t,b) = CAst.make @@ GProd(n,Explicit,t,b) +let mkGLetIn(n,b,t,c) = CAst.make @@ GLetIn(n,b,t,c) +let mkGCases(rto,l,brl) = CAst.make @@ GCases(Term.RegularStyle,rto,l,brl) +let mkGSort s = CAst.make @@ GSort(s) +let mkGHole () = CAst.make @@ GHole(Evar_kinds.BinderType Anonymous,Misctypes.IntroAnonymous,None) +let mkGCast(b,t) = CAst.make @@ GCast(b,CastConv t) (* Some basic functions to decompose glob_constrs @@ -27,7 +27,7 @@ let mkGCast(b,t) = Loc.tag @@ GCast(b,CastConv t) *) let glob_decompose_prod = let rec glob_decompose_prod args = function - | _, GProd(n,k,t,b) -> + | { CAst.v = GProd(n,k,t,b) } -> glob_decompose_prod ((n,t)::args) b | rt -> args,rt in @@ -35,9 +35,9 @@ let glob_decompose_prod = let glob_decompose_prod_or_letin = let rec glob_decompose_prod args = function - | _, GProd(n,k,t,b) -> + | { CAst.v = GProd(n,k,t,b) } -> glob_decompose_prod ((n,None,Some t)::args) b - | _, GLetIn(n,b,t,c) -> + | { CAst.v = GLetIn(n,b,t,c) } -> glob_decompose_prod ((n,Some b,t)::args) c | rt -> args,rt in @@ -59,7 +59,7 @@ let glob_decompose_prod_n n = if i<=0 then args,c else match c with - | _, GProd(n,_,t,b) -> + | { CAst.v = GProd(n,_,t,b) } -> glob_decompose_prod (i-1) ((n,t)::args) b | rt -> args,rt in @@ -71,9 +71,9 @@ let glob_decompose_prod_or_letin_n n = if i<=0 then args,c else match c with - | _, GProd(n,_,t,b) -> + | { CAst.v = GProd(n,_,t,b) } -> glob_decompose_prod (i-1) ((n,None,Some t)::args) b - | _, GLetIn(n,b,t,c) -> + | { CAst.v = GLetIn(n,b,t,c) } -> glob_decompose_prod (i-1) ((n,Some b,t)::args) c | rt -> args,rt in @@ -84,7 +84,7 @@ let glob_decompose_app = let rec decompose_rapp acc rt = (* msgnl (str "glob_decompose_app on : "++ Printer.pr_glob_constr rt); *) match rt with - | _, GApp(rt,rtl) -> + | { CAst.v = GApp(rt,rtl) } -> decompose_rapp (List.fold_left (fun y x -> x::y) acc rtl) rt | rt -> rt,List.rev acc in @@ -120,7 +120,7 @@ let remove_name_from_mapping mapping na = let change_vars = let rec change_vars mapping rt = - Loc.map (function + CAst.map (function | GRef _ as x -> x | GVar id -> let new_id = @@ -189,18 +189,19 @@ let change_vars = -let rec alpha_pat excluded (loc, pat) = - match pat with +let rec alpha_pat excluded pat = + let loc = pat.CAst.loc in + match pat.CAst.v with | PatVar Anonymous -> let new_id = Indfun_common.fresh_id excluded "_x" in - (Loc.tag ?loc @@ PatVar(Name new_id)),(new_id::excluded),Id.Map.empty + (CAst.make ?loc @@ PatVar(Name new_id)),(new_id::excluded),Id.Map.empty | PatVar(Name id) -> if Id.List.mem id excluded then let new_id = Namegen.next_ident_away id excluded in - (Loc.tag ?loc @@ PatVar(Name new_id)),(new_id::excluded), + (CAst.make ?loc @@ PatVar(Name new_id)),(new_id::excluded), (Id.Map.add id new_id Id.Map.empty) - else (Loc.tag ?loc pat),excluded,Id.Map.empty + else pat, excluded,Id.Map.empty | PatCstr(constr,patl,na) -> let new_na,new_excluded,map = match na with @@ -218,7 +219,7 @@ let rec alpha_pat excluded (loc, pat) = ([],new_excluded,map) patl in - (Loc.tag ?loc @@ PatCstr(constr,List.rev new_patl,new_na)),new_excluded,new_map + (CAst.make ?loc @@ PatCstr(constr,List.rev new_patl,new_na)),new_excluded,new_map let alpha_patl excluded patl = let patl,new_excluded,map = @@ -236,8 +237,8 @@ let alpha_patl excluded patl = let raw_get_pattern_id pat acc = - let rec get_pattern_id (loc, pat) = - match pat with + let rec get_pattern_id pat = + match pat.CAst.v with | PatVar(Anonymous) -> assert false | PatVar(Name id) -> [id] @@ -254,10 +255,11 @@ let raw_get_pattern_id pat acc = let get_pattern_id pat = raw_get_pattern_id pat [] -let rec alpha_rt excluded (loc, rt) = - let new_rt = Loc.tag ?loc @@ - match rt with - | GRef _ | GVar _ | GEvar _ | GPatVar _ -> rt +let rec alpha_rt excluded rt = + let loc = rt.CAst.loc in + let new_rt = CAst.make ?loc @@ + match rt.CAst.v with + | GRef _ | GVar _ | GEvar _ | GPatVar _ as rt -> rt | GLambda(Anonymous,k,t,b) -> let new_id = Namegen.next_ident_away (Id.of_string "_x") excluded in let new_excluded = new_id :: excluded in @@ -351,8 +353,8 @@ let rec alpha_rt excluded (loc, rt) = alpha_rt excluded rhs ) | GRec _ -> error "Not handled GRec" - | GSort _ -> rt - | GHole _ -> rt + | GSort _ + | GHole _ as rt -> rt | GCast (b,c) -> GCast(alpha_rt excluded b, Miscops.map_cast_type (alpha_rt excluded) c) @@ -375,7 +377,7 @@ and alpha_br excluded (loc,(ids,patl,res)) = [is_free_in id rt] checks if [id] is a free variable in [rt] *) let is_free_in id = - let rec is_free_in (loc, gt) = match gt with + let rec is_free_in x = CAst.with_loc_val (fun ?loc -> function | GRef _ -> false | GVar id' -> Id.compare id' id == 0 | GEvar _ -> false @@ -411,6 +413,7 @@ let is_free_in id = | GHole _ -> false | GCast (b,(CastConv t|CastVM t|CastNative t)) -> is_free_in b || is_free_in t | GCast (b,CastCoerce) -> is_free_in b + ) x and is_free_in_br (_,(ids,_,rt)) = (not (Id.List.mem id ids)) && is_free_in rt in @@ -418,7 +421,7 @@ let is_free_in id = -let rec pattern_to_term pt = Loc.with_unloc (function +let rec pattern_to_term pt = CAst.with_val (function | PatVar Anonymous -> assert false | PatVar(Name id) -> mkGVar id @@ -445,39 +448,38 @@ let rec pattern_to_term pt = Loc.with_unloc (function let replace_var_by_term x_id term = - let rec replace_var_by_pattern (loc, rt) = Loc.tag ?loc @@ - match rt with - | GRef _ -> rt - | GVar id when Id.compare id x_id == 0 -> Loc.obj term - | GVar _ -> rt - | GEvar _ -> rt - | GPatVar _ -> rt + let rec replace_var_by_pattern x = CAst.map (function + | GVar id when Id.compare id x_id == 0 -> term.CAst.v + | GRef _ + | GVar _ + | GEvar _ + | GPatVar _ as rt -> rt | GApp(rt',rtl) -> GApp(replace_var_by_pattern rt', List.map replace_var_by_pattern rtl ) - | GLambda(Name id,_,_,_) when Id.compare id x_id == 0 -> rt + | GLambda(Name id,_,_,_) as rt when Id.compare id x_id == 0 -> rt | GLambda(name,k,t,b) -> GLambda(name, k, replace_var_by_pattern t, replace_var_by_pattern b ) - | GProd(Name id,_,_,_) when Id.compare id x_id == 0 -> rt + | GProd(Name id,_,_,_) as rt when Id.compare id x_id == 0 -> rt | GProd(name,k,t,b) -> GProd( name, k, replace_var_by_pattern t, replace_var_by_pattern b ) - | GLetIn(Name id,_,_,_) when Id.compare id x_id == 0 -> rt + | GLetIn(Name id,_,_,_) as rt when Id.compare id x_id == 0 -> rt | GLetIn(name,def,typ,b) -> GLetIn(name, replace_var_by_pattern def, Option.map (replace_var_by_pattern) typ, replace_var_by_pattern b ) - | GLetTuple(nal,_,_,_) + | GLetTuple(nal,_,_,_) as rt when List.exists (function Name id -> Id.equal id x_id | _ -> false) nal -> rt | GLetTuple(nal,(na,rto),def,b) -> @@ -499,11 +501,12 @@ let replace_var_by_term x_id term = replace_var_by_pattern rhs ) | GRec _ -> raise (UserError(None,str "Not handled GRec")) - | GSort _ -> rt - | GHole _ -> rt + | GSort _ + | GHole _ as rt -> rt | GCast(b,c) -> GCast(replace_var_by_pattern b, Miscops.map_cast_type replace_var_by_pattern c) + ) x and replace_var_by_pattern_br ((loc,(idl,patl,res)) as br) = if List.exists (fun id -> Id.compare id x_id == 0) idl then br @@ -520,9 +523,10 @@ exception NotUnifiable let rec are_unifiable_aux = function | [] -> () | eq::eqs -> + let open CAst in match eq with - | (_,PatVar _),_ | _,(_,PatVar _) -> are_unifiable_aux eqs - | (_,PatCstr(constructor1,cpl1,_)),(_,PatCstr(constructor2,cpl2,_)) -> + | { v = PatVar _ },_ | _, { v = PatVar _ } -> are_unifiable_aux eqs + | { v = PatCstr(constructor1,cpl1,_) }, { v = PatCstr(constructor2,cpl2,_) } -> if not (eq_constructor constructor2 constructor1) then raise NotUnifiable else @@ -542,9 +546,10 @@ let are_unifiable pat1 pat2 = let rec eq_cases_pattern_aux = function | [] -> () | eq::eqs -> + let open CAst in match eq with - | (_,PatVar _),(_,PatVar _) -> eq_cases_pattern_aux eqs - | (_,PatCstr(constructor1,cpl1,_)),(_,PatCstr(constructor2,cpl2,_)) -> + | { v = PatVar _ }, { v = PatVar _ } -> eq_cases_pattern_aux eqs + | { v = PatCstr(constructor1,cpl1,_) }, { v = PatCstr(constructor2,cpl2,_) } -> if not (eq_constructor constructor2 constructor1) then raise NotUnifiable else @@ -564,7 +569,7 @@ let eq_cases_pattern pat1 pat2 = let ids_of_pat = - let rec ids_of_pat ids = Loc.with_unloc (function + let rec ids_of_pat ids = CAst.with_val (function | PatVar Anonymous -> ids | PatVar(Name id) -> Id.Set.add id ids | PatCstr(_,patl,_) -> List.fold_left ids_of_pat ids patl @@ -578,7 +583,7 @@ let id_of_name = function (* TODO: finish Rec caes *) let ids_of_glob_constr c = - let rec ids_of_glob_constr acc (loc, c) = + let rec ids_of_glob_constr acc {loc; CAst.v = c} = let idof = id_of_name in match c with | GVar id -> id::acc @@ -605,12 +610,11 @@ let ids_of_glob_constr c = let zeta_normalize = - let rec zeta_normalize_term (loc, rt) = Loc.tag ?loc @@ - match rt with - | GRef _ -> rt - | GVar _ -> rt - | GEvar _ -> rt - | GPatVar _ -> rt + let rec zeta_normalize_term x = CAst.map (function + | GRef _ + | GVar _ + | GEvar _ + | GPatVar _ as rt -> rt | GApp(rt',rtl) -> GApp(zeta_normalize_term rt', List.map zeta_normalize_term rtl @@ -628,9 +632,9 @@ let zeta_normalize = zeta_normalize_term b ) | GLetIn(Name id,def,typ,b) -> - Loc.obj @@ zeta_normalize_term (replace_var_by_term id def b) + (zeta_normalize_term (replace_var_by_term id def b)).CAst.v | GLetIn(Anonymous,def,typ,b) -> - Loc.obj @@ zeta_normalize_term b + (zeta_normalize_term b).CAst.v | GLetTuple(nal,(na,rto),def,b) -> GLetTuple(nal, (na,Option.map zeta_normalize_term rto), @@ -650,11 +654,12 @@ let zeta_normalize = zeta_normalize_term rhs ) | GRec _ -> raise (UserError(None,str "Not handled GRec")) - | GSort _ -> rt - | GHole _ -> rt + | GSort _ + | GHole _ as rt -> rt | GCast(b,c) -> GCast(zeta_normalize_term b, Miscops.map_cast_type zeta_normalize_term c) + ) x and zeta_normalize_br (loc,(idl,patl,res)) = (loc,(idl,patl,zeta_normalize_term res)) in @@ -665,21 +670,19 @@ let zeta_normalize = let expand_as = - let rec add_as map (loc, pat) = + let rec add_as map ({loc; CAst.v = pat } as rt) = match pat with | PatVar _ -> map | PatCstr(_,patl,Name id) -> - Id.Map.add id (pattern_to_term (loc, pat)) (List.fold_left add_as map patl) + Id.Map.add id (pattern_to_term rt) (List.fold_left add_as map patl) | PatCstr(_,patl,_) -> List.fold_left add_as map patl in - let rec expand_as map (loc, rt) = - Loc.tag ?loc @@ - match rt with - | GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ -> rt - | GVar id -> + let rec expand_as map = CAst.map (function + | GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ as rt -> rt + | GVar id as rt -> begin try - Loc.obj @@ Id.Map.find id map + (Id.Map.find id map).CAst.v with Not_found -> rt end | GApp(f,args) -> GApp(expand_as map f,List.map (expand_as map) args) @@ -699,6 +702,7 @@ let expand_as = | GCases(sty,po,el,brl) -> GCases(sty, Option.map (expand_as map) po, List.map (fun (rt,t) -> expand_as map rt,t) el, List.map (expand_as_br map) brl) + ) and expand_as_br map (loc,(idl,cpl,rt)) = (loc,(idl,cpl, expand_as (List.fold_left add_as map cpl) rt)) in diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index f4e9aa372..ab83cb15a 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -190,7 +190,7 @@ let build_newrecursive l = let is_rec names = let names = List.fold_right Id.Set.add names Id.Set.empty in let check_id id names = Id.Set.mem id names in - let rec lookup names (loc, gt) = match gt with + let rec lookup names gt = match gt.CAst.v with | GVar(id) -> check_id id names | GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ -> false | GCast(b,_) -> lookup names b diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml index de8dc53f1..394b252aa 100644 --- a/plugins/funind/indfun_common.ml +++ b/plugins/funind/indfun_common.ml @@ -69,7 +69,7 @@ let chop_rlambda_n = if n == 0 then List.rev acc,rt else - match Loc.obj rt with + match rt.CAst.v with | Glob_term.GLambda(name,k,t,b) -> chop_lambda_n ((name,t,None)::acc) (n-1) b | Glob_term.GLetIn(name,v,t,b) -> chop_lambda_n ((name,v,t)::acc) (n-1) b | _ -> @@ -83,7 +83,7 @@ let chop_rprod_n = if n == 0 then List.rev acc,rt else - match Loc.obj rt with + match rt.CAst.v with | Glob_term.GProd(name,k,t,b) -> chop_prod_n ((name,t)::acc) (n-1) b | _ -> raise (CErrors.UserError(Some "chop_rprod_n",str "chop_rprod_n: Not enough products")) in diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml index 5b51a213a..d4865bf5e 100644 --- a/plugins/funind/merge.ml +++ b/plugins/funind/merge.ml @@ -66,7 +66,7 @@ let string_of_name = id_of_name %> Id.to_string (** [isVarf f x] returns [true] if term [x] is of the form [(Var f)]. *) let isVarf f x = match x with - | _, GVar x -> Id.equal x f + | { CAst.v = GVar x } -> Id.equal x f | _ -> false (** [ident_global_exist id] returns true if identifier [id] is linked @@ -504,40 +504,40 @@ let shift_linked_params mib1 mib2 (lnk1:linked_var array) (lnk2:linked_var array exception NoMerge -let rec merge_app (loc1, c1) (loc2, c2) id1 id2 shift filter_shift_stable = +let rec merge_app c1 c2 id1 id2 shift filter_shift_stable = let lnk = Array.append shift.lnk1 shift.lnk2 in - match c1, c2 with + match CAst.(c1.v, c2.v) with | GApp(f1, arr1), GApp(f2,arr2) when isVarf id1 f1 && isVarf id2 f2 -> let _ = prstr "\nICI1!\n" in let args = filter_shift_stable lnk (arr1 @ arr2) in - Loc.tag @@ GApp ((Loc.tag @@ GVar shift.ident) , args) + CAst.make @@ GApp ((CAst.make @@ GVar shift.ident) , args) | GApp(f1, arr1), GApp(f2,arr2) -> raise NoMerge | GLetIn(nme,bdy,typ,trm) , _ -> let _ = prstr "\nICI2!\n" in - let newtrm = merge_app trm (loc2, c2) id1 id2 shift filter_shift_stable in - Loc.tag @@ GLetIn(nme,bdy,typ,newtrm) + let newtrm = merge_app trm c2 id1 id2 shift filter_shift_stable in + CAst.make @@ GLetIn(nme,bdy,typ,newtrm) | _, GLetIn(nme,bdy,typ,trm) -> let _ = prstr "\nICI3!\n" in - let newtrm = merge_app (loc1, c1) trm id1 id2 shift filter_shift_stable in - Loc.tag @@ GLetIn(nme,bdy,typ,newtrm) + let newtrm = merge_app c1 trm id1 id2 shift filter_shift_stable in + CAst.make @@ GLetIn(nme,bdy,typ,newtrm) | _ -> let _ = prstr "\nICI4!\n" in raise NoMerge -let rec merge_app_unsafe (l1, c1) (l2, c2) shift filter_shift_stable = +let rec merge_app_unsafe c1 c2 shift filter_shift_stable = let lnk = Array.append shift.lnk1 shift.lnk2 in - match c1 , c2 with + match CAst.(c1.v, c2.v) with | GApp(f1, arr1), GApp(f2,arr2) -> let args = filter_shift_stable lnk (arr1 @ arr2) in - Loc.tag @@ GApp (Loc.tag @@ GVar shift.ident, args) + CAst.make @@ GApp (CAst.make @@ GVar shift.ident, args) (* FIXME: what if the function appears in the body of the let? *) | GLetIn(nme,bdy,typ,trm) , _ -> let _ = prstr "\nICI2 '!\n" in - let newtrm = merge_app_unsafe trm (l2, c2) shift filter_shift_stable in - Loc.tag @@ GLetIn(nme,bdy,typ,newtrm) + let newtrm = merge_app_unsafe trm c2 shift filter_shift_stable in + CAst.make @@ GLetIn(nme,bdy,typ,newtrm) | _, GLetIn(nme,bdy,typ,trm) -> let _ = prstr "\nICI3 '!\n" in - let newtrm = merge_app_unsafe (l1, c1) trm shift filter_shift_stable in - Loc.tag @@ GLetIn(nme,bdy,typ,newtrm) + let newtrm = merge_app_unsafe c1 trm shift filter_shift_stable in + CAst.make @@ GLetIn(nme,bdy,typ,newtrm) | _ -> let _ = prstr "\nICI4 '!\n" in raise NoMerge @@ -550,14 +550,14 @@ let rec merge_rec_hyps shift accrec filter_shift_stable : (Name.t * glob_constr option * glob_constr option) list = let mergeonehyp t reldecl = match reldecl with - | (nme,x,Some (_, GApp(i,args) as ind)) + | (nme,x,Some ({ CAst.v = GApp(i,args)} as ind)) -> nme,x, Some (merge_app_unsafe ind t shift filter_shift_stable) | (nme,Some _,None) -> error "letins with recursive calls not treated yet" | (nme,None,Some _) -> assert false | (nme,None,None) | (nme,Some _,Some _) -> assert false in match ltyp with | [] -> [] - | (nme,None,Some (_, GApp(f, largs) as t)) :: lt when isVarf ind2name f -> + | (nme,None,Some ({ CAst. v = GApp(f, largs) } as t)) :: lt when isVarf ind2name f -> let rechyps = List.map (mergeonehyp t) accrec in rechyps @ merge_rec_hyps shift accrec lt filter_shift_stable | e::lt -> e :: merge_rec_hyps shift accrec lt filter_shift_stable @@ -573,7 +573,7 @@ let find_app (nme:Id.t) ltyp = (List.map (fun x -> match x with - | _,None,Some (_, GApp(f,_)) when isVarf nme f -> raise (Found 0) + | _,None,Some { CAst.v = GApp(f,_)} when isVarf nme f -> raise (Found 0) | _ -> ()) ltyp); false @@ -632,8 +632,8 @@ let rec merge_types shift accrec1 rechyps , concl | (nme,None, Some t1)as e ::lt1 -> - (match t1 with - | _, GApp(f,carr) when isVarf ind1name f -> + (match t1.CAst.v with + | GApp(f,carr) when isVarf ind1name f -> merge_types shift (e::accrec1) lt1 concl1 ltyp2 concl2 | _ -> let recres, recconcl2 = @@ -864,7 +864,7 @@ let mkProd_reldecl (rdecl:Context.Rel.Declaration.t) (t2:glob_constr) = | LocalAssum (nme,t) -> let t = EConstr.of_constr t in let traw = Detyping.detype false [] (Global.env()) Evd.empty t in - Loc.tag @@ GProd (nme,Explicit,traw,t2) + CAst.make @@ GProd (nme,Explicit,traw,t2) | LocalDef _ -> assert false diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 9e469c756..8c8839944 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -188,15 +188,15 @@ let (value_f:Constr.constr list -> global_reference -> Constr.constr) = in let env = Environ.push_rel_context context (Global.env ()) in let glob_body = - Loc.tag @@ + CAst.make @@ GCases (RegularStyle,None, - [Loc.tag @@ GApp(Loc.tag @@ GRef(fterm,None), List.rev_map (fun x_id -> Loc.tag @@ GVar x_id) rev_x_id_l), + [CAst.make @@ GApp(CAst.make @@ GRef(fterm,None), List.rev_map (fun x_id -> CAst.make @@ GVar x_id) rev_x_id_l), (Anonymous,None)], - [Loc.tag ([v_id], [Loc.tag @@ PatCstr ((destIndRef (delayed_force coq_sig_ref),1), - [Loc.tag @@ PatVar(Name v_id); Loc.tag @@ PatVar Anonymous], + [Loc.tag ([v_id], [CAst.make @@ PatCstr ((destIndRef (delayed_force coq_sig_ref),1), + [CAst.make @@ PatVar(Name v_id); CAst.make @@ PatVar Anonymous], Anonymous)], - Loc.tag @@ GVar v_id)]) + CAst.make @@ GVar v_id)]) in let body = fst (understand env (Evd.from_env env) glob_body)(*FIXME*) in it_mkLambda_or_LetIn body context diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4 index dc43930e4..9d50b6e6f 100644 --- a/plugins/ltac/extratactics.ml4 +++ b/plugins/ltac/extratactics.ml4 @@ -631,14 +631,14 @@ let subst_var_with_hole occ tid t = let occref = if occ > 0 then ref occ else Find_subterm.error_invalid_occurrence [occ] in let locref = ref 0 in let rec substrec = function - | (_, GVar id) as x -> + | { CAst.v = GVar id } as x -> if Id.equal id tid then (decr occref; if Int.equal !occref 0 then x else (incr locref; - Loc.tag ~loc:(Loc.make_loc (!locref,0)) @@ + CAst.make ~loc:(Loc.make_loc (!locref,0)) @@ GHole (Evar_kinds.QuestionMark(Evar_kinds.Define true), Misctypes.IntroAnonymous, None))) else x @@ -651,12 +651,12 @@ let subst_hole_with_term occ tc t = let locref = ref 0 in let occref = ref occ in let rec substrec = function - | _, GHole (Evar_kinds.QuestionMark(Evar_kinds.Define true),Misctypes.IntroAnonymous,s) -> + | { CAst.v = GHole (Evar_kinds.QuestionMark(Evar_kinds.Define true),Misctypes.IntroAnonymous,s) } -> decr occref; if Int.equal !occref 0 then tc else (incr locref; - Loc.tag ~loc:(Loc.make_loc (!locref,0)) @@ + CAst.make ~loc:(Loc.make_loc (!locref,0)) @@ GHole (Evar_kinds.QuestionMark(Evar_kinds.Define true),Misctypes.IntroAnonymous,s)) | c -> map_glob_constr_left_to_right substrec c in diff --git a/plugins/ltac/g_rewrite.ml4 b/plugins/ltac/g_rewrite.ml4 index 1f40c67b5..b4a0e46ae 100644 --- a/plugins/ltac/g_rewrite.ml4 +++ b/plugins/ltac/g_rewrite.ml4 @@ -123,7 +123,7 @@ TACTIC EXTEND rewrite_strat END let clsubstitute o c = - let is_tac id = match fst (fst (snd c)) with (_, GVar id') when Id.equal id' id -> true | _ -> false in + let is_tac id = match fst (fst (snd c)) with { CAst.v = GVar id'} when Id.equal id' id -> true | _ -> false in Tacticals.onAllHypsAndConcl (fun cl -> match cl with diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml index 58473d7dd..87b79374e 100644 --- a/plugins/ltac/pptactic.ml +++ b/plugins/ltac/pptactic.ml @@ -1085,7 +1085,7 @@ type 'a extra_genarg_printer = let strip_prod_binders_glob_constr n (ty,_) = let rec strip_ty acc n ty = if Int.equal n 0 then (List.rev acc, (ty,None)) else - match Loc.obj ty with + match ty.CAst.v with Glob_term.GProd(na,Explicit,a,b) -> strip_ty (([Loc.tag na],(a,None))::acc) (n-1) b | _ -> error "Cannot translate fix tactic: not enough products" in diff --git a/plugins/ltac/tacintern.ml b/plugins/ltac/tacintern.ml index 566dd8ed7..8751a14c7 100644 --- a/plugins/ltac/tacintern.ml +++ b/plugins/ltac/tacintern.ml @@ -108,12 +108,12 @@ let intern_ltac_variable ist = function let intern_constr_reference strict ist = function | Ident (_,id) as r when not strict && find_hyp id ist -> - (Loc.tag @@ GVar id), Some (CAst.make @@ CRef (r,None)) + (CAst.make @@ GVar id), Some (CAst.make @@ CRef (r,None)) | Ident (_,id) as r when find_var id ist -> - (Loc.tag @@ GVar id), if strict then None else Some (CAst.make @@ CRef (r,None)) + (CAst.make @@ GVar id), if strict then None else Some (CAst.make @@ CRef (r,None)) | r -> let loc,_ as lqid = qualid_of_reference r in - Loc.tag @@ GRef (locate_global_with_alias lqid,None), + CAst.make @@ GRef (locate_global_with_alias lqid,None), if strict then None else Some (CAst.make @@ CRef (r,None)) let intern_move_location ist = function @@ -272,7 +272,7 @@ let intern_destruction_arg ist = function if !strict_check then (* If in a defined tactic, no intros-until *) match intern_constr ist (CAst.make @@ CRef (Ident (Loc.tag id), None)) with - | (loc, GVar id), _ -> clear,ElimOnIdent (loc,id) + | {loc; CAst.v = GVar id}, _ -> clear,ElimOnIdent (loc,id) | c -> clear,ElimOnConstr (c,NoBindings) else clear,ElimOnIdent (loc,id) @@ -350,7 +350,7 @@ let intern_typed_pattern_or_ref_with_occurrences ist (l,p) = | _ -> Qualid (loc,qualid_of_path (path_of_global (smart_global r))) in let sign = { Constrintern.ltac_vars = ist.ltacvars; Constrintern.ltac_bound = Id.Set.empty } in let c = Constrintern.interp_reference sign r in - match Loc.obj c with + match c.CAst.v with | GRef (r,None) -> Inl (ArgArg (evaluable_of_global_reference ist.genv r,None)) | GVar id -> diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml index 449027b52..91bc46fe7 100644 --- a/plugins/ltac/tacinterp.ml +++ b/plugins/ltac/tacinterp.ml @@ -697,7 +697,7 @@ let interp_typed_pattern ist env sigma (_,c,_) = let interp_constr_in_compound_list inj_fun dest_fun interp_fun ist env sigma l = let try_expand_ltac_var sigma x = try match dest_fun x with - | (_, GVar id), _ -> + | { CAst.v = GVar id }, _ -> let v = Id.Map.find id ist.lfun in sigma, List.map inj_fun (coerce_to_constr_list env v) | _ -> @@ -1072,7 +1072,7 @@ let interp_destruction_arg ist gl arg = if Tactics.is_quantified_hypothesis id gl then keep,ElimOnIdent (loc,id) else - let c = (Loc.tag ?loc @@ GVar id,Some (CAst.make @@ CRef (Ident (loc,id),None))) in + let c = (CAst.make ?loc @@ GVar id,Some (CAst.make @@ CRef (Ident (loc,id),None))) in let f = { delayed = fun env sigma -> let sigma = Sigma.to_evar_map sigma in let (sigma,c) = interp_open_constr ist env sigma c in diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index e2a6ad55c..2b64204c9 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -131,7 +131,7 @@ let closed_term_ast l = let l = List.map (fun gr -> ArgArg(Loc.tag gr)) l in TacFun([Name(Id.of_string"t")], TacML(Loc.tag (tacname, - [TacGeneric (Genarg.in_gen (Genarg.glbwit Stdarg.wit_constr) (Loc.tag @@ GVar(Id.of_string"t"),None)); + [TacGeneric (Genarg.in_gen (Genarg.glbwit Stdarg.wit_constr) (CAst.make @@ GVar(Id.of_string"t"),None)); TacGeneric (Genarg.in_gen (Genarg.glbwit (Genarg.wit_list Stdarg.wit_ref)) l)]))) (* let _ = add_tacdef false ((Loc.ghost,Id.of_string"ring_closed_term" diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index c11c9f83b..9a0dfbf1c 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -158,10 +158,10 @@ let mkCLetIn ?loc name bo t = CAst.make ?loc @@ CLetIn ((Loc.tag ?loc name), bo, None, t) let mkCCast ?loc t ty = CAst.make ?loc @@ CCast (t, dC ty) (** Constructors for rawconstr *) -let mkRHole = Loc.tag @@ GHole (InternalHole, IntroAnonymous, None) -let mkRApp f args = if args = [] then f else Loc.tag @@ GApp (f, args) -let mkRCast rc rt = Loc.tag @@ GCast (rc, dC rt) -let mkRLambda n s t = Loc.tag @@ GLambda (n, Explicit, s, t) +let mkRHole = CAst.make @@ GHole (InternalHole, IntroAnonymous, None) +let mkRApp f args = if args = [] then f else CAst.make @@ GApp (f, args) +let mkRCast rc rt = CAst.make @@ GCast (rc, dC rt) +let mkRLambda n s t = CAst.make @@ GLambda (n, Explicit, s, t) (* ssrterm conbinators *) let combineCG t1 t2 f g = match t1, t2 with @@ -1022,7 +1022,7 @@ type pattern = Evd.evar_map * (constr, constr) ssrpattern let id_of_cpattern = let open CAst in function | _,(_,Some { v = CRef (Ident (_, x), _) } ) -> Some x | _,(_,Some { v = CAppExpl ((_, Ident (_, x), _), []) } ) -> Some x - | _,((_, GRef (VarRef x, _)) ,None) -> Some x + | _,({ v = GRef (VarRef x, _)} ,None) -> Some x | _ -> None let id_of_Cterm t = match id_of_cpattern t with | Some x -> x @@ -1121,9 +1121,10 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty = let eAsXInT e x t = E_As_X_In_T(e,x,t) in let mkG ?(k=' ') x = k,(x,None) in let decode ist t ?reccall f g = + let open CAst in try match (pf_intern_term ist gl t) with - | _, GCast((_, GHole _),CastConv((_, GLambda(Name x,_,_,c)))) -> f x (' ',(c,None)) - | _, GVar id + | { v = GCast({ v = GHole _},CastConv({ v = GLambda(Name x,_,_,c)})) } -> f x (' ',(c,None)) + | { v = GVar id } when Id.Map.mem id ist.lfun && not(Option.is_empty reccall) && not(Option.is_empty wit_ssrpatternarg) -> @@ -1164,18 +1165,18 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty = thin name sigma e) sigma new_evars in sigma in - let red = let rec decode_red (ist,red) = match red with - | T(k,((_, GCast ((_, GHole _),CastConv((_, GLambda (Name id,_,_,t))))),None)) + let red = let rec decode_red (ist,red) = let open CAst in match red with + | T(k,({ v = GCast ({ v = GHole _ },CastConv({ v = GLambda (Name id,_,_,t)}))},None)) when let id = string_of_id id in let len = String.length id in (len > 8 && String.sub id 0 8 = "_ssrpat_") -> let id = string_of_id id in let len = String.length id in (match String.sub id 8 (len - 8), t with - | "In", (_, GApp( _, [t])) -> decodeG t xInT (fun x -> T x) - | "In", (_, GApp( _, [e; t])) -> decodeG t (eInXInT (mkG e)) (bad_enc id) - | "In", (_, GApp( _, [e; t; e_in_t])) -> + | "In", { v = GApp( _, [t]) } -> decodeG t xInT (fun x -> T x) + | "In", { v = GApp( _, [e; t]) } -> decodeG t (eInXInT (mkG e)) (bad_enc id) + | "In", { v = GApp( _, [e; t; e_in_t]) } -> decodeG t (eInXInT (mkG e)) (fun _ -> decodeG e_in_t xInT (fun _ -> assert false)) - | "As", (_, GApp(_, [e; t])) -> decodeG t (eAsXInT (mkG e)) (bad_enc id) + | "As", { v = GApp(_, [e; t]) } -> decodeG t (eAsXInT (mkG e)) (bad_enc id) | _ -> bad_enc id ()) | T t -> decode ist ~reccall:decode_red t xInT (fun x -> T x) | In_T t -> decode ist t inXInT inT @@ -1201,7 +1202,7 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty = pp(lazy(str"typed as: " ++ pr_pattern_w_ids red)); let mkXLetIn ?loc x (a,(g,c)) = match c with | Some b -> a,(g,Some (mkCLetIn ?loc x (mkCHole ?loc) b)) - | None -> a,(Loc.tag ?loc @@ GLetIn (x, Loc.tag ?loc @@ GHole (BinderType x, IntroAnonymous, None), None, g), None) in + | None -> a,(CAst.make ?loc @@ GLetIn (x, CAst.make ?loc @@ GHole (BinderType x, IntroAnonymous, None), None, g), None) in match red with | T t -> let sigma, t = interp_term ist gl t in sigma, T t | In_T t -> let sigma, t = interp_term ist gl t in sigma, In_T t @@ -1374,10 +1375,10 @@ let pf_fill_occ_term gl occ t = let cl,(_,t) = fill_occ_term env concl occ sigma0 t in cl, t -let cpattern_of_id id = ' ', (Loc.tag @@ GRef (VarRef id, None), None) +let cpattern_of_id id = ' ', (CAst.make @@ GRef (VarRef id, None), None) let is_wildcard : cpattern -> bool = function - | _,(_,Some { CAst.v = CHole _ } | (_, GHole _),None) -> true + | _,(_,Some { CAst.v = CHole _ } | { CAst.v = GHole _ } ,None) -> true | _ -> false (* "ssrpattern" *) diff --git a/plugins/syntax/ascii_syntax.ml b/plugins/syntax/ascii_syntax.ml index ed977c416..e7eea0284 100644 --- a/plugins/syntax/ascii_syntax.ml +++ b/plugins/syntax/ascii_syntax.ml @@ -41,9 +41,9 @@ let interp_ascii ?loc p = let rec aux n p = if Int.equal n 0 then [] else let mp = p mod 2 in - (Loc.tag ?loc @@ GRef ((if Int.equal mp 0 then glob_false else glob_true),None)) + (CAst.make ?loc @@ GRef ((if Int.equal mp 0 then glob_false else glob_true),None)) :: (aux (n-1) (p/2)) in - Loc.tag ?loc @@ GApp (Loc.tag ?loc @@ GRef(force glob_Ascii,None), aux 8 p) + CAst.make ?loc @@ GApp (CAst.make ?loc @@ GRef(force glob_Ascii,None), aux 8 p) let interp_ascii_string ?loc s = let p = @@ -59,12 +59,12 @@ let interp_ascii_string ?loc s = let uninterp_ascii r = let rec uninterp_bool_list n = function | [] when Int.equal n 0 -> 0 - | (_, GRef (k,_))::l when Globnames.eq_gr k glob_true -> 1+2*(uninterp_bool_list (n-1) l) - | (_, GRef (k,_))::l when Globnames.eq_gr k glob_false -> 2*(uninterp_bool_list (n-1) l) + | { CAst.v = GRef (k,_)}::l when Globnames.eq_gr k glob_true -> 1+2*(uninterp_bool_list (n-1) l) + | { CAst.v = GRef (k,_)}::l when Globnames.eq_gr k glob_false -> 2*(uninterp_bool_list (n-1) l) | _ -> raise Non_closed_ascii in try let aux = function - | _, GApp ((_, GRef (k,_)),l) when Globnames.eq_gr k (force glob_Ascii) -> uninterp_bool_list 8 l + | { CAst.v = GApp ({ CAst.v = GRef (k,_)},l) } when Globnames.eq_gr k (force glob_Ascii) -> uninterp_bool_list 8 l | _ -> raise Non_closed_ascii in Some (aux r) with @@ -80,4 +80,4 @@ let _ = Notation.declare_string_interpreter "char_scope" (ascii_path,ascii_module) interp_ascii_string - ([Loc.tag @@ GRef (static_glob_Ascii,None)], uninterp_ascii_string, true) + ([CAst.make @@ GRef (static_glob_Ascii,None)], uninterp_ascii_string, true) diff --git a/plugins/syntax/nat_syntax.ml b/plugins/syntax/nat_syntax.ml index 5cdd82024..9a4cd6c25 100644 --- a/plugins/syntax/nat_syntax.ml +++ b/plugins/syntax/nat_syntax.ml @@ -36,11 +36,11 @@ let warn_large_nat = let nat_of_int ?loc n = if is_pos_or_zero n then begin if less_than threshold n then warn_large_nat (); - let ref_O = Loc.tag ?loc @@ GRef (glob_O, None) in - let ref_S = Loc.tag ?loc @@ GRef (glob_S, None) in + let ref_O = CAst.make ?loc @@ GRef (glob_O, None) in + let ref_S = CAst.make ?loc @@ GRef (glob_S, None) in let rec mk_nat acc n = if n <> zero then - mk_nat (Loc.tag ?loc @@ GApp (ref_S, [acc])) (sub_1 n) + mk_nat (CAst.make ?loc @@ GApp (ref_S, [acc])) (sub_1 n) else acc in @@ -55,8 +55,8 @@ let nat_of_int ?loc n = exception Non_closed_number -let rec int_of_nat x = Loc.with_unloc (function - | GApp ((_, GRef (s,_)),[a]) when Globnames.eq_gr s glob_S -> add_1 (int_of_nat a) +let rec int_of_nat x = CAst.with_val (function + | GApp ({ CAst.v = GRef (s,_) } ,[a]) when Globnames.eq_gr s glob_S -> add_1 (int_of_nat a) | GRef (z,_) when Globnames.eq_gr z glob_O -> zero | _ -> raise Non_closed_number ) x @@ -74,4 +74,4 @@ let _ = Notation.declare_numeral_interpreter "nat_scope" (nat_path,datatypes_module_name) nat_of_int - ([Loc.tag @@ GRef (glob_S,None); Loc.tag @@ GRef (glob_O,None)], uninterp_nat, true) + ([CAst.make @@ GRef (glob_S,None); CAst.make @@ GRef (glob_O,None)], uninterp_nat, true) diff --git a/plugins/syntax/numbers_syntax.ml b/plugins/syntax/numbers_syntax.ml index 3ee64ba7e..e23852bf8 100644 --- a/plugins/syntax/numbers_syntax.ml +++ b/plugins/syntax/numbers_syntax.ml @@ -87,9 +87,9 @@ exception Non_closed (* parses a *non-negative* integer (from bigint.ml) into an int31 wraps modulo 2^31 *) let int31_of_pos_bigint ?loc n = - let ref_construct = Loc.tag ?loc @@ GRef (int31_construct, None) in - let ref_0 = Loc.tag ?loc @@ GRef (int31_0, None) in - let ref_1 = Loc.tag ?loc @@ GRef (int31_1, None) in + let ref_construct = CAst.make ?loc @@ GRef (int31_construct, None) in + let ref_0 = CAst.make ?loc @@ GRef (int31_0, None) in + let ref_1 = CAst.make ?loc @@ GRef (int31_1, None) in let rec args counter n = if counter <= 0 then [] @@ -97,7 +97,7 @@ let int31_of_pos_bigint ?loc n = let (q,r) = div2_with_rest n in (if r then ref_1 else ref_0)::(args (counter-1) q) in - Loc.tag ?loc @@ GApp (ref_construct, List.rev (args 31 n)) + CAst.make ?loc @@ GApp (ref_construct, List.rev (args 31 n)) let error_negative ?loc = CErrors.user_err ?loc ~hdr:"interp_int31" (Pp.str "int31 are only non-negative numbers.") @@ -114,12 +114,12 @@ let bigint_of_int31 = let rec args_parsing args cur = match args with | [] -> cur - | (_, GRef (b,_))::l when eq_gr b int31_0 -> args_parsing l (mult_2 cur) - | (_, GRef (b,_))::l when eq_gr b int31_1 -> args_parsing l (add_1 (mult_2 cur)) + | { CAst.v = GRef (b,_) }::l when eq_gr b int31_0 -> args_parsing l (mult_2 cur) + | { CAst.v = GRef (b,_) }::l when eq_gr b int31_1 -> args_parsing l (add_1 (mult_2 cur)) | _ -> raise Non_closed in function - | _, GApp ((_, GRef (c, _)), args) when eq_gr c int31_construct -> args_parsing args zero + | { CAst.v = GApp ({ CAst.v = GRef (c, _)}, args) } when eq_gr c int31_construct -> args_parsing args zero | _ -> raise Non_closed let uninterp_int31 i = @@ -132,7 +132,7 @@ let uninterp_int31 i = let _ = Notation.declare_numeral_interpreter int31_scope (int31_path, int31_module) interp_int31 - ([Loc.tag @@ GRef (int31_construct, None)], + ([CAst.make @@ GRef (int31_construct, None)], uninterp_int31, true) @@ -163,16 +163,16 @@ let height bi = (* n must be a non-negative integer (from bigint.ml) *) let word_of_pos_bigint ?loc hght n = - let ref_W0 = Loc.tag ?loc @@ GRef (zn2z_W0, None) in - let ref_WW = Loc.tag ?loc @@ GRef (zn2z_WW, None) in + let ref_W0 = CAst.make ?loc @@ GRef (zn2z_W0, None) in + let ref_WW = CAst.make ?loc @@ GRef (zn2z_WW, None) in let rec decomp hgt n = if hgt <= 0 then int31_of_pos_bigint ?loc n else if equal n zero then - Loc.tag ?loc @@ GApp (ref_W0, [Loc.tag ?loc @@ GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None)]) + CAst.make ?loc @@ GApp (ref_W0, [CAst.make ?loc @@ GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None)]) else let (h,l) = split_at hgt n in - Loc.tag ?loc @@ GApp (ref_WW, [Loc.tag ?loc @@ GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None); + CAst.make ?loc @@ GApp (ref_WW, [CAst.make ?loc @@ GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None); decomp (hgt-1) h; decomp (hgt-1) l]) in @@ -180,13 +180,13 @@ let word_of_pos_bigint ?loc hght n = let bigN_of_pos_bigint ?loc n = let h = height n in - let ref_constructor = Loc.tag ?loc @@ GRef (bigN_constructor h, None) in + let ref_constructor = CAst.make ?loc @@ GRef (bigN_constructor h, None) in let word = word_of_pos_bigint ?loc h n in let args = if h < n_inlined then [word] else [Nat_syntax_plugin.Nat_syntax.nat_of_int ?loc (of_int (h-n_inlined));word] in - Loc.tag ?loc @@ GApp (ref_constructor, args) + CAst.make ?loc @@ GApp (ref_constructor, args) let bigN_error_negative ?loc = CErrors.user_err ?loc ~hdr:"interp_bigN" (Pp.str "bigN are only non-negative numbers.") @@ -203,14 +203,14 @@ let interp_bigN ?loc n = let bigint_of_word = let rec get_height rc = match rc with - | _, GApp ((_, GRef(c,_)), [_;lft;rght]) when eq_gr c zn2z_WW -> + | { CAst.v = GApp ({ CAst.v = GRef(c,_)}, [_;lft;rght]) } when eq_gr c zn2z_WW -> 1+max (get_height lft) (get_height rght) | _ -> 0 in let rec transform hght rc = match rc with - | _, GApp ((_, GRef(c,_)),_) when eq_gr c zn2z_W0-> zero - | _, GApp ((_, GRef(c,_)), [_;lft;rght]) when eq_gr c zn2z_WW-> + | { CAst.v = GApp ({ CAst.v = GRef(c,_)},_)} when eq_gr c zn2z_W0-> zero + | { CAst.v = GApp ({ CAst.v = GRef(c,_)}, [_;lft;rght]) } when eq_gr c zn2z_WW-> let new_hght = hght-1 in add (mult (rank new_hght) (transform new_hght lft)) @@ -223,8 +223,8 @@ let bigint_of_word = let bigint_of_bigN rc = match rc with - | _, GApp (_,[one_arg]) -> bigint_of_word one_arg - | _, GApp (_,[_;second_arg]) -> bigint_of_word second_arg + | { CAst.v = GApp (_,[one_arg]) } -> bigint_of_word one_arg + | { CAst.v = GApp (_,[_;second_arg]) } -> bigint_of_word second_arg | _ -> raise Non_closed let uninterp_bigN rc = @@ -240,7 +240,7 @@ let uninterp_bigN rc = let bigN_list_of_constructors = let rec build i = if i < n_inlined+1 then - (Loc.tag @@ GRef (bigN_constructor i,None))::(build (i+1)) + (CAst.make @@ GRef (bigN_constructor i,None))::(build (i+1)) else [] in @@ -257,17 +257,17 @@ let _ = Notation.declare_numeral_interpreter bigN_scope (*** Parsing for bigZ in digital notation ***) let interp_bigZ ?loc n = - let ref_pos = Loc.tag ?loc @@ GRef (bigZ_pos, None) in - let ref_neg = Loc.tag ?loc @@ GRef (bigZ_neg, None) in + let ref_pos = CAst.make ?loc @@ GRef (bigZ_pos, None) in + let ref_neg = CAst.make ?loc @@ GRef (bigZ_neg, None) in if is_pos_or_zero n then - Loc.tag ?loc @@ GApp (ref_pos, [bigN_of_pos_bigint ?loc n]) + CAst.make ?loc @@ GApp (ref_pos, [bigN_of_pos_bigint ?loc n]) else - Loc.tag ?loc @@ GApp (ref_neg, [bigN_of_pos_bigint ?loc (neg n)]) + CAst.make ?loc @@ GApp (ref_neg, [bigN_of_pos_bigint ?loc (neg n)]) (* pretty printing functions for bigZ *) let bigint_of_bigZ = function - | _, GApp ((_, GRef(c,_)), [one_arg]) when eq_gr c bigZ_pos -> bigint_of_bigN one_arg - | _, GApp ((_, GRef(c,_)), [one_arg]) when eq_gr c bigZ_neg -> + | { CAst.v = GApp ({ CAst.v = GRef(c,_) }, [one_arg])} when eq_gr c bigZ_pos -> bigint_of_bigN one_arg + | { CAst.v = GApp ({ CAst.v = GRef(c,_) }, [one_arg])} when eq_gr c bigZ_neg -> let opp_val = bigint_of_bigN one_arg in if equal opp_val zero then raise Non_closed @@ -286,19 +286,19 @@ let uninterp_bigZ rc = let _ = Notation.declare_numeral_interpreter bigZ_scope (bigZ_path, bigZ_module) interp_bigZ - ([Loc.tag @@ GRef (bigZ_pos, None); - Loc.tag @@ GRef (bigZ_neg, None)], + ([CAst.make @@ GRef (bigZ_pos, None); + CAst.make @@ GRef (bigZ_neg, None)], uninterp_bigZ, true) (*** Parsing for bigQ in digital notation ***) let interp_bigQ ?loc n = - let ref_z = Loc.tag ?loc @@ GRef (bigQ_z, None) in - Loc.tag ?loc @@ GApp (ref_z, [interp_bigZ ?loc n]) + let ref_z = CAst.make ?loc @@ GRef (bigQ_z, None) in + CAst.make ?loc @@ GApp (ref_z, [interp_bigZ ?loc n]) let uninterp_bigQ rc = try match rc with - | _, GApp ((_, GRef(c,_)), [one_arg]) when eq_gr c bigQ_z -> + | { CAst.v = GApp ({ CAst.v = GRef(c,_)}, [one_arg]) } when eq_gr c bigQ_z -> Some (bigint_of_bigZ one_arg) | _ -> None (* we don't pretty-print yet fractions *) with Non_closed -> None @@ -307,5 +307,5 @@ let uninterp_bigQ rc = let _ = Notation.declare_numeral_interpreter bigQ_scope (bigQ_path, bigQ_module) interp_bigQ - ([Loc.tag @@ GRef (bigQ_z, None)], uninterp_bigQ, + ([CAst.make @@ GRef (bigQ_z, None)], uninterp_bigQ, true) diff --git a/plugins/syntax/r_syntax.ml b/plugins/syntax/r_syntax.ml index b7041d045..7ce066c59 100644 --- a/plugins/syntax/r_syntax.ml +++ b/plugins/syntax/r_syntax.ml @@ -42,13 +42,13 @@ let glob_xO = ConstructRef path_of_xO let glob_xH = ConstructRef path_of_xH let pos_of_bignat ?loc x = - let ref_xI = Loc.tag @@ GRef (glob_xI, None) in - let ref_xH = Loc.tag @@ GRef (glob_xH, None) in - let ref_xO = Loc.tag @@ GRef (glob_xO, None) in + let ref_xI = CAst.make @@ GRef (glob_xI, None) in + let ref_xH = CAst.make @@ GRef (glob_xH, None) in + let ref_xO = CAst.make @@ GRef (glob_xO, None) in let rec pos_of x = match div2_with_rest x with - | (q,false) -> Loc.tag @@ GApp (ref_xO,[pos_of q]) - | (q,true) when not (Bigint.equal q zero) -> Loc.tag @@ GApp (ref_xI,[pos_of q]) + | (q,false) -> CAst.make @@ GApp (ref_xO,[pos_of q]) + | (q,true) when not (Bigint.equal q zero) -> CAst.make @@ GApp (ref_xI,[pos_of q]) | (q,true) -> ref_xH in pos_of x @@ -58,9 +58,9 @@ let pos_of_bignat ?loc x = (**********************************************************************) let rec bignat_of_pos = function - | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_xO -> mult_2(bignat_of_pos a) - | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_xI -> add_1(mult_2(bignat_of_pos a)) - | _, GRef (a, _) when Globnames.eq_gr a glob_xH -> Bigint.one + | { CAst.v = GApp ({ CAst.v = GRef (b,_)},[a]) } when Globnames.eq_gr b glob_xO -> mult_2(bignat_of_pos a) + | { CAst.v = GApp ({ CAst.v = GRef (b,_)},[a]) } when Globnames.eq_gr b glob_xI -> add_1(mult_2(bignat_of_pos a)) + | { CAst.v = GRef (a, _) } when Globnames.eq_gr a glob_xH -> Bigint.one | _ -> raise Non_closed_number (**********************************************************************) @@ -81,18 +81,18 @@ let z_of_int ?loc n = if not (Bigint.equal n zero) then let sgn, n = if is_pos_or_zero n then glob_POS, n else glob_NEG, Bigint.neg n in - Loc.tag @@ GApp(Loc.tag @@ GRef (sgn,None), [pos_of_bignat ?loc n]) + CAst.make @@ GApp(CAst.make @@ GRef (sgn,None), [pos_of_bignat ?loc n]) else - Loc.tag @@ GRef (glob_ZERO, None) + CAst.make @@ GRef (glob_ZERO, None) (**********************************************************************) (* Printing Z via scopes *) (**********************************************************************) let bigint_of_z = function - | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_POS -> bignat_of_pos a - | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_NEG -> Bigint.neg (bignat_of_pos a) - | _, GRef (a, _) when Globnames.eq_gr a glob_ZERO -> Bigint.zero + | { CAst.v = GApp ({ CAst.v = GRef (b,_)},[a]) } when Globnames.eq_gr b glob_POS -> bignat_of_pos a + | { CAst.v = GApp ({ CAst.v = GRef (b,_)},[a]) } when Globnames.eq_gr b glob_NEG -> Bigint.neg (bignat_of_pos a) + | { CAst.v = GRef (a, _) } when Globnames.eq_gr a glob_ZERO -> Bigint.zero | _ -> raise Non_closed_number (**********************************************************************) @@ -108,14 +108,14 @@ let make_path dir id = Globnames.encode_con dir (Id.of_string id) let glob_IZR = ConstRef (make_path (make_dir rdefinitions) "IZR") let r_of_int ?loc z = - Loc.tag @@ GApp (Loc.tag @@ GRef(glob_IZR,None), [z_of_int ?loc z]) + CAst.make @@ GApp (CAst.make @@ GRef(glob_IZR,None), [z_of_int ?loc z]) (**********************************************************************) (* Printing R via scopes *) (**********************************************************************) let bigint_of_r = function - | _, GApp ((_, GRef (o,_)), [a]) when Globnames.eq_gr o glob_IZR -> + | { CAst.v = GApp ({ CAst.v = GRef (o,_) }, [a]) } when Globnames.eq_gr o glob_IZR -> bigint_of_z a | _ -> raise Non_closed_number @@ -128,6 +128,6 @@ let uninterp_r p = let _ = Notation.declare_numeral_interpreter "R_scope" (r_path,["Coq";"Reals";"Rdefinitions"]) r_of_int - ([Loc.tag @@ GRef (glob_IZR, None)], + ([CAst.make @@ GRef (glob_IZR, None)], uninterp_r, false) diff --git a/plugins/syntax/string_syntax.ml b/plugins/syntax/string_syntax.ml index 49cb9355c..b7f13b040 100644 --- a/plugins/syntax/string_syntax.ml +++ b/plugins/syntax/string_syntax.ml @@ -36,8 +36,8 @@ open Lazy let interp_string ?loc s = let le = String.length s in let rec aux n = - if n = le then Loc.tag ?loc @@ GRef (force glob_EmptyString, None) else - Loc.tag ?loc @@ GApp (Loc.tag ?loc @@ GRef (force glob_String, None), + if n = le then CAst.make ?loc @@ GRef (force glob_EmptyString, None) else + CAst.make ?loc @@ GApp (CAst.make ?loc @@ GRef (force glob_String, None), [interp_ascii ?loc (int_of_char s.[n]); aux (n+1)]) in aux 0 @@ -45,11 +45,11 @@ let uninterp_string r = try let b = Buffer.create 16 in let rec aux = function - | _, GApp ((_, GRef (k,_)),[a;s]) when eq_gr k (force glob_String) -> + | { CAst.v = GApp ({ CAst.v = GRef (k,_) },[a;s]) } when eq_gr k (force glob_String) -> (match uninterp_ascii a with | Some c -> Buffer.add_char b (Char.chr c); aux s | _ -> raise Non_closed_string) - | _, GRef (z,_) when eq_gr z (force glob_EmptyString) -> + | { CAst.v = GRef (z,_) } when eq_gr z (force glob_EmptyString) -> Some (Buffer.contents b) | _ -> raise Non_closed_string @@ -61,6 +61,6 @@ let _ = Notation.declare_string_interpreter "string_scope" (string_path,["Coq";"Strings";"String"]) interp_string - ([Loc.tag @@ GRef (static_glob_String,None); - Loc.tag @@ GRef (static_glob_EmptyString,None)], + ([CAst.make @@ GRef (static_glob_String,None); + CAst.make @@ GRef (static_glob_EmptyString,None)], uninterp_string, true) diff --git a/plugins/syntax/z_syntax.ml b/plugins/syntax/z_syntax.ml index 96c1f3e39..479448e06 100644 --- a/plugins/syntax/z_syntax.ml +++ b/plugins/syntax/z_syntax.ml @@ -45,13 +45,13 @@ let glob_xO = ConstructRef path_of_xO let glob_xH = ConstructRef path_of_xH let pos_of_bignat ?loc x = - let ref_xI = Loc.tag ?loc @@ GRef (glob_xI, None) in - let ref_xH = Loc.tag ?loc @@ GRef (glob_xH, None) in - let ref_xO = Loc.tag ?loc @@ GRef (glob_xO, None) in + let ref_xI = CAst.make ?loc @@ GRef (glob_xI, None) in + let ref_xH = CAst.make ?loc @@ GRef (glob_xH, None) in + let ref_xO = CAst.make ?loc @@ GRef (glob_xO, None) in let rec pos_of x = match div2_with_rest x with - | (q,false) -> Loc.tag ?loc @@ GApp (ref_xO,[pos_of q]) - | (q,true) when not (Bigint.equal q zero) -> Loc.tag ?loc @@ GApp (ref_xI,[pos_of q]) + | (q,false) -> CAst.make ?loc @@ GApp (ref_xO,[pos_of q]) + | (q,true) when not (Bigint.equal q zero) -> CAst.make ?loc @@ GApp (ref_xI,[pos_of q]) | (q,true) -> ref_xH in pos_of x @@ -68,11 +68,12 @@ let interp_positive ?loc n = (* Printing positive via scopes *) (**********************************************************************) -let rec bignat_of_pos = function - | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_xO -> mult_2(bignat_of_pos a) - | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_xI -> add_1(mult_2(bignat_of_pos a)) - | _, GRef (a, _) when Globnames.eq_gr a glob_xH -> Bigint.one +let rec bignat_of_pos x = CAst.with_val (function + | GApp ({ CAst.v = GRef (b,_) },[a]) when Globnames.eq_gr b glob_xO -> mult_2(bignat_of_pos a) + | GApp ({ CAst.v = GRef (b,_) },[a]) when Globnames.eq_gr b glob_xI -> add_1(mult_2(bignat_of_pos a)) + | GRef (a, _) when Globnames.eq_gr a glob_xH -> Bigint.one | _ -> raise Non_closed_number + ) x let uninterp_positive p = try @@ -87,9 +88,9 @@ let uninterp_positive p = let _ = Notation.declare_numeral_interpreter "positive_scope" (positive_path,binnums) interp_positive - ([Loc.tag @@ GRef (glob_xI, None); - Loc.tag @@ GRef (glob_xO, None); - Loc.tag @@ GRef (glob_xH, None)], + ([CAst.make @@ GRef (glob_xI, None); + CAst.make @@ GRef (glob_xO, None); + CAst.make @@ GRef (glob_xH, None)], uninterp_positive, true) @@ -106,9 +107,9 @@ let glob_Npos = ConstructRef path_of_Npos let n_path = make_path binnums "N" -let n_of_binnat ?loc pos_or_neg n = Loc.tag ?loc @@ +let n_of_binnat ?loc pos_or_neg n = CAst.make ?loc @@ if not (Bigint.equal n zero) then - GApp(Loc.tag @@ GRef (glob_Npos,None), [pos_of_bignat ?loc n]) + GApp(CAst.make @@ GRef (glob_Npos,None), [pos_of_bignat ?loc n]) else GRef(glob_N0, None) @@ -123,10 +124,11 @@ let n_of_int ?loc n = (* Printing N via scopes *) (**********************************************************************) -let bignat_of_n = function - | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_Npos -> bignat_of_pos a - | _, GRef (a,_) when Globnames.eq_gr a glob_N0 -> Bigint.zero +let bignat_of_n = CAst.with_val (function + | GApp ({ CAst.v = GRef (b,_)},[a]) when Globnames.eq_gr b glob_Npos -> bignat_of_pos a + | GRef (a,_) when Globnames.eq_gr a glob_N0 -> Bigint.zero | _ -> raise Non_closed_number + ) let uninterp_n p = try Some (bignat_of_n p) @@ -138,8 +140,8 @@ let uninterp_n p = let _ = Notation.declare_numeral_interpreter "N_scope" (n_path,binnums) n_of_int - ([Loc.tag @@ GRef (glob_N0, None); - Loc.tag @@ GRef (glob_Npos, None)], + ([CAst.make @@ GRef (glob_N0, None); + CAst.make @@ GRef (glob_Npos, None)], uninterp_n, true) @@ -161,19 +163,20 @@ let z_of_int ?loc n = if not (Bigint.equal n zero) then let sgn, n = if is_pos_or_zero n then glob_POS, n else glob_NEG, Bigint.neg n in - Loc.tag ?loc @@ GApp(Loc.tag ?loc @@ GRef(sgn,None), [pos_of_bignat ?loc n]) + CAst.make ?loc @@ GApp(CAst.make ?loc @@ GRef(sgn,None), [pos_of_bignat ?loc n]) else - Loc.tag ?loc @@ GRef(glob_ZERO, None) + CAst.make ?loc @@ GRef(glob_ZERO, None) (**********************************************************************) (* Printing Z via scopes *) (**********************************************************************) -let bigint_of_z = function - | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_POS -> bignat_of_pos a - | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_NEG -> Bigint.neg (bignat_of_pos a) - | _, GRef (a, _) when Globnames.eq_gr a glob_ZERO -> Bigint.zero +let bigint_of_z = CAst.with_val (function + | GApp ({ CAst.v = GRef (b,_)},[a]) when Globnames.eq_gr b glob_POS -> bignat_of_pos a + | GApp ({ CAst.v = GRef (b,_)},[a]) when Globnames.eq_gr b glob_NEG -> Bigint.neg (bignat_of_pos a) + | GRef (a, _) when Globnames.eq_gr a glob_ZERO -> Bigint.zero | _ -> raise Non_closed_number + ) let uninterp_z p = try @@ -186,8 +189,8 @@ let uninterp_z p = let _ = Notation.declare_numeral_interpreter "Z_scope" (z_path,binnums) z_of_int - ([Loc.tag @@ GRef (glob_ZERO, None); - Loc.tag @@ GRef (glob_POS, None); - Loc.tag @@ GRef (glob_NEG, None)], + ([CAst.make @@ GRef (glob_ZERO, None); + CAst.make @@ GRef (glob_POS, None); + CAst.make @@ GRef (glob_NEG, None)], uninterp_z, true) diff --git a/pretyping/cases.ml b/pretyping/cases.ml index eb0d01718..3beef7773 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -95,7 +95,7 @@ let msg_may_need_inversion () = (* Utils *) let make_anonymous_patvars n = - List.make n (Loc.tag @@ PatVar Anonymous) + List.make n (CAst.make @@ PatVar Anonymous) (* We have x1:t1...xn:tn,xi':ti,y1..yk |- c and re-generalize over xi:ti to get x1:t1...xn:tn,xi':ti,y1..yk |- c[xi:=xi'] *) @@ -178,7 +178,7 @@ and build_glob_pattern args = function | Top -> args | MakeConstructor (pci, rh) -> glob_pattern_of_partial_history - [Loc.tag @@ PatCstr (pci, args, Anonymous)] rh + [CAst.make @@ PatCstr (pci, args, Anonymous)] rh let complete_history = glob_pattern_of_partial_history [] @@ -188,12 +188,12 @@ let pop_history_pattern = function | Continuation (0, l, Top) -> Result (List.rev l) | Continuation (0, l, MakeConstructor (pci, rh)) -> - feed_history (Loc.tag @@ PatCstr (pci,List.rev l,Anonymous)) rh + feed_history (CAst.make @@ PatCstr (pci,List.rev l,Anonymous)) rh | _ -> anomaly (Pp.str "Constructor not yet filled with its arguments") let pop_history h = - feed_history (Loc.tag @@ PatVar Anonymous) h + feed_history (CAst.make @@ PatVar Anonymous) h (* Builds a continuation expecting [n] arguments and building [ci] applied to this [n] arguments *) @@ -273,8 +273,8 @@ type 'a pattern_matching_problem = let rec find_row_ind = function [] -> None - | (_, PatVar _) :: l -> find_row_ind l - | (loc, PatCstr(c,_,_)) :: _ -> Some (loc,c) + | { CAst.v = PatVar _ } :: l -> find_row_ind l + | { CAst.v = PatCstr(c,_,_) ; loc } :: _ -> Some (loc,c) let inductive_template evdref env tmloc ind = let indu = evd_comb1 (Evd.fresh_inductive_instance env) evdref ind in @@ -427,7 +427,7 @@ let current_pattern eqn = | pat::_ -> pat | [] -> anomaly (Pp.str "Empty list of patterns") -let alias_of_pat = Loc.with_loc (fun ?loc -> function +let alias_of_pat = CAst.with_val (function | PatVar name -> name | PatCstr(_,_,name) -> name ) @@ -473,13 +473,13 @@ let rec adjust_local_defs ?loc = function | (pat :: pats, LocalAssum _ :: decls) -> pat :: adjust_local_defs ?loc (pats,decls) | (pats, LocalDef _ :: decls) -> - (Loc.tag ?loc @@ PatVar Anonymous) :: adjust_local_defs ?loc (pats,decls) + (CAst.make ?loc @@ PatVar Anonymous) :: adjust_local_defs ?loc (pats,decls) | [], [] -> [] | _ -> raise NotAdjustable let check_and_adjust_constructor env ind cstrs = function - | _, PatVar _ as pat -> pat - | loc, PatCstr (((_,i) as cstr),args,alias) as pat -> + | { CAst.v = PatVar _ } as pat -> pat + | { CAst.v = PatCstr (((_,i) as cstr),args,alias) ; loc } as pat -> (* Check it is constructor of the right type *) let ind' = inductive_of_constructor cstr in if eq_ind ind' ind then @@ -490,7 +490,7 @@ let check_and_adjust_constructor env ind cstrs = function else try let args' = adjust_local_defs ?loc (args, List.rev ci.cs_args) - in Loc.tag ?loc @@ PatCstr (cstr, args', alias) + in CAst.make ?loc @@ PatCstr (cstr, args', alias) with NotAdjustable -> error_wrong_numarg_constructor ?loc env cstr nb_args_constr else @@ -503,8 +503,8 @@ let check_and_adjust_constructor env ind cstrs = function let check_all_variables env sigma typ mat = List.iter (fun eqn -> match current_pattern eqn with - | _, PatVar id -> () - | loc, PatCstr (cstr_sp,_,_) -> + | { CAst.v = PatVar id } -> () + | { CAst.v = PatCstr (cstr_sp,_,_); loc } -> error_bad_pattern ?loc env sigma cstr_sp typ) mat @@ -530,8 +530,8 @@ let occur_in_rhs na rhs = | Name id -> Id.List.mem id rhs.rhs_vars let is_dep_patt_in eqn = function - | _, PatVar name -> Flags.is_program_mode () || occur_in_rhs name eqn.rhs - | _, PatCstr _ -> true + | { CAst.v = PatVar name } -> Flags.is_program_mode () || occur_in_rhs name eqn.rhs + | { CAst.v = PatCstr _ } -> true let mk_dep_patt_row (pats,_,eqn) = List.map (is_dep_patt_in eqn) pats @@ -751,7 +751,7 @@ let recover_and_adjust_alias_names names sign = | x::names, LocalAssum (_,t)::sign -> (x, LocalAssum (alias_of_pat x,t)) :: aux (names,sign) | names, (LocalDef (na,_,_) as decl)::sign -> - (Loc.tag @@ PatVar na, decl) :: aux (names,sign) + (CAst.make @@ PatVar na, decl) :: aux (names,sign) | _ -> assert false in List.split (aux (names,sign)) @@ -968,7 +968,7 @@ let use_unit_judge evd = evd', j let add_assert_false_case pb tomatch = - let pats = List.map (fun _ -> Loc.tag @@ PatVar Anonymous) tomatch in + let pats = List.map (fun _ -> CAst.make @@ PatVar Anonymous) tomatch in let aliasnames = List.map_filter (function Alias _ | NonDepAlias -> Some Anonymous | _ -> None) tomatch in @@ -1166,8 +1166,8 @@ let postprocess_dependencies evd tocheck brs tomatch pred deps cs = (* Sorting equations by constructor *) let rec irrefutable env = function - | _, PatVar name -> true - | _, PatCstr (cstr,args,_) -> + | { CAst.v = PatVar name } -> true + | { CAst.v = PatCstr (cstr,args,_) } -> let ind = inductive_of_constructor cstr in let (_,mip) = Inductive.lookup_mind_specif env ind in let one_constr = Int.equal (Array.length mip.mind_user_lc) 1 in @@ -1188,14 +1188,14 @@ let group_equations pb ind current cstrs mat = let rest = remove_current_pattern eqn in let pat = current_pattern eqn in match check_and_adjust_constructor pb.env ind cstrs pat with - | _, PatVar name -> + | { CAst.v = PatVar name } -> (* This is a default clause that we expand *) for i=1 to Array.length cstrs do let args = make_anonymous_patvars cstrs.(i-1).cs_nargs in brs.(i-1) <- (args, name, rest) :: brs.(i-1) done; if !only_default == None then only_default := Some true - | loc, PatCstr (((_,i)),args,name) -> + | { CAst.v = PatCstr (((_,i)),args,name) ; loc } -> (* This is a regular clause *) only_default := Some false; brs.(i-1) <- (args, name, rest) :: brs.(i-1)) mat () in @@ -1719,16 +1719,16 @@ let build_tycon ?loc env tycon_env s subst tycon extenv evdref t = let build_inversion_problem loc env sigma tms t = let make_patvar t (subst,avoid) = let id = next_name_away (named_hd env sigma t Anonymous) avoid in - Loc.tag @@ PatVar (Name id), ((id,t)::subst, id::avoid) in + CAst.make @@ PatVar (Name id), ((id,t)::subst, id::avoid) in let rec reveal_pattern t (subst,avoid as acc) = match EConstr.kind sigma (whd_all env sigma t) with - | Construct (cstr,u) -> Loc.tag (PatCstr (cstr,[],Anonymous)), acc + | Construct (cstr,u) -> CAst.make (PatCstr (cstr,[],Anonymous)), acc | App (f,v) when isConstruct sigma f -> let cstr,u = destConstruct sigma f in let n = constructor_nrealargs_env env cstr in let l = List.lastn n (Array.to_list v) in let l,acc = List.fold_map' reveal_pattern l acc in - Loc.tag (PatCstr (cstr,l,Anonymous)), acc + CAst.make (PatCstr (cstr,l,Anonymous)), acc | _ -> make_patvar t acc in let rec aux n env acc_sign tms acc = match tms with @@ -1804,7 +1804,7 @@ let build_inversion_problem loc env sigma tms t = (* No need for a catch all clause *) [] else - [ { patterns = List.map (fun _ -> Loc.tag @@ PatVar Anonymous) patl; + [ { patterns = List.map (fun _ -> CAst.make @@ PatVar Anonymous) patl; alias_stack = []; eqn_loc = None; used = ref false; @@ -2059,13 +2059,14 @@ let mk_JMeq evdref typ x typ' y = let mk_JMeq_refl evdref typ x = papp evdref coq_JMeq_refl [| typ; x |] -let hole = Loc.tag @@ +let hole = CAst.make @@ GHole (Evar_kinds.QuestionMark (Evar_kinds.Define false), Misctypes.IntroAnonymous, None) let constr_of_pat env evdref arsign pat avoid = - let rec typ env (ty, realargs) (loc, pat) avoid = - match pat with + let rec typ env (ty, realargs) pat avoid = + let loc = pat.CAst.loc in + match pat.CAst.v with | PatVar name -> let name, avoid = match name with Name n -> name, avoid @@ -2073,7 +2074,7 @@ let constr_of_pat env evdref arsign pat avoid = let previd, id = prime avoid (Name (Id.of_string "wildcard")) in Name id, id :: avoid in - ((Loc.tag ?loc @@ PatVar name), [LocalAssum (name, ty)] @ realargs, mkRel 1, ty, + ((CAst.make ?loc @@ PatVar name), [LocalAssum (name, ty)] @ realargs, mkRel 1, ty, (List.map (fun x -> mkRel 1) realargs), 1, avoid) | PatCstr (((_, i) as cstr),args,alias) -> let cind = inductive_of_constructor cstr in @@ -2104,7 +2105,7 @@ let constr_of_pat env evdref arsign pat avoid = in let args = List.rev args in let patargs = List.rev patargs in - let pat' = Loc.tag ?loc @@ PatCstr (cstr, patargs, alias) in + let pat' = CAst.make ?loc @@ PatCstr (cstr, patargs, alias) in let cstr = mkConstructU (on_snd EInstance.make ci.cs_cstr) in let app = applist (cstr, List.map (lift (List.length sign)) params) in let app = applist (app, args) in @@ -2160,18 +2161,18 @@ let vars_of_ctx sigma ctx = match decl with | LocalDef (na,t',t) when is_topvar sigma t' -> prev, - (Loc.tag @@ GApp ( - (Loc.tag @@ GRef (delayed_force coq_eq_refl_ref, None)), - [hole; Loc.tag @@ GVar prev])) :: vars + (CAst.make @@ GApp ( + (CAst.make @@ GRef (delayed_force coq_eq_refl_ref, None)), + [hole; CAst.make @@ GVar prev])) :: vars | _ -> match RelDecl.get_name decl with Anonymous -> invalid_arg "vars_of_ctx" - | Name n -> n, (Loc.tag @@ GVar n) :: vars) + | Name n -> n, (CAst.make @@ GVar n) :: vars) ctx (Id.of_string "vars_of_ctx_error", []) in List.rev y -let rec is_included (loc_x, x) (loc_y, y) = - match x, y with +let rec is_included x y = + match CAst.(x.v, y.v) with | PatVar _, _ -> true | _, PatVar _ -> true | PatCstr ((_, i), args, alias), PatCstr ((_, i'), args', alias') -> @@ -2289,13 +2290,13 @@ let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity = let branch_name = Id.of_string ("program_branch_" ^ (string_of_int !i)) in let branch_decl = LocalDef (Name branch_name, lift !i bbody, lift !i btype) in let branch = - let bref = Loc.tag @@ GVar branch_name in + let bref = CAst.make @@ GVar branch_name in match vars_of_ctx !evdref rhs_rels with [] -> bref - | l -> Loc.tag @@ GApp (bref, l) + | l -> CAst.make @@ GApp (bref, l) in let branch = match ineqs with - Some _ -> Loc.tag @@ GApp (branch, [ hole ]) + Some _ -> CAst.make @@ GApp (branch, [ hole ]) | None -> branch in incr i; diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index acccfc125..a93653f32 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -79,8 +79,8 @@ let apply_pattern_coercion ?loc pat p = List.fold_left (fun pat (co,n) -> let f i = - if i pat::hd,rest) lines) clauses) | _ -> - let pat = Loc.tag @@ PatVar(update_name sigma na rhs) in + let pat = CAst.make @@ PatVar(update_name sigma na rhs) in let mat = align_tree nal isgoal rhs sigma in List.map (fun (hd,rest) -> pat::hd,rest) mat @@ -329,7 +329,7 @@ let is_nondep_branch sigma c l = let extract_nondep_branches test c b l = let rec strip l r = - match snd r,l with + match r.CAst.v, l with | r', [] -> r | GLambda (_,_,_,t), false::l -> strip l t | GLetIn (_,_,_,t), true::l -> strip l t @@ -339,7 +339,7 @@ let extract_nondep_branches test c b l = let it_destRLambda_or_LetIn_names l c = let rec aux l nal c = - match snd c, l with + match c.CAst.v, l with | _, [] -> (List.rev nal,c) | GLambda (na,_,_,c), false::l -> aux l (na::nal) c | GLetIn (na,_,_,c), true::l -> aux l (na::nal) c @@ -353,11 +353,11 @@ let it_destRLambda_or_LetIn_names l c = x in let x = next (free_glob_vars c) in - let a = Loc.tag @@ GVar x in + let a = CAst.make @@ GVar x in aux l (Name x :: nal) (match c with - | loc, GApp (p,l) -> (loc, GApp (p,l@[a])) - | _ -> Loc.tag @@ GApp (c,[a])) + | { loc; CAst.v = GApp (p,l) } -> CAst.make ?loc @@ GApp (p,l@[a]) + | _ -> CAst.make @@ GApp (c,[a])) in aux l [] c let detype_case computable detype detype_eqns testdep avoid data p c bl = @@ -373,7 +373,7 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl = | None -> Anonymous, None, None | Some p -> let nl,typ = it_destRLambda_or_LetIn_names k p in - let n,typ = match snd typ with + let n,typ = match typ.CAst.v with | GLambda (x,_,t,c) -> x, c | _ -> Anonymous, typ in let aliastyp = @@ -395,7 +395,7 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl = else st with Not_found -> st - in Loc.tag @@ + in match tag, aliastyp with | LetStyle, None -> let bl' = Array.map detype bl in @@ -440,12 +440,12 @@ let detype_instance sigma l = if Univ.Instance.is_empty l then None else Some (List.map (detype_level sigma) (Array.to_list (Univ.Instance.to_array l))) -let rec detype flags avoid env sigma t = Loc.tag @@ +let rec detype flags avoid env sigma t = CAst.make @@ match EConstr.kind sigma (collapse_appl sigma t) with | Rel n -> (try match lookup_name_of_rel n (fst env) with | Name id -> GVar id - | Anonymous -> snd @@ !detype_anonymous n + | Anonymous -> (!detype_anonymous n).CAst.v with Not_found -> let s = "_UNBOUND_REL_"^(string_of_int n) in GVar (Id.of_string s)) @@ -458,7 +458,7 @@ let rec detype flags avoid env sigma t = Loc.tag @@ with Not_found -> GVar id) | Sort s -> GSort (detype_sort sigma (ESorts.kind sigma s)) | Cast (c1,REVERTcast,c2) when not !Flags.raw_print -> - snd (detype flags avoid env sigma c1) + (detype flags avoid env sigma c1).CAst.v | Cast (c1,k,c2) -> let d1 = detype flags avoid env sigma c1 in let d2 = detype flags avoid env sigma c2 in @@ -468,12 +468,12 @@ let rec detype flags avoid env sigma t = Loc.tag @@ | _ -> CastConv d2 in GCast(d1,cast) - | Prod (na,ty,c) -> snd @@ detype_binder flags BProd avoid env sigma na None ty c - | Lambda (na,ty,c) -> snd @@ detype_binder flags BLambda avoid env sigma na None ty c - | LetIn (na,b,ty,c) -> snd @@ detype_binder flags BLetIn avoid env sigma na (Some b) ty c + | Prod (na,ty,c) -> detype_binder flags BProd avoid env sigma na None ty c + | Lambda (na,ty,c) -> detype_binder flags BLambda avoid env sigma na None ty c + | LetIn (na,b,ty,c) -> detype_binder flags BLetIn avoid env sigma na (Some b) ty c | App (f,args) -> let mkapp f' args' = - match snd f' with + match f'.CAst.v with | GApp (f',args'') -> GApp (f',args''@args') | _ -> GApp (f',args') @@ -485,16 +485,16 @@ let rec detype flags avoid env sigma t = Loc.tag @@ let noparams () = let pb = Environ.lookup_projection p (snd env) in let pars = pb.Declarations.proj_npars in - let hole = Loc.tag @@ GHole(Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None) in + let hole = CAst.make @@ GHole(Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None) in let args = List.make pars hole in - GApp (Loc.tag @@ GRef (ConstRef (Projection.constant p), None), + GApp (CAst.make @@ GRef (ConstRef (Projection.constant p), None), (args @ [detype flags avoid env sigma c])) in if fst flags || !Flags.in_debugger || !Flags.in_toplevel then try noparams () with _ -> (* lax mode, used by debug printers only *) - GApp (Loc.tag @@ GRef (ConstRef (Projection.constant p), None), + GApp (CAst.make @@ GRef (ConstRef (Projection.constant p), None), [detype flags avoid env sigma c]) else if print_primproj_compatibility () && Projection.unfolded p then @@ -512,12 +512,12 @@ let rec detype flags avoid env sigma t = Loc.tag @@ substl (c :: List.rev args) body' with Retyping.RetypeError _ | Not_found -> anomaly (str"Cannot detype an unfolded primitive projection.") - in snd @@ detype flags avoid env sigma c' + in (detype flags avoid env sigma c').CAst.v else if print_primproj_params () then try let c = Retyping.expand_projection (snd env) sigma p c [] in - snd @@ detype flags avoid env sigma c + (detype flags avoid env sigma c).CAst.v with Retyping.RetypeError _ -> noparams () else noparams () @@ -552,7 +552,6 @@ let rec detype flags avoid env sigma t = Loc.tag @@ GRef (ConstructRef cstr_sp, detype_instance sigma u) | Case (ci,p,c,bl) -> let comp = computable sigma p (List.length (ci.ci_pp_info.ind_tags)) in - snd @@ detype_case comp (detype flags avoid env sigma) (detype_eqns flags avoid env sigma ci comp) (is_nondep_branch sigma) avoid @@ -643,17 +642,17 @@ and detype_eqns flags avoid env sigma ci computable constructs consnargsl bl = and detype_eqn (lax,isgoal as flags) avoid env sigma constr construct_nargs branch = let make_pat x avoid env b body ty ids = if force_wildcard () && noccurn sigma 1 b then - Loc.tag @@ PatVar Anonymous,avoid,(add_name Anonymous body ty env),ids + CAst.make @@ PatVar Anonymous,avoid,(add_name Anonymous body ty env),ids else let flag = if isgoal then RenamingForGoal else RenamingForCasesPattern (fst env,b) in let na,avoid' = compute_displayed_name_in sigma flag avoid x b in - Loc.tag @@ PatVar na,avoid',(add_name na body ty env),add_vname ids na + CAst.make (PatVar na),avoid',(add_name na body ty env),add_vname ids na in let rec buildrec ids patlist avoid env l b = match EConstr.kind sigma b, l with | _, [] -> Loc.tag @@ (Id.Set.elements ids, - [Loc.tag @@ PatCstr(constr, List.rev patlist,Anonymous)], + [CAst.make @@ PatCstr(constr, List.rev patlist,Anonymous)], detype flags avoid env sigma b) | Lambda (x,t,b), false::l -> let pat,new_avoid,new_env,new_ids = make_pat x avoid env b None t ids in @@ -667,7 +666,7 @@ and detype_eqn (lax,isgoal as flags) avoid env sigma constr construct_nargs bran buildrec ids patlist avoid env l c | _, true::l -> - let pat = Loc.tag @@ PatVar Anonymous in + let pat = CAst.make @@ PatVar Anonymous in buildrec ids (pat::patlist) avoid env l b | _, false::l -> @@ -682,7 +681,7 @@ and detype_eqn (lax,isgoal as flags) avoid env sigma constr construct_nargs bran in buildrec Id.Set.empty [] avoid env construct_nargs branch -and detype_binder (lax,isgoal as flags) bk avoid env sigma na body ty c = Loc.tag @@ +and detype_binder (lax,isgoal as flags) bk avoid env sigma na body ty c = let flag = if isgoal then RenamingForGoal else RenamingElsewhereFor (fst env,c) in let na',avoid' = match bk with | BLetIn -> compute_displayed_let_name_in sigma flag avoid na c @@ -740,7 +739,7 @@ let detype_closed_glob ?lax isgoal avoid env sigma t = | Name id -> Name (convert_id cl id) | Anonymous -> Anonymous in - let rec detype_closed_glob cl cg = Loc.map (function + let rec detype_closed_glob cl cg : Glob_term.glob_constr = CAst.map (function | GVar id -> (* if [id] is bound to a name. *) begin try @@ -754,11 +753,11 @@ let detype_closed_glob ?lax isgoal avoid env sigma t = [Printer.pr_constr_under_binders_env] does. *) let assums = List.map (fun id -> LocalAssum (Name id,(* dummy *) mkProp)) b in let env = push_rel_context assums env in - snd @@ detype ?lax isgoal avoid env sigma c + (detype ?lax isgoal avoid env sigma c).CAst.v (* if [id] is bound to a [closed_glob_constr]. *) with Not_found -> try let {closure;term} = Id.Map.find id cl.untyped in - snd @@ detype_closed_glob closure term + (detype_closed_glob closure term).CAst.v (* Otherwise [id] stands for itself *) with Not_found -> GVar id @@ -785,7 +784,7 @@ let detype_closed_glob ?lax isgoal avoid env sigma t = in GCases(sty,po,tml,eqns) | c -> - snd @@ Glob_ops.map_glob_constr (detype_closed_glob cl) cg + (Glob_ops.map_glob_constr (detype_closed_glob cl) cg).CAst.v ) cg in detype_closed_glob t.closure t.term @@ -793,52 +792,52 @@ let detype_closed_glob ?lax isgoal avoid env sigma t = (**********************************************************************) (* Module substitution: relies on detyping *) -let rec subst_cases_pattern subst (loc, pat) = Loc.tag ?loc @@ - match pat with - | PatVar _ -> pat - | PatCstr (((kn,i),j),cpl,n) -> +let rec subst_cases_pattern subst = CAst.map (function + | PatVar _ as pat -> pat + | PatCstr (((kn,i),j),cpl,n) as pat -> let kn' = subst_mind subst kn and cpl' = List.smartmap (subst_cases_pattern subst) cpl in if kn' == kn && cpl' == cpl then pat else PatCstr (((kn',i),j),cpl',n) + ) let (f_subst_genarg, subst_genarg_hook) = Hook.make () -let rec subst_glob_constr subst (loc, raw) = Loc.tag ?loc @@ - match raw with - | GRef (ref,u) -> +let rec subst_glob_constr subst = CAst.map (function + | GRef (ref,u) as raw -> let ref',t = subst_global subst ref in if ref' == ref then raw else - snd @@ detype false [] (Global.env()) Evd.empty (EConstr.of_constr t) + (detype false [] (Global.env()) Evd.empty (EConstr.of_constr t)).CAst.v - | GVar _ -> raw - | GEvar _ -> raw - | GPatVar _ -> raw + | GSort _ + | GVar _ + | GEvar _ + | GPatVar _ as raw -> raw - | GApp (r,rl) -> + | GApp (r,rl) as raw -> let r' = subst_glob_constr subst r and rl' = List.smartmap (subst_glob_constr subst) rl in if r' == r && rl' == rl then raw else GApp(r',rl') - | GLambda (n,bk,r1,r2) -> + | GLambda (n,bk,r1,r2) as raw -> let r1' = subst_glob_constr subst r1 and r2' = subst_glob_constr subst r2 in if r1' == r1 && r2' == r2 then raw else GLambda (n,bk,r1',r2') - | GProd (n,bk,r1,r2) -> + | GProd (n,bk,r1,r2) as raw -> let r1' = subst_glob_constr subst r1 and r2' = subst_glob_constr subst r2 in if r1' == r1 && r2' == r2 then raw else GProd (n,bk,r1',r2') - | GLetIn (n,r1,t,r2) -> + | GLetIn (n,r1,t,r2) as raw -> let r1' = subst_glob_constr subst r1 in let r2' = subst_glob_constr subst r2 in let t' = Option.smartmap (subst_glob_constr subst) t in if r1' == r1 && t == t' && r2' == r2 then raw else GLetIn (n,r1',t',r2') - | GCases (sty,rtno,rl,branches) -> + | GCases (sty,rtno,rl,branches) as raw -> let rtno' = Option.smartmap (subst_glob_constr subst) rtno and rl' = List.smartmap (fun (a,x as y) -> let a' = subst_glob_constr subst a in @@ -860,14 +859,14 @@ let rec subst_glob_constr subst (loc, raw) = Loc.tag ?loc @@ if rtno' == rtno && rl' == rl && branches' == branches then raw else GCases (sty,rtno',rl',branches') - | GLetTuple (nal,(na,po),b,c) -> + | GLetTuple (nal,(na,po),b,c) as raw -> let po' = Option.smartmap (subst_glob_constr subst) po and b' = subst_glob_constr subst b and c' = subst_glob_constr subst c in if po' == po && b' == b && c' == c then raw else GLetTuple (nal,(na,po'),b',c') - | GIf (c,(na,po),b1,b2) -> + | GIf (c,(na,po),b1,b2) as raw -> let po' = Option.smartmap (subst_glob_constr subst) po and b1' = subst_glob_constr subst b1 and b2' = subst_glob_constr subst b2 @@ -875,7 +874,7 @@ let rec subst_glob_constr subst (loc, raw) = Loc.tag ?loc @@ if c' == c && po' == po && b1' == b1 && b2' == b2 then raw else GIf (c',(na,po'),b1',b2') - | GRec (fix,ida,bl,ra1,ra2) -> + | GRec (fix,ida,bl,ra1,ra2) as raw -> let ra1' = Array.smartmap (subst_glob_constr subst) ra1 and ra2' = Array.smartmap (subst_glob_constr subst) ra2 in let bl' = Array.smartmap @@ -887,9 +886,7 @@ let rec subst_glob_constr subst (loc, raw) = Loc.tag ?loc @@ if ra1' == ra1 && ra2' == ra2 && bl'==bl then raw else GRec (fix,ida,bl',ra1',ra2') - | GSort _ -> raw - - | GHole (knd, naming, solve) -> + | GHole (knd, naming, solve) as raw -> let nknd = match knd with | Evar_kinds.ImplicitArg (ref, i, b) -> let nref, _ = subst_global subst ref in @@ -900,18 +897,19 @@ let rec subst_glob_constr subst (loc, raw) = Loc.tag ?loc @@ if nsolve == solve && nknd == knd then raw else GHole (nknd, naming, nsolve) - | GCast (r1,k) -> + | GCast (r1,k) as raw -> let r1' = subst_glob_constr subst r1 in let k' = Miscops.smartmap_cast_type (subst_glob_constr subst) k in if r1' == r1 && k' == k then raw else GCast (r1',k') + ) (* Utilities to transform kernel cases to simple pattern-matching problem *) let simple_cases_matrix_of_branches ind brs = List.map (fun (i,n,b) -> let nal,c = it_destRLambda_or_LetIn_names n b in - let mkPatVar na = Loc.tag @@ PatVar na in - let p = Loc.tag @@ PatCstr ((ind,i+1),List.map mkPatVar nal,Anonymous) in + let mkPatVar na = CAst.make @@ PatVar na in + let p = CAst.make @@ PatCstr ((ind,i+1),List.map mkPatVar nal,Anonymous) in let map name = try Some (Nameops.out_name name) with Failure _ -> None in let ids = List.map_filter map nal in Loc.tag @@ (ids,[p],c)) diff --git a/pretyping/detyping.mli b/pretyping/detyping.mli index 84da3652f..da287ae9f 100644 --- a/pretyping/detyping.mli +++ b/pretyping/detyping.mli @@ -35,14 +35,6 @@ val detype_names : bool -> Id.t list -> names_context -> env -> evar_map -> cons val detype : ?lax:bool -> bool -> Id.t list -> env -> evar_map -> constr -> glob_constr -val detype_case : - bool -> (constr -> glob_constr) -> - (constructor array -> bool list array -> constr array -> - (Id.t list * cases_pattern list * glob_constr) Loc.located list) -> - (constr -> bool list -> bool) -> - Id.t list -> inductive * case_style * bool list array * bool list -> - constr option -> constr -> constr array -> glob_constr - val detype_sort : evar_map -> sorts -> glob_sort val detype_rel_context : ?lax:bool -> constr option -> Id.t list -> (names_context * env) -> diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml index 7e6b063d0..94bc24e3c 100644 --- a/pretyping/glob_ops.ml +++ b/pretyping/glob_ops.ml @@ -15,15 +15,15 @@ open Glob_term (* Untyped intermediate terms, after ASTs and before constr. *) -let cases_pattern_loc (loc, _) = loc +let cases_pattern_loc c = c.CAst.loc let cases_predicate_names tml = List.flatten (List.map (function | (tm,(na,None)) -> [na] | (tm,(na,Some (_,(_,nal)))) -> na::nal) tml) -let mkGApp ?loc p t = Loc.tag ?loc @@ - match snd p with +let mkGApp ?loc p t = CAst.make ?loc @@ + match p.CAst.v with | GApp (f,l) -> GApp (f,l@[t]) | _ -> GApp (p,[t]) @@ -45,7 +45,7 @@ let case_style_eq s1 s2 = match s1, s2 with | RegularStyle, RegularStyle -> true | _ -> false -let rec cases_pattern_eq (_loc1, p1) (_loc2, p2) = match p1, p2 with +let rec cases_pattern_eq { CAst.v = p1} { CAst.v = p2 } = match p1, p2 with | PatVar na1, PatVar na2 -> Name.equal na1 na2 | PatCstr (c1, pl1, na1), PatCstr (c2, pl2, na2) -> eq_constructor c1 c2 && List.equal cases_pattern_eq pl1 pl2 && @@ -59,7 +59,7 @@ let cast_type_eq eq t1 t2 = match t1, t2 with | CastNative t1, CastNative t2 -> eq t1 t2 | _ -> false -let rec glob_constr_eq (_loc1, c1) (_loc2, c2) = match c1, c2 with +let rec glob_constr_eq { CAst.v = c1 } { CAst.v = c2 } = match c1, c2 with | GRef (gr1, _), GRef (gr2, _) -> eq_gr gr1 gr2 | GVar id1, GVar id2 -> Id.equal id1 id2 | GEvar (id1, arg1), GEvar (id2, arg2) -> @@ -137,7 +137,7 @@ and fix_recursion_order_eq o1 o2 = match o1, o2 with and instance_eq (x1,c1) (x2,c2) = Id.equal x1 x2 && glob_constr_eq c1 c2 -let map_glob_constr_left_to_right f = Loc.map (function +let map_glob_constr_left_to_right f = CAst.map (function | GApp (g,args) -> let comp1 = f g in let comp2 = Util.List.map_left f args in @@ -186,7 +186,7 @@ let map_glob_constr = map_glob_constr_left_to_right let fold_return_type f acc (na,tyopt) = Option.fold_left f acc tyopt -let fold_glob_constr f acc = Loc.with_unloc (function +let fold_glob_constr f acc = CAst.with_val (function | GVar _ -> acc | GApp (c,args) -> List.fold_left f (f acc c) args | GLambda (_,_,b,c) | GProd (_,_,b,c) -> @@ -221,7 +221,7 @@ let same_id na id = match na with | Name id' -> Id.equal id id' let occur_glob_constr id = - let rec occur gt = Loc.with_unloc (function + let rec occur gt = CAst.with_val (function | GVar (id') -> Id.equal id id' | GApp (f,args) -> (occur f) || (List.exists occur args) | GLambda (na,bk,ty,c) -> @@ -270,7 +270,7 @@ let add_name_to_ids set na = | Name id -> Id.Set.add id set let free_glob_vars = - let rec vars bounded vs = Loc.with_unloc @@ (function + let rec vars bounded vs = CAst.with_val @@ (function | GVar (id') -> if Id.Set.mem id' bounded then vs else Id.Set.add id' vs | GApp (f,args) -> List.fold_left (vars bounded) vs (f::args) | GLambda (na,_,ty,c) | GProd (na,_,ty,c) -> @@ -335,7 +335,7 @@ let free_glob_vars = let glob_visible_short_qualid c = let rec aux acc = function - | _, GRef (c,_) -> + | { CAst.v = GRef (c,_) } -> let qualid = Nametab.shortest_qualid_of_global Id.Set.empty c in let dir,id = Libnames.repr_qualid qualid in if DirPath.is_empty dir then id :: acc else acc @@ -354,10 +354,10 @@ let add_and_check_ident id set = Id.Set.add id set let bound_glob_vars = - let rec vars bound = Loc.with_loc (fun ?loc -> function - | GLambda (na,_,_,_) | GProd (na,_,_,_) | GLetIn (na,_,_,_) as c -> + let rec vars bound c = match c.CAst.v with + | GLambda (na,_,_,_) | GProd (na,_,_,_) | GLetIn (na,_,_,_) -> let bound = name_fold add_and_check_ident na bound in - fold_glob_constr vars bound (loc, c) + fold_glob_constr vars bound c | GCases (sty,rtntypopt,tml,pl) -> let bound = vars_option bound rtntypopt in let bound = @@ -391,8 +391,7 @@ let bound_glob_vars = in Array.fold_left_i vars_fix bound idl | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GVar _) -> bound - | GApp _ | GCast _ as c -> fold_glob_constr vars bound (loc, c) - ) + | GApp _ | GCast _ -> fold_glob_constr vars bound c and vars_pattern bound (loc,(idl,p,c)) = let bound = List.fold_right add_and_check_ident idl bound in @@ -425,7 +424,7 @@ let map_tomatch_binders f ((c,(na,inp)) as x) : tomatch_tuple = if r == inp then x else c,(f na, r) -let rec map_case_pattern_binders f = Loc.map (function +let rec map_case_pattern_binders f = CAst.map (function | PatVar na as x -> let r = f na in if r == na then x @@ -463,7 +462,7 @@ let map_pattern f tomatch branches = List.map (fun tm -> map_tomatch f tm) tomatch, List.map (fun br -> map_cases_branch f br) branches -let loc_of_glob_constr (loc, _) = loc +let loc_of_glob_constr c = c.CAst.loc (**********************************************************************) (* Alpha-renaming *) @@ -495,7 +494,7 @@ let rename_var l id = if List.exists (fun (_,id') -> Id.equal id id') l then raise UnsoundRenaming else id -let rec rename_glob_vars l = Loc.map_with_loc (fun ?loc -> function +let rec rename_glob_vars l c = CAst.map_with_loc (fun ?loc -> function | GVar id as r -> let id' = rename_var l id in if id == id' then r else GVar id' @@ -535,14 +534,13 @@ let rec rename_glob_vars l = Loc.map_with_loc (fun ?loc -> function test_na l na; (na,k,Option.map (rename_glob_vars l) bbd,rename_glob_vars l bty))) decls, Array.map (rename_glob_vars l) bs, Array.map (rename_glob_vars l) ts) - (* XXX: This located use case should be improved. *) - | r -> snd @@ map_glob_constr (rename_glob_vars l) (loc, r) - ) + | _ -> (map_glob_constr (rename_glob_vars l) c).CAst.v + ) c (**********************************************************************) (* Conversion from glob_constr to cases pattern, if possible *) -let rec cases_pattern_of_glob_constr na = Loc.map (function +let rec cases_pattern_of_glob_constr na = CAst.map (function | GVar id -> begin match na with | Name _ -> @@ -552,22 +550,22 @@ let rec cases_pattern_of_glob_constr na = Loc.map (function end | GHole (_,_,_) -> PatVar na | GRef (ConstructRef cstr,_) -> PatCstr (cstr,[],na) - | GApp ((_loc, GRef (ConstructRef cstr,_)),l) -> + | GApp ( { CAst.v = GRef (ConstructRef cstr,_) }, l) -> PatCstr (cstr,List.map (cases_pattern_of_glob_constr Anonymous) l,na) | _ -> raise Not_found ) (* Turn a closed cases pattern into a glob_constr *) -let rec glob_constr_of_closed_cases_pattern_aux x = Loc.map_with_loc (fun ?loc -> function +let rec glob_constr_of_closed_cases_pattern_aux x = CAst.map_with_loc (fun ?loc -> function | PatCstr (cstr,[],Anonymous) -> GRef (ConstructRef cstr,None) | PatCstr (cstr,l,Anonymous) -> - let ref = Loc.tag ?loc @@ GRef (ConstructRef cstr,None) in + let ref = CAst.make ?loc @@ GRef (ConstructRef cstr,None) in GApp (ref, List.map glob_constr_of_closed_cases_pattern_aux l) | _ -> raise Not_found ) x let glob_constr_of_closed_cases_pattern = function - | loc, PatCstr (cstr,l,na) -> - na,glob_constr_of_closed_cases_pattern_aux (loc, PatCstr (cstr,l,Anonymous)) + | { CAst.loc ; v = PatCstr (cstr,l,na) } -> + na,glob_constr_of_closed_cases_pattern_aux (CAst.make ?loc @@ PatCstr (cstr,l,Anonymous)) | _ -> raise Not_found diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index 84d846afd..1884b6927 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -324,7 +324,7 @@ let warn_cast_in_pattern = CWarnings.create ~name:"cast-in-pattern" ~category:"automation" (fun () -> Pp.strbrk "Casts are ignored in patterns") -let rec pat_of_raw metas vars = Loc.with_loc (fun ?loc -> function +let rec pat_of_raw metas vars = CAst.with_loc_val (fun ?loc -> function | GVar id -> (try PRel (List.index Name.equal (Name id) vars) with Not_found -> PVar id) @@ -333,7 +333,7 @@ let rec pat_of_raw metas vars = Loc.with_loc (fun ?loc -> function | GRef (gr,_) -> PRef (canonical_gr gr) (* Hack to avoid rewriting a complete interpretation of patterns *) - | GApp ((_, GPatVar (true,n)), cl) -> + | GApp ({ CAst.v = GPatVar (true,n) }, cl) -> metas := n::!metas; PSoApp (n, List.map (pat_of_raw metas vars) cl) | GApp (c,cl) -> PApp (pat_of_raw metas vars c, @@ -362,8 +362,8 @@ let rec pat_of_raw metas vars = Loc.with_loc (fun ?loc -> function PIf (pat_of_raw metas vars c, pat_of_raw metas vars b1,pat_of_raw metas vars b2) | GLetTuple (nal,(_,None),b,c) -> - let mkGLambda c na = Loc.tag ?loc @@ - GLambda (na,Explicit, Loc.tag @@ GHole (Evar_kinds.InternalHole, IntroAnonymous, None),c) in + let mkGLambda c na = CAst.make ?loc @@ + GLambda (na,Explicit, CAst.make @@ GHole (Evar_kinds.InternalHole, IntroAnonymous, None),c) in let c = List.fold_left mkGLambda c nal in let cip = { cip_style = LetStyle; @@ -376,7 +376,7 @@ let rec pat_of_raw metas vars = Loc.with_loc (fun ?loc -> function [0,tags,pat_of_raw metas vars c]) | GCases (sty,p,[c,(na,indnames)],brs) -> let get_ind = function - | (_,(_,[_, PatCstr((ind,_),_,_)],_))::_ -> Some ind + | (_,(_,[{ CAst.v = PatCstr((ind,_),_,_) }],_))::_ -> Some ind | _ -> None in let ind_tags,ind = match indnames with @@ -389,7 +389,7 @@ let rec pat_of_raw metas vars = Loc.with_loc (fun ?loc -> function | Some p, Some (_,(_,nal)) -> let nvars = na :: List.rev nal @ vars in rev_it_mkPLambda nal (mkPLambda na (pat_of_raw metas nvars p)) - | (None | Some (_, GHole _)), _ -> PMeta None + | (None | Some { CAst.v = GHole _}), _ -> PMeta None | Some p, None -> user_err ?loc (strbrk "Clause \"in\" expected in patterns over \"match\" expressions with an explicit \"return\" clause.") in @@ -409,15 +409,15 @@ let rec pat_of_raw metas vars = Loc.with_loc (fun ?loc -> function and pats_of_glob_branches loc metas vars ind brs = let get_arg = function - | _, PatVar na -> + | { CAst.v = PatVar na } -> name_iter (fun n -> metas := n::!metas) na; na - | loc, PatCstr(_,_,_) -> err ?loc (Pp.str "Non supported pattern.") + | { CAst.v = PatCstr(_,_,_) ; loc } -> err ?loc (Pp.str "Non supported pattern.") in let rec get_pat indexes = function | [] -> false, [] - | [(_,(_,[_, PatVar(Anonymous)],(_,GHole _)))] -> true, [] (* ends with _ => _ *) - | (_,(_,[_, PatCstr((indsp,j),lv,_)],br)) :: brs -> + | [(_,(_,[{ CAst.v = PatVar Anonymous }], { CAst.v = GHole _}))] -> true, [] (* ends with _ => _ *) + | (_,(_,[{ CAst.v = PatCstr((indsp,j),lv,_) }],br)) :: brs -> let () = match ind with | Some sp when eq_ind sp indsp -> () | _ -> diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index a256eaa5d..76df89a26 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -567,12 +567,13 @@ let (f_genarg_interp, genarg_interp_hook) = Hook.make () (* in environment [env], with existential variables [evdref] and *) (* the type constraint tycon *) -let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdref (lvar : ltac_var_map) (loc, t) = +let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdref (lvar : ltac_var_map) t = let inh_conv_coerce_to_tycon ?loc = inh_conv_coerce_to_tycon ?loc resolve_tc in let pretype_type = pretype_type k0 resolve_tc in let pretype = pretype k0 resolve_tc in let open Context.Rel.Declaration in - match t with + let loc = t.CAst.loc in + match t.CAst.v with | GRef (ref,u) -> inh_conv_coerce_to_tycon ?loc env evdref (pretype_ref ?loc evdref env ref u) @@ -1097,7 +1098,7 @@ and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update = (* [pretype_type valcon env evdref lvar c] coerces [c] into a type *) and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function - | loc, GHole (knd, naming, None) -> + | { loc; CAst.v = GHole (knd, naming, None) } -> let rec is_Type c = match EConstr.kind !evdref c with | Sort s -> begin match ESorts.kind !evdref s with diff --git a/printing/ppvernac.ml b/printing/ppvernac.ml index 8928d83b0..fded0a60b 100644 --- a/printing/ppvernac.ml +++ b/printing/ppvernac.ml @@ -201,19 +201,19 @@ open Decl_kinds keyword "Module" ++ spc() ++ pr_lfqid id ++ str" := " ++ pr_located pr_qualid qid - let rec pr_module_ast leading_space pr_c = function - | loc, CMident qid -> + let rec pr_module_ast leading_space pr_c = let open CAst in function + | { loc ; v = CMident qid } -> if leading_space then spc () ++ pr_located pr_qualid (loc, qid) else pr_located pr_qualid (loc,qid) - | _loc, CMwith (mty,decl) -> + | { v = CMwith (mty,decl) } -> let m = pr_module_ast leading_space pr_c mty in let p = pr_with_declaration pr_c decl in m ++ spc() ++ keyword "with" ++ spc() ++ p - | _loc, CMapply (me1,(_, CMident _ as me2)) -> + | { v = CMapply (me1, ( { v = CMident _ } as me2 ) ) } -> pr_module_ast leading_space pr_c me1 ++ spc() ++ pr_module_ast false pr_c me2 - | _loc, CMapply (me1,me2) -> + | { v = CMapply (me1,me2) } -> pr_module_ast leading_space pr_c me1 ++ spc() ++ hov 1 (str"(" ++ pr_module_ast false pr_c me2 ++ str")") diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index 2c331ba56..309691e39 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -253,16 +253,16 @@ open Decl_kinds open Evar_kinds let mkPattern c = snd (Patternops.pattern_of_glob_constr c) -let mkGApp f args = Loc.tag @@ GApp (f, args) -let mkGHole = Loc.tag @@ +let mkGApp f args = CAst.make @@ GApp (f, args) +let mkGHole = CAst.make @@ GHole (QuestionMark (Define false), Misctypes.IntroAnonymous, None) -let mkGProd id c1 c2 = Loc.tag @@ +let mkGProd id c1 c2 = CAst.make @@ GProd (Name (Id.of_string id), Explicit, c1, c2) -let mkGArrow c1 c2 = Loc.tag @@ +let mkGArrow c1 c2 = CAst.make @@ GProd (Anonymous, Explicit, c1, c2) -let mkGVar id = Loc.tag @@ GVar (Id.of_string id) -let mkGPatVar id = Loc.tag @@ GPatVar((false, Id.of_string id)) -let mkGRef r = Loc.tag @@ GRef (Lazy.force r, None) +let mkGVar id = CAst.make @@ GVar (Id.of_string id) +let mkGPatVar id = CAst.make @@ GPatVar((false, Id.of_string id)) +let mkGRef r = CAst.make @@ GRef (Lazy.force r, None) let mkGAppRef r args = mkGApp (mkGRef r) args (** forall x : _, _ x x *) diff --git a/vernac/command.ml b/vernac/command.ml index 12df344c2..cae33f316 100644 --- a/vernac/command.ml +++ b/vernac/command.ml @@ -422,7 +422,7 @@ let prepare_param = function let rec check_anonymous_type ind = let open Glob_term in - match snd ind with + match ind.CAst.v with | GSort (GType []) -> true | GProd ( _, _, _, e) | GLetIn (_, _, _, e) -- cgit v1.2.3 From 209956322367e5a4a4c8c78c053ea9352a9a16c8 Mon Sep 17 00:00:00 2001 From: Matej Košík Date: Fri, 28 Apr 2017 14:31:14 +0200 Subject: [location] Renaming "CAst.ast" to "CAst.t" --- interp/constrintern.ml | 2 +- intf/constrexpr.mli | 6 +++--- intf/glob_term.mli | 6 +++--- lib/cAst.ml | 2 +- lib/cAst.mli | 14 +++++++------- 5 files changed, 15 insertions(+), 15 deletions(-) (limited to 'intf') diff --git a/interp/constrintern.ml b/interp/constrintern.ml index 7dc364e5d..3b3dccc99 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -893,7 +893,7 @@ type raw_cases_pattern_expr_r = (** [RCPatCstr (loc, c, l1, l2)] represents ((@c l1) l2) *) | RCPatAtom of Id.t option | RCPatOr of raw_cases_pattern_expr list -and raw_cases_pattern_expr = raw_cases_pattern_expr_r CAst.ast +and raw_cases_pattern_expr = raw_cases_pattern_expr_r CAst.t (** {6 Elementary bricks } *) let apply_scope_env env = function diff --git a/intf/constrexpr.mli b/intf/constrexpr.mli index d54ad8bee..614c097b5 100644 --- a/intf/constrexpr.mli +++ b/intf/constrexpr.mli @@ -52,7 +52,7 @@ type cases_pattern_expr_r = | CPatRecord of (reference * cases_pattern_expr) list | CPatDelimiters of string * cases_pattern_expr | CPatCast of cases_pattern_expr * constr_expr -and cases_pattern_expr = cases_pattern_expr_r CAst.ast +and cases_pattern_expr = cases_pattern_expr_r CAst.t and cases_pattern_notation_substitution = cases_pattern_expr list * (** for constr subterms *) @@ -89,7 +89,7 @@ and constr_expr_r = | CGeneralization of binding_kind * abstraction_kind option * constr_expr | CPrim of prim_token | CDelimiters of string * constr_expr -and constr_expr = constr_expr_r CAst.ast +and constr_expr = constr_expr_r CAst.t and case_expr = constr_expr (* expression that is being matched *) * Name.t Loc.located option (* as-clause *) @@ -140,4 +140,4 @@ type module_ast_r = | CMident of qualid | CMapply of module_ast * module_ast | CMwith of module_ast * with_declaration_ast -and module_ast = module_ast_r CAst.ast +and module_ast = module_ast_r CAst.t diff --git a/intf/glob_term.mli b/intf/glob_term.mli index aefccd518..33c71884a 100644 --- a/intf/glob_term.mli +++ b/intf/glob_term.mli @@ -28,7 +28,7 @@ type cases_pattern_r = | PatVar of Name.t | PatCstr of constructor * cases_pattern list * Name.t (** [PatCstr(p,C,l,x)] = "|'C' 'l' as 'x'" *) -and cases_pattern = cases_pattern_r CAst.ast +and cases_pattern = cases_pattern_r CAst.t (** Representation of an internalized (or in other words globalized) term. *) type glob_constr_r = @@ -53,7 +53,7 @@ type glob_constr_r = | GSort of glob_sort | GHole of Evar_kinds.t * intro_pattern_naming_expr * Genarg.glob_generic_argument option | GCast of glob_constr * glob_constr cast_type -and glob_constr = glob_constr_r CAst.ast +and glob_constr = glob_constr_r CAst.t and glob_decl = Name.t * binding_kind * glob_constr option * glob_constr @@ -83,7 +83,7 @@ type extended_glob_local_binder_r = | GLocalAssum of Name.t * binding_kind * glob_constr | GLocalDef of Name.t * binding_kind * glob_constr * glob_constr option | GLocalPattern of (cases_pattern * Id.t list) * Id.t * binding_kind * glob_constr -and extended_glob_local_binder = extended_glob_local_binder_r CAst.ast +and extended_glob_local_binder = extended_glob_local_binder_r CAst.t (** A globalised term together with a closure representing the value of its free variables. Intended for use when these variables are taken diff --git a/lib/cAst.ml b/lib/cAst.ml index f0a405776..301a6bac8 100644 --- a/lib/cAst.ml +++ b/lib/cAst.ml @@ -7,7 +7,7 @@ (************************************************************************) (** The ast type contains generic metadata for AST nodes. *) -type 'a ast = { +type 'a t = { v : 'a; loc : Loc.t option; } diff --git a/lib/cAst.mli b/lib/cAst.mli index 291536d12..700a06ce8 100644 --- a/lib/cAst.mli +++ b/lib/cAst.mli @@ -7,16 +7,16 @@ (************************************************************************) (** The ast type contains generic metadata for AST nodes *) -type 'a ast = private { +type 'a t = private { v : 'a; loc : Loc.t option; } -val make : ?loc:Loc.t -> 'a -> 'a ast +val make : ?loc:Loc.t -> 'a -> 'a t -val map : ('a -> 'b) -> 'a ast -> 'b ast -val map_with_loc : (?loc:Loc.t -> 'a -> 'b) -> 'a ast -> 'b ast -val map_from_loc : (?loc:Loc.t -> 'a -> 'b) -> 'a Loc.located -> 'b ast +val map : ('a -> 'b) -> 'a t -> 'b t +val map_with_loc : (?loc:Loc.t -> 'a -> 'b) -> 'a t -> 'b t +val map_from_loc : (?loc:Loc.t -> 'a -> 'b) -> 'a Loc.located -> 'b t -val with_val : ('a -> 'b) -> 'a ast -> 'b -val with_loc_val : (?loc:Loc.t -> 'a -> 'b) -> 'a ast -> 'b +val with_val : ('a -> 'b) -> 'a t -> 'b +val with_loc_val : (?loc:Loc.t -> 'a -> 'b) -> 'a t -> 'b -- cgit v1.2.3