summaryrefslogtreecommitdiff
path: root/interp/constrintern.ml
diff options
context:
space:
mode:
Diffstat (limited to 'interp/constrintern.ml')
-rw-r--r--interp/constrintern.ml1780
1 files changed, 1018 insertions, 762 deletions
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index e6340646..83ace93c 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -1,14 +1,17 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Pp
open CErrors
open Util
+open CAst
open Names
open Nameops
open Namegen
@@ -24,7 +27,6 @@ open Constrexpr
open Constrexpr_ops
open Notation_term
open Notation_ops
-open Topconstr
open Nametab
open Notation
open Inductiveops
@@ -45,7 +47,7 @@ open Context.Rel.Declaration
types and recursive definitions and of projection names in records *)
type var_internalization_type =
- | Inductive of Id.t list (* list of params *)
+ | Inductive of Id.t list (* list of params *) * bool (* true = check for possible capture *)
| Recursive
| Method
| Variable
@@ -65,11 +67,10 @@ type var_internalization_data =
type internalization_env =
(var_internalization_data) Id.Map.t
-type glob_binder = (Name.t * binding_kind * glob_constr option * glob_constr)
-
type ltac_sign = {
ltac_vars : Id.Set.t;
ltac_bound : Id.Set.t;
+ ltac_extra : Genintern.Store.t;
}
let interning_grammar = ref false
@@ -95,19 +96,19 @@ let is_global id =
false
let global_reference_of_reference ref =
- locate_reference (snd (qualid_of_reference ref))
+ locate_reference (qualid_of_reference ref).CAst.v
let global_reference id =
- Universes.constr_of_global (locate_reference (qualid_of_ident id))
+ locate_reference (qualid_of_ident id)
let construct_reference ctx id =
try
- Term.mkVar (let _ = Context.Named.lookup id ctx in id)
+ VarRef (let _ = Context.Named.lookup id ctx in id)
with Not_found ->
global_reference id
let global_reference_in_absolute_module dir id =
- Universes.constr_of_global (Nametab.global_of_path (Libnames.make_path dir id))
+ Nametab.global_of_path (Libnames.make_path dir id)
(**********************************************************************)
(* Internalization errors *)
@@ -120,10 +121,10 @@ type internalization_error =
| NonLinearPattern of Id.t
| BadPatternsNumber of int * int
-exception InternalizationError of Loc.t * internalization_error
+exception InternalizationError of internalization_error Loc.located
let explain_variable_capture id id' =
- pr_id id ++ str " is dependent in the type of " ++ pr_id id' ++
+ Id.print id ++ str " is dependent in the type of " ++ Id.print id' ++
strbrk ": cannot interpret both of them with the same type"
let explain_illegal_metavariable =
@@ -133,12 +134,12 @@ let explain_not_a_constructor ref =
str "Unknown constructor: " ++ pr_reference ref
let explain_unbound_fix_name is_cofix id =
- str "The name" ++ spc () ++ pr_id id ++
+ str "The name" ++ spc () ++ Id.print id ++
spc () ++ str "is not bound in the corresponding" ++ spc () ++
str (if is_cofix then "co" else "") ++ str "fixpoint definition"
let explain_non_linear_pattern id =
- str "The variable " ++ pr_id id ++ str " is bound several times in pattern"
+ str "The variable " ++ Id.print id ++ str " is bound several times in pattern"
let explain_bad_patterns_number n1 n2 =
str "Expecting " ++ int n1 ++ str (String.plural n1 " pattern") ++
@@ -154,17 +155,17 @@ let explain_internalization_error e =
| BadPatternsNumber (n1,n2) -> explain_bad_patterns_number n1 n2
in pp ++ str "."
-let error_bad_inductive_type loc =
- user_err_loc (loc,"",str
+let error_bad_inductive_type ?loc =
+ user_err ?loc (str
"This should be an inductive type applied to patterns.")
-let error_parameter_not_implicit loc =
- user_err_loc (loc,"", str
+let error_parameter_not_implicit ?loc =
+ user_err ?loc (str
"The parameters do not bind in patterns;" ++ spc () ++ str
"they must be replaced by '_'.")
-let error_ldots_var loc =
- user_err_loc (loc,"",str "Special token " ++ pr_id ldots_var ++
+let error_ldots_var ?loc =
+ user_err ?loc (str "Special token " ++ Id.print ldots_var ++
str " is for use in the Notation command.")
(**********************************************************************)
@@ -176,7 +177,7 @@ let parsing_explicit = ref false
let empty_internalization_env = Id.Map.empty
let compute_explicitable_implicit imps = function
- | Inductive params ->
+ | Inductive (params,_) ->
(* In inductive types, the parameters are fixed implicit arguments *)
let sub_impl,_ = List.chop (List.length params) imps in
let sub_impl' = List.filter is_status_implicit sub_impl in
@@ -185,15 +186,15 @@ let compute_explicitable_implicit imps = function
(* Unable to know in advance what the implicit arguments will be *)
[]
-let compute_internalization_data env ty typ impl =
- let impl = compute_implicits_with_manual env typ (is_implicit_args()) impl in
+let compute_internalization_data env sigma ty typ impl =
+ let impl = compute_implicits_with_manual env sigma typ (is_implicit_args()) impl in
let expls_impl = compute_explicitable_implicit impl ty in
- (ty, expls_impl, impl, compute_arguments_scope typ)
+ (ty, expls_impl, impl, compute_arguments_scope sigma typ)
-let compute_internalization_env env ty =
+let compute_internalization_env env sigma ?(impls=empty_internalization_env) ty =
List.fold_left3
- (fun map id typ impl -> Id.Map.add id (compute_internalization_data env ty typ impl) map)
- empty_internalization_env
+ (fun map id typ impl -> Id.Map.add id (compute_internalization_data env sigma ty typ impl) map)
+ impls
(**********************************************************************)
(* Contracting "{ _ }" in notations *)
@@ -215,24 +216,24 @@ let expand_notation_string ntn n =
(* This contracts the special case of "{ _ }" for sumbool, sumor notations *)
(* Remark: expansion of squash at definition is done in metasyntax.ml *)
-let contract_notation ntn (l,ll,bll) =
+let contract_curly_brackets ntn (l,ll,bl,bll) =
let ntn' = ref ntn in
let rec contract_squash n = function
| [] -> []
- | CNotation (_,"{ _ }",([a],[],[])) :: l ->
+ | { CAst.v = CNotation ("{ _ }",([a],[],[],[])) } :: l ->
ntn' := expand_notation_string !ntn' n;
contract_squash n (a::l)
| a :: l ->
a::contract_squash (n+1) l in
let l = contract_squash 0 l in
(* side effect; don't inline *)
- !ntn',(l,ll,bll)
+ !ntn',(l,ll,bl,bll)
-let contract_pat_notation ntn (l,ll) =
+let contract_curly_brackets_pat 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 ->
@@ -262,40 +263,35 @@ let pr_scope_stack = function
| l -> str "scope stack " ++
str "[" ++ prlist_with_sep pr_comma str l ++ str "]"
-let error_inconsistent_scope loc id scopes1 scopes2 =
- user_err_loc (loc,"set_var_scope",
- pr_id id ++ str " is here used in " ++
+let error_inconsistent_scope ?loc id scopes1 scopes2 =
+ user_err ?loc ~hdr:"set_var_scope"
+ (Id.print id ++ str " is here used in " ++
pr_scope_stack scopes2 ++ strbrk " while it was elsewhere used in " ++
pr_scope_stack scopes1)
-let error_expect_binder_notation_type loc id =
- user_err_loc (loc,"",
- pr_id id ++
+let error_expect_binder_notation_type ?loc id =
+ user_err ?loc
+ (Id.print id ++
str " is expected to occur in binding position in the right-hand side.")
-let set_var_scope loc id istermvar env ntnvars =
+let set_var_scope ?loc id istermvar (tmp_scope,subscopes as scopes) ntnvars =
try
- let isonlybinding,idscopes,typ = Id.Map.find id ntnvars in
- if istermvar then isonlybinding := false;
+ let used_as_binder,idscopes,typ = Id.Map.find id ntnvars in
+ if not istermvar then used_as_binder := true;
let () = if istermvar then
(* scopes have no effect on the interpretation of identifiers *)
begin match !idscopes with
- | None -> idscopes := Some (env.tmp_scope, env.scopes)
- | Some (tmp, scope) ->
- let s1 = make_current_scope tmp scope in
- let s2 = make_current_scope env.tmp_scope env.scopes in
- if not (List.equal String.equal s1 s2) then error_inconsistent_scope loc id s1 s2
+ | None -> idscopes := Some scopes
+ | Some (tmp_scope', subscopes') ->
+ let s' = make_current_scope tmp_scope' subscopes' in
+ let s = make_current_scope tmp_scope subscopes in
+ if not (List.equal String.equal s' s) then error_inconsistent_scope ?loc id s' s
end
in
match typ with
- | NtnInternTypeBinder ->
- if istermvar then error_expect_binder_notation_type loc id
- | NtnInternTypeConstr ->
- (* We need sometimes to parse idents at a constr level for
- factorization and we cannot enforce this constraint:
- if not istermvar then error_expect_constr_notation_type loc id *)
- ()
- | NtnInternTypeIdent -> ()
+ | Notation_term.NtnInternTypeOnlyBinder ->
+ if istermvar then error_expect_binder_notation_type ?loc id
+ | Notation_term.NtnInternTypeAny -> ()
with Not_found ->
(* Not in a notation *)
()
@@ -304,15 +300,11 @@ let set_type_scope env = {env with tmp_scope = Notation.current_type_scope_name
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))
- | [] -> body
+let set_env_scopes env (scopt,subscopes) =
+ {env with tmp_scope = scopt; scopes = subscopes @ env.scopes}
-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))
- | [] -> body
+let mkGProd ?loc (na,bk,t) body = DAst.make ?loc @@ GProd (na, bk, t, body)
+let mkGLambda ?loc (na,bk,t) body = DAst.make ?loc @@ GLambda (na, bk, t, body)
(**********************************************************************)
(* Utilities for binders *)
@@ -323,15 +315,15 @@ let build_impls = function
|Explicit -> fun _ -> None
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),[])
+ let rec aux acc c = match DAst.get c with
+ | 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) ->
+ let rec aux acc c = match DAst.get c with
+ | 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)
@@ -339,53 +331,55 @@ let impls_term_list ?(args = []) =
in aux []
(* Check if in binder "(x1 x2 .. xn : t)", none of x1 .. xn-1 occurs in t *)
-let rec check_capture ty = function
- | (loc,Name id)::(_,Name id')::_ when occur_glob_constr id ty ->
+let rec check_capture ty = let open CAst in function
+ | { loc; v = Name id } :: { v = Name id' } :: _ when occur_glob_constr id ty ->
raise (InternalizationError (loc,VariableCapture (id,id')))
| _::nal ->
check_capture ty nal
| [] ->
()
-let locate_if_hole loc na = function
- | GHole (_,_,naming,arg) ->
+let locate_if_hole ?loc na c = match DAst.get c with
+ | 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 -> GHole (loc, Evar_kinds.BinderType na, naming, arg))
- | x -> x
+ with Not_found -> DAst.make ?loc @@ GHole (Evar_kinds.BinderType na, naming, arg))
+ | _ -> c
let reset_hidden_inductive_implicit_test env =
{ env with impls = Id.Map.map (function
- | (Inductive _,b,c,d) -> (Inductive [],b,c,d)
+ | (Inductive (params,_),b,c,d) -> (Inductive (params,false),b,c,d)
| x -> x) env.impls }
-let check_hidden_implicit_parameters id impls =
+let check_hidden_implicit_parameters ?loc id impls =
if Id.Map.exists (fun _ -> function
- | (Inductive indparams,_,_,_) -> Id.List.mem id indparams
+ | (Inductive (indparams,check),_,_,_) when check -> Id.List.mem id indparams
| _ -> false) impls
then
- errorlabstrm "" (strbrk "A parameter of an inductive type " ++
- pr_id id ++ strbrk " is not allowed to be used as a bound variable in the type of its constructor.")
+ user_err ?loc (Id.print id ++ strbrk " is already used as name of " ++
+ strbrk "a parameter of the inductive type; bound variables in " ++
+ strbrk "the type of a constructor shall use a different name.")
let push_name_env ?(global_level=false) ntnvars implargs env =
+ let open CAst in
function
- | loc,Anonymous ->
+ | { loc; v = Anonymous } ->
if global_level then
- user_err_loc (loc,"", str "Anonymous variables not allowed");
+ user_err ?loc (str "Anonymous variables not allowed");
env
- | loc,Name id ->
- check_hidden_implicit_parameters id env.impls ;
+ | { loc; v = Name id } ->
+ check_hidden_implicit_parameters ?loc id env.impls ;
if Id.Map.is_empty ntnvars && Id.equal id ldots_var
- then error_ldots_var loc;
- set_var_scope loc id false env ntnvars;
- if global_level then Dumpglob.dump_definition (loc,id) true "var"
- else Dumpglob.dump_binding loc id;
+ then error_ldots_var ?loc;
+ set_var_scope ?loc id false (env.tmp_scope,env.scopes) ntnvars;
+ if global_level then Dumpglob.dump_definition CAst.(make ?loc id) true "var"
+ else Dumpglob.dump_binding ?loc id;
{env with ids = Id.Set.add id env.ids; impls = Id.Map.add id implargs env.impls}
-let intern_generalized_binder ?(global_level=false) intern_type lvar
- env (loc, na) b b' t ty =
+let intern_generalized_binder ?(global_level=false) intern_type ntnvars
+ env {loc;v=na} b b' t ty =
let ids = (match na with Anonymous -> fun x -> x | Name na -> Id.Set.add na) env.ids in
let ty, ids' =
if t then ty, ids else
@@ -395,11 +389,11 @@ let intern_generalized_binder ?(global_level=false) intern_type lvar
let ty' = intern_type {env with ids = ids; unb = true} ty in
let fvs = Implicit_quantifiers.generalizable_vars_of_glob_constr ~bound:ids ~allowed:ids' ty' in
let env' = List.fold_left
- (fun env (x, l) -> push_name_env ~global_level lvar (Variable,[],[],[])(*?*) env (l, Name x))
+ (fun env {loc;v=x} -> push_name_env ~global_level ntnvars (Variable,[],[],[])(*?*) env (make ?loc @@ Name x))
env fvs in
let bl = List.map
- (fun (id, loc) ->
- (loc, (Name id, b, None, GHole (loc, Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None))))
+ CAst.(map (fun id ->
+ (Name id, b, DAst.make ?loc @@ GHole (Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None))))
fvs
in
let na = match na with
@@ -409,14 +403,14 @@ let intern_generalized_binder ?(global_level=false) intern_type lvar
let name =
let id =
match ty with
- | CApp (_, (_, CRef (Ident (loc,id),_)), _) -> id
+ | { v = CApp ((_, { v = CRef ({v=Ident id},_) } ), _) } -> id
| _ -> default_non_dependent_ident
in Implicit_quantifiers.make_fresh ids' (Global.env ()) id
in Name name
| _ -> na
- in (push_name_env ~global_level lvar (impls_type_list ty')(*?*) env' (loc,na)), (loc,(na,b',None,ty')) :: List.rev bl
+ in (push_name_env ~global_level ntnvars (impls_type_list ty')(*?*) env' (make ?loc na)), (make ?loc (na,b',ty')) :: List.rev bl
-let intern_assumption intern lvar env nal bk ty =
+let intern_assumption intern ntnvars env nal bk ty =
let intern_type env = intern (set_type_scope env) in
match bk with
| Default k ->
@@ -424,82 +418,68 @@ let intern_assumption intern lvar env nal bk ty =
check_capture ty nal;
let impls = impls_type_list ty in
List.fold_left
- (fun (env, bl) (loc, na as locna) ->
- (push_name_env lvar impls env locna,
- (loc,(na,k,None,locate_if_hole loc na ty))::bl))
+ (fun (env, bl) ({loc;v=na} as locna) ->
+ (push_name_env ntnvars impls env locna,
+ (make ?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
+ let env, b = intern_generalized_binder intern_type ntnvars env (List.hd nal) b b' t ty in
env, b
-let rec free_vars_of_pat il =
- function
- | CPatCstr (loc, 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
- | CPatAtom (loc, ro) ->
- begin match ro with
- | Some (Ident (loc, i)) -> (loc, i) :: il
- | Some _ | None -> il
- end
- | CPatNotation (loc, n, l1, l2) ->
- let il = List.fold_left free_vars_of_pat il (fst l1) in
- List.fold_left (List.fold_left free_vars_of_pat) il (snd l1)
- | _ -> anomaly (str "free_vars_of_pat")
-
-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 n = Name i in
- let env, _ = intern_assumption intern lvar env [(loc, n)] bk ty in
- env)
- env (free_vars_of_pat [] p)
-
-type binder_data =
- | BDRawDef of (Loc.t * glob_binder)
- | BDPattern of
- (Loc.t * (cases_pattern * Id.t list) *
- (bool ref *
- (Notation_term.tmp_scope_name option *
- Notation_term.tmp_scope_name list)
- option ref * Notation_term.notation_var_internalization_type)
- Names.Id.Map.t *
- intern_env * constr_expr)
+let glob_local_binder_of_extended = DAst.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 = DAst.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")
+ )
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
- | LocalRawAssum(nal,bk,ty) ->
- let env, bl' = intern_assumption intern lvar env nal bk ty in
- let bl' = List.map (fun a -> BDRawDef a) bl' in
+let intern_letin_binder intern ntnvars env (({loc;v=na} as locna),def,ty) =
+ let term = intern env def in
+ let ty = Option.map (intern env) ty in
+ (push_name_env ntnvars (impls_term_list term) env locna,
+ (na,Explicit,term,ty))
+
+let intern_cases_pattern_as_binder ?loc ntnvars env p =
+ let il,disjpat =
+ let (il, subst_disjpat) = !intern_cases_pattern_fwd ntnvars (None,env.scopes) p in
+ let substl,disjpat = List.split subst_disjpat in
+ if not (List.for_all (fun subst -> Id.Map.equal Id.equal subst Id.Map.empty) substl) then
+ user_err ?loc (str "Unsupported nested \"as\" clause.");
+ il,disjpat
+ in
+ let env = List.fold_right (fun {loc;v=id} env -> push_name_env ntnvars (Variable,[],[],[]) env (make ?loc @@ Name id)) il env in
+ let na = alias_of_pat (List.hd disjpat) in
+ let ienv = Name.fold_right Id.Set.remove na env.ids in
+ let id = Namegen.next_name_away_with_default "pat" na ienv in
+ let na = make ?loc @@ Name id in
+ env,((disjpat,il),id),na
+
+let intern_local_binder_aux ?(global_level=false) intern ntnvars (env,bl) = function
+ | CLocalAssum(nal,bk,ty) ->
+ let env, bl' = intern_assumption intern ntnvars env nal bk ty in
+ let bl' = List.map (fun {loc;v=(na,c,t)} -> DAst.make ?loc @@ GLocalAssum (na,c,t)) bl' in
env, bl' @ bl
- | LocalRawDef((loc,na as locna),def) ->
- let indef = intern env def in
- let term, ty =
- match indef with
- | GCast (loc, b, Misctypes.CastConv t) -> b, t
- | _ -> indef, GHole(loc,Evar_kinds.BinderType na,Misctypes.IntroAnonymous,None)
- in
- (push_name_env lvar (impls_term_list indef) env locna,
- (BDRawDef ((loc,(na,Explicit,Some(term),ty))))::bl)
- | LocalPattern (loc,p,ty) ->
+ | CLocalDef( {loc; v=na} as locna,def,ty) ->
+ let env,(na,bk,def,ty) = intern_letin_binder intern ntnvars env (locna,def,ty) in
+ env, (DAst.make ?loc @@ GLocalDef (na,bk,def,ty)) :: bl
+ | CLocalPattern {loc;v=(p,ty)} ->
let tyc =
match ty with
| Some ty -> ty
- | None -> CHole(loc,None,Misctypes.IntroAnonymous,None)
- in
- let env = intern_local_pattern intern lvar env p in
- let cp =
- match !intern_cases_pattern_fwd (None,env.scopes) p with
- | (_, [(_, cp)]) -> cp
- | _ -> assert false
+ | None -> CAst.make ?loc @@ CHole(None,Misctypes.IntroAnonymous,None)
in
- let il = List.map snd (free_vars_of_pat [] p) in
- (env, BDPattern(loc,(cp,il),lvar,env,tyc) :: bl)
+ let env, ((disjpat,il),id),na = intern_cases_pattern_as_binder ?loc ntnvars env p in
+ let bk = Default Explicit in
+ let _, bl' = intern_assumption intern ntnvars env [na] bk tyc in
+ let {v=(_,bk,t)} = List.hd bl' in
+ (env, (DAst.make ?loc @@ GLocalPattern((disjpat,List.map (fun x -> x.v) il),id,bk,t)) :: bl)
-let intern_generalization intern env lvar loc bk ak c =
+let intern_generalization intern env ntnvars loc bk ak c =
let c = intern {env with unb = true} c in
let fvs = Implicit_quantifiers.generalizable_vars_of_glob_constr ~bound:env.ids c in
let env', c' =
@@ -519,17 +499,35 @@ let intern_generalization intern env lvar loc bk ak c =
| None -> false
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))
+ (fun {loc=loc';v=id} acc ->
+ DAst.make ?loc:(Loc.merge_opt loc' loc) @@
+ GProd (Name id, bk, DAst.make ?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))
+ (fun {loc=loc';v=id} acc ->
+ DAst.make ?loc:(Loc.merge_opt loc' loc) @@
+ GLambda (Name id, bk, DAst.make ?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
+ List.fold_right (fun ({loc;v=id} as lid) (env, acc) ->
+ let env' = push_name_env ntnvars (Variable,[],[],[]) env CAst.(make @@ Name id) in
(env', abs lid acc)) fvs (env,c)
in c'
+let rec expand_binders ?loc mk bl c =
+ match bl with
+ | [] -> c
+ | b :: bl ->
+ match DAst.get b with
+ | GLocalDef (n, bk, b, oty) ->
+ expand_binders ?loc mk bl (DAst.make ?loc @@ GLetIn (n, b, oty, c))
+ | GLocalAssum (n, bk, t) ->
+ expand_binders ?loc mk bl (mk ?loc (n,bk,t) c)
+ | GLocalPattern ((disjpat,ids), id, bk, ty) ->
+ let tm = DAst.make ?loc (GVar id) in
+ (* Distribute the disjunctive patterns over the shared right-hand side *)
+ let eqnl = List.map (fun pat -> CAst.make ?loc (ids,[pat],c)) disjpat in
+ let c = DAst.make ?loc @@ GCases (Misctypes.LetPatternStyle, None, [tm,(Anonymous,None)], eqnl) in
+ expand_binders ?loc mk bl (mk ?loc (Name id,Explicit,ty) c)
+
(**********************************************************************)
(* Syntax extensions *)
@@ -537,7 +535,7 @@ let option_mem_assoc id = function
| Some (id',c) -> Id.equal id id'
| None -> false
-let find_fresh_name renaming (terms,termlists,binders) avoid id =
+let find_fresh_name renaming (terms,termlists,binders,binderlists) avoid id =
let fold1 _ (c, _) accu = Id.Set.union (free_vars_of_constr_expr c) accu in
let fold2 _ (l, _) accu =
let fold accu c = Id.Set.union (free_vars_of_constr_expr c) accu in
@@ -550,13 +548,53 @@ let find_fresh_name renaming (terms,termlists,binders) avoid id =
(* TODO binders *)
next_ident_away_from id (fun id -> Id.Set.mem id fvs3)
-let traverse_binder (terms,_,_ as subst) avoid (renaming,env) = function
- | Anonymous -> (renaming,env),Anonymous
+let is_var store pat =
+ match DAst.get pat with
+ | PatVar na -> store na; true
+ | _ -> false
+
+let out_var pat =
+ match pat.v with
+ | CPatAtom (Some ({v=Ident id})) -> Name id
+ | CPatAtom None -> Anonymous
+ | _ -> assert false
+
+let term_of_name = function
+ | Name id -> DAst.make (GVar id)
+ | Anonymous ->
+ let st = Evar_kinds.Define (not (Program.get_proofs_transparency ())) in
+ DAst.make (GHole (Evar_kinds.QuestionMark (st,Anonymous), Misctypes.IntroAnonymous, None))
+
+let traverse_binder intern_pat ntnvars (terms,_,binders,_ as subst) avoid (renaming,env) = function
+ | Anonymous -> (renaming,env), None, Anonymous
| Name id ->
+ let store,get = set_temporary_memory () in
+ try
+ (* We instantiate binder name with patterns which may be parsed as terms *)
+ let pat = coerce_to_cases_pattern_expr (fst (Id.Map.find id terms)) in
+ let env,((disjpat,ids),id),na = intern_pat ntnvars env pat in
+ let pat, na = match disjpat with
+ | [pat] when is_var store pat -> let na = get () in None, na
+ | _ -> Some ((List.map (fun x -> x.v) ids,disjpat),id), na.v in
+ (renaming,env), pat, na
+ with Not_found ->
try
- (* Binders bound in the notation are considered first-order objects *)
- let _,na = coerce_to_name (fst (Id.Map.find id terms)) in
- (renaming,{env with ids = name_fold Id.Set.add na env.ids}), na
+ (* Trying to associate a pattern *)
+ let pat,(onlyident,scopes) = Id.Map.find id binders in
+ let env = set_env_scopes env scopes in
+ if onlyident then
+ (* Do not try to interpret a variable as a constructor *)
+ let na = out_var pat in
+ let env = push_name_env ntnvars (Variable,[],[],[]) env (make ?loc:pat.loc na) in
+ (renaming,env), None, na
+ else
+ (* Interpret as a pattern *)
+ let env,((disjpat,ids),id),na = intern_pat ntnvars env pat in
+ let pat, na =
+ match disjpat with
+ | [pat] when is_var store pat -> let na = get () in None, na
+ | _ -> Some ((List.map (fun x -> x.v) ids,disjpat),id), na.v in
+ (renaming,env), pat, na
with Not_found ->
(* Binders not bound in the notation do not capture variables *)
(* outside the notation (i.e. in the substitution) *)
@@ -564,92 +602,101 @@ let traverse_binder (terms,_,_ as subst) avoid (renaming,env) = function
let renaming' =
if Id.equal id id' then renaming else Id.Map.add id id' renaming
in
- (renaming',env), Name id'
-
-type letin_param =
- | LPLetIn of Loc.t * (Name.t * glob_constr)
- | LPCases of Loc.t * (cases_pattern * Id.t list) * Id.t
-
-let make_letins =
- List.fold_right
- (fun a c ->
- match a with
- | LPLetIn (loc,(na,b)) ->
- GLetIn(loc,na,b,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 rec subordinate_letins intern letins = function
- (* binders come in reverse order; the non-let are returned in reverse order together *)
- (* with the subordinated let-in in writing order *)
- | BDRawDef (loc,(na,_,Some b,t))::l ->
- subordinate_letins intern (LPLetIn (loc,(na,b))::letins) l
- | BDRawDef (loc,(na,bk,None,t))::l ->
- let letins',rest = subordinate_letins intern [] l in
- letins',((loc,(na,bk,t)),letins)::rest
- | BDPattern (loc,u,lvar,env,tyc) :: l ->
- let ienv = Id.Set.elements env.ids in
- let id = Namegen.next_ident_away (Id.of_string "pat") ienv in
- let na = (loc, Name id) in
- let bk = Default Explicit in
- let _, bl' = intern_assumption intern lvar env [na] bk tyc in
- let bl' = List.map (fun a -> BDRawDef a) bl' in
- subordinate_letins intern (LPCases (loc,u,id)::letins) (bl'@ l)
- | [] ->
- letins,[]
+ (renaming',env), None, Name id'
+
+type binder_action =
+| AddLetIn of Misctypes.lname * constr_expr * constr_expr option
+| AddTermIter of (constr_expr * subscopes) Names.Id.Map.t
+| AddPreBinderIter of Id.t * local_binder_expr (* A binder to be internalized *)
+| AddBinderIter of Id.t * extended_glob_local_binder (* A binder already internalized - used for generalized binders *)
+
+let dmap_with_loc f n =
+ CAst.map_with_loc (fun ?loc c -> f ?loc (DAst.get_thunk c)) n
+
+let error_cannot_coerce_wildcard_term ?loc () =
+ user_err ?loc Pp.(str "Cannot turn \"_\" into a term.")
+
+let error_cannot_coerce_disjunctive_pattern_term ?loc () =
+ user_err ?loc Pp.(str "Cannot turn a disjunctive pattern into a term.")
let terms_of_binders bl =
- let rec term_of_pat = function
- | PatVar (loc,Name id) -> 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 rec term_of_pat pt = dmap_with_loc (fun ?loc -> function
+ | PatVar (Name id) -> CRef (make @@ Ident id, None)
+ | PatVar (Anonymous) -> error_cannot_coerce_wildcard_term ?loc ()
+ | PatCstr (c,l,_) ->
+ let r = make ?loc @@ Qualid (qualid_of_path (path_of_global (ConstructRef c))) in
+ let hole = CAst.make ?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
- let rec extract_variables = function
- | BDRawDef (loc,(Name id,_,None,_))::l -> CRef (Ident (loc,id), None) :: extract_variables l
- | BDRawDef (loc,(Name id,_,Some _,_))::l -> extract_variables l
- | BDRawDef (loc,(Anonymous,_,_,_))::l -> error "Cannot turn \"_\" into a term."
- | BDPattern (loc,(u,_),lvar,env,tyc) :: l -> term_of_pat u :: extract_variables l
+ CAppExpl ((None,r,None),params @ List.map term_of_pat l)) pt in
+ let rec extract_variables l = match l with
+ | bnd :: l ->
+ let loc = bnd.loc in
+ begin match DAst.get bnd with
+ | GLocalAssum (Name id,_,_) -> (CAst.make ?loc @@ CRef (make ?loc @@ Ident id, None)) :: extract_variables l
+ | GLocalDef (Name id,_,_,_) -> extract_variables l
+ | GLocalDef (Anonymous,_,_,_)
+ | GLocalAssum (Anonymous,_,_) -> user_err Pp.(str "Cannot turn \"_\" into a term.")
+ | GLocalPattern (([u],_),_,_,_) -> term_of_pat u :: extract_variables l
+ | GLocalPattern ((_,_),_,_,_) -> error_cannot_coerce_disjunctive_pattern_term ?loc ()
+ end
| [] -> [] in
extract_variables bl
-let instantiate_notation_constr loc intern ntnvars subst infos c =
- let (terms,termlists,binders) = subst in
+let flatten_generalized_binders_if_any y l =
+ match List.rev l with
+ | [] -> assert false
+ | a::l -> a, List.map (fun a -> AddBinderIter (y,a)) l (* if l not empty, this means we had a generalized binder *)
+
+let flatten_binders bl =
+ let dispatch = function
+ | CLocalAssum (nal,bk,t) -> List.map (fun na -> CLocalAssum ([na],bk,t)) nal
+ | a -> [a] in
+ List.flatten (List.map dispatch bl)
+
+let instantiate_notation_constr loc intern intern_pat ntnvars subst infos c =
+ let (terms,termlists,binders,binderlists) = subst in
(* when called while defining a notation, avoid capturing the private binders
of the expression by variables bound by the notation (see #3892) *)
let avoid = Id.Map.domain ntnvars in
- let rec aux (terms,binderopt,terminopt as subst') (renaming,env) c =
+ let rec aux (terms,binderopt,iteropt as subst') (renaming,env) c =
let subinfos = renaming,{env with tmp_scope = None} in
match c with
- | NVar id when Id.equal id ldots_var -> Option.get terminopt
+ | NVar id when Id.equal id ldots_var ->
+ let rec aux_letin env = function
+ | [],terminator,_ -> aux (terms,None,None) (renaming,env) terminator
+ | AddPreBinderIter (y,binder)::rest,terminator,iter ->
+ let env,binders = intern_local_binder_aux intern ntnvars (env,[]) binder in
+ let binder,extra = flatten_generalized_binders_if_any y binders in
+ aux (terms,Some (y,binder),Some (extra@rest,terminator,iter)) (renaming,env) iter
+ | AddBinderIter (y,binder)::rest,terminator,iter ->
+ aux (terms,Some (y,binder),Some (rest,terminator,iter)) (renaming,env) iter
+ | AddTermIter nterms::rest,terminator,iter ->
+ aux (nterms,None,Some (rest,terminator,iter)) (renaming,env) iter
+ | AddLetIn (na,c,t)::rest,terminator,iter ->
+ let env,(na,_,c,t) = intern_letin_binder intern ntnvars env (na,c,t) in
+ DAst.make ?loc (GLetIn (na,c,t,aux_letin env (rest,terminator,iter))) in
+ aux_letin env (Option.get iteropt)
| NVar id -> subst_var subst' (renaming, env) id
- | NList (x,y,iter,terminator,lassoc) ->
+ | NList (x,y,iter,terminator,revert) ->
let l,(scopt,subscopes) =
(* All elements of the list are in scopes (scopt,subscopes) *)
try
let l,scopes = Id.Map.find x termlists in
- (if lassoc then List.rev l else l),scopes
+ (if revert then List.rev l else l),scopes
with Not_found ->
try
- let (bl,(scopt,subscopes)) = Id.Map.find x binders in
+ let (bl,(scopt,subscopes)) = Id.Map.find x binderlists in
let env,bl' = List.fold_left (intern_local_binder_aux intern ntnvars) (env,[]) bl in
- terms_of_binders (if lassoc then bl' else List.rev bl'),(None,[])
+ terms_of_binders (if revert then bl' else List.rev bl'),(None,[])
with Not_found ->
- anomaly (Pp.str "Inconsistent substitution of recursive notation") in
- let termin = aux (terms,None,None) subinfos terminator in
- let fold a t =
- let nterms = Id.Map.add y (a, (scopt, subscopes)) terms in
- aux (nterms,None,Some t) subinfos iter
- in
- List.fold_right fold l termin
+ anomaly (Pp.str "Inconsistent substitution of recursive notation.") in
+ let l = List.map (fun a -> AddTermIter ((Id.Map.add y (a,(scopt,subscopes)) terms))) l in
+ aux (terms,None,Some (l,terminator,iter)) subinfos (NVar ldots_var)
| NHole (knd, naming, arg) ->
let knd = match knd with
| Evar_kinds.BinderType (Name id as na) ->
let na =
- try snd (coerce_to_name (fst (Id.Map.find id terms)))
+ try (coerce_to_name (fst (Id.Map.find id terms))).v
with Not_found ->
try Name (Id.Map.find id renaming)
with Not_found -> na
@@ -660,62 +707,62 @@ let instantiate_notation_constr loc intern ntnvars subst infos c =
let arg = match arg with
| None -> None
| Some arg ->
- let open Tacexpr in
- let open Genarg in
- let wit = glbwit Constrarg.wit_tactic in
- let body =
- if has_type arg wit then out_gen wit arg
- else assert false (** FIXME *)
- in
- let mk_env id (c, (tmp_scope, subscopes)) accu =
+ let mk_env (c, (tmp_scope, subscopes)) =
let nenv = {env with tmp_scope; scopes = subscopes @ env.scopes} in
let gc = intern nenv c in
- let c = ConstrMayEval (Genredexpr.ConstrTerm (gc, Some c)) in
- ((loc, id), c) :: accu
+ (gc, Some c)
+ in
+ let mk_env' (c, (onlyident,(tmp_scope,subscopes))) =
+ let nenv = {env with tmp_scope; scopes = subscopes @ env.scopes} in
+ if onlyident then
+ let na = out_var c in term_of_name na, None
+ else
+ let _,((disjpat,_),_),_ = intern_pat ntnvars nenv c in
+ match disjpat with
+ | [pat] -> (glob_constr_of_cases_pattern pat, None)
+ | _ -> error_cannot_coerce_disjunctive_pattern_term ?loc:c.loc ()
in
- let bindings = Id.Map.fold mk_env terms [] in
- let tac = TacLetIn (false, bindings, body) in
- let arg = in_gen wit tac in
- Some arg
+ let terms = Id.Map.map mk_env terms in
+ let binders = Id.Map.map mk_env' binders in
+ let bindings = Id.Map.fold Id.Map.add terms binders in
+ Some (Genintern.generic_substitute_notation bindings arg)
in
- GHole (loc, knd, naming, arg)
- | NBinderList (x,y,iter,terminator) ->
+ DAst.make ?loc @@ GHole (knd, naming, arg)
+ | NBinderList (x,y,iter,terminator,revert) ->
(try
(* All elements of the list are in scopes (scopt,subscopes) *)
- let (bl,(scopt,subscopes)) = Id.Map.find x binders in
- let env,bl = List.fold_left (intern_local_binder_aux intern ntnvars) (env,[]) bl in
- let letins,bl = subordinate_letins intern [] bl in
- let termin = aux (terms,None,None) (renaming,env) terminator in
- let res = List.fold_left (fun t binder ->
- aux (terms,Some(y,binder),Some t) subinfos iter)
- termin bl in
- make_letins letins res
+ let (bl,(scopt,subscopes)) = Id.Map.find x binderlists in
+ (* We flatten binders so that we can interpret them at substitution time *)
+ let bl = flatten_binders bl in
+ let bl = if revert then List.rev bl else bl in
+ (* We isolate let-ins which do not contribute to the repeated pattern *)
+ let l = List.map (function | CLocalDef (na,c,t) -> AddLetIn (na,c,t)
+ | binder -> AddPreBinderIter (y,binder)) bl in
+ (* We stack the binders to iterate or let-ins to insert *)
+ aux (terms,None,Some (l,terminator,iter)) subinfos (NVar ldots_var)
with Not_found ->
- anomaly (Pp.str "Inconsistent substitution of recursive notation"))
+ anomaly (Pp.str "Inconsistent substitution of recursive notation."))
| NProd (Name id, NHole _, c') when option_mem_assoc id binderopt ->
- 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)
+ let binder = snd (Option.get binderopt) in
+ expand_binders ?loc mkGProd [binder] (aux subst' (renaming,env) c')
| 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'))
+ let binder = snd (Option.get binderopt) in
+ expand_binders ?loc mkGLambda [binder] (aux subst' (renaming,env) 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 subinfos,disjpat,na = traverse_binder intern_pat ntnvars subst avoid subinfos na in
+ let ty = DAst.make ?loc @@ GHole (Evar_kinds.BinderType na,naming,arg) in
+ DAst.make ?loc @@ GProd (na,Explicit,ty,Option.fold_right apply_cases_pattern disjpat (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 subinfos,disjpat,na = traverse_binder intern_pat ntnvars subst avoid subinfos na in
+ let ty = DAst.make ?loc @@ GHole (Evar_kinds.BinderType na,naming,arg) in
+ DAst.make ?loc @@ GLambda (na,Explicit,ty,Option.fold_right apply_cases_pattern disjpat (aux subst' subinfos c'))
| t ->
- 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 =
+ glob_constr_of_notation_constr_with_binders ?loc
+ (traverse_binder intern_pat ntnvars subst avoid) (aux subst') subinfos t
+ and subst_var (terms, binderopt, _terminopt) (renaming, env) id =
(* subst remembers the delimiters stack in the interpretation *)
(* of the notations *)
try
@@ -724,33 +771,109 @@ let instantiate_notation_constr loc intern ntnvars subst infos c =
scopes = subscopes @ env.scopes} a
with Not_found ->
try
- GVar (loc, Id.Map.find id renaming)
+ let pat,(onlyident,scopes) = Id.Map.find id binders in
+ let env = set_env_scopes env scopes in
+ (* We deactivate impls to avoid the check on hidden parameters *)
+ (* and since we are only interested in the pattern as a term *)
+ let env = reset_hidden_inductive_implicit_test env in
+ if onlyident then
+ term_of_name (out_var pat)
+ else
+ let env,((disjpat,ids),id),na = intern_pat ntnvars env pat in
+ match disjpat with
+ | [pat] -> glob_constr_of_cases_pattern pat
+ | _ -> user_err Pp.(str "Cannot turn a disjunctive pattern into a term.")
+ with Not_found ->
+ try
+ match binderopt with
+ | Some (x,binder) when Id.equal x id ->
+ let terms = terms_of_binders [binder] in
+ assert (List.length terms = 1);
+ intern env (List.hd terms)
+ | _ -> raise Not_found
+ with Not_found ->
+ DAst.make ?loc (
+ try
+ 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 =
- List.fold_right (fun (x,(scl,typ)) (l1,l2,l3) ->
+(* Turning substitution coming from parsing and based on production
+ into a substitution for interpretation and based on binding/constr
+ distinction *)
+
+let cases_pattern_of_name {loc;v=na} =
+ let atom = match na with Name id -> Some (make ?loc @@ Ident id) | Anonymous -> None in
+ CAst.make ?loc (CPatAtom atom)
+
+let split_by_type ids subst =
+ let bind id scl l s =
+ match l with
+ | [] -> assert false
+ | a::l -> l, Id.Map.add id (a,scl) s in
+ let (terms,termlists,binders,binderlists),subst =
+ List.fold_left (fun ((terms,termlists,binders,binderlists),(terms',termlists',binders',binderlists')) (id,(scl,typ)) ->
match typ with
- | NtnTypeConstr | NtnTypeOnlyBinder -> ((x,scl)::l1,l2,l3)
- | NtnTypeConstrList -> (l1,(x,scl)::l2,l3)
- | NtnTypeBinderList -> (l1,l2,(x,scl)::l3)) ids ([],[],[])
+ | NtnTypeConstr ->
+ let terms,terms' = bind id scl terms terms' in
+ (terms,termlists,binders,binderlists),(terms',termlists',binders',binderlists')
+ | NtnTypeBinder NtnBinderParsedAsConstr (Extend.AsIdentOrPattern | Extend.AsStrictPattern) ->
+ let a,terms = match terms with a::terms -> a,terms | _ -> assert false in
+ let binders' = Id.Map.add id (coerce_to_cases_pattern_expr a,(false,scl)) binders' in
+ (terms,termlists,binders,binderlists),(terms',termlists',binders',binderlists')
+ | NtnTypeBinder NtnBinderParsedAsConstr Extend.AsIdent ->
+ let a,terms = match terms with a::terms -> a,terms | _ -> assert false in
+ let binders' = Id.Map.add id (cases_pattern_of_name (coerce_to_name a),(true,scl)) binders' in
+ (terms,termlists,binders,binderlists),(terms',termlists',binders',binderlists')
+ | NtnTypeBinder (NtnParsedAsIdent | NtnParsedAsPattern _ as x) ->
+ let onlyident = (x = NtnParsedAsIdent) in
+ let binders,binders' = bind id (onlyident,scl) binders binders' in
+ (terms,termlists,binders,binderlists),(terms',termlists',binders',binderlists')
+ | NtnTypeConstrList ->
+ let termlists,termlists' = bind id scl termlists termlists' in
+ (terms,termlists,binders,binderlists),(terms',termlists',binders',binderlists')
+ | NtnTypeBinderList ->
+ let binderlists,binderlists' = bind id scl binderlists binderlists' in
+ (terms,termlists,binders,binderlists),(terms',termlists',binders',binderlists'))
+ (subst,(Id.Map.empty,Id.Map.empty,Id.Map.empty,Id.Map.empty)) ids in
+ assert (terms = [] && termlists = [] && binders = [] && binderlists = []);
+ subst
+
+let split_by_type_pat ?loc ids subst =
+ let bind id scl l s =
+ match l with
+ | [] -> assert false
+ | a::l -> l, Id.Map.add id (a,scl) s in
+ let (terms,termlists),subst =
+ List.fold_left (fun ((terms,termlists),(terms',termlists')) (id,(scl,typ)) ->
+ match typ with
+ | NtnTypeConstr | NtnTypeBinder _ ->
+ let terms,terms' = bind id scl terms terms' in
+ (terms,termlists),(terms',termlists')
+ | NtnTypeConstrList ->
+ let termlists,termlists' = bind id scl termlists termlists' in
+ (terms,termlists),(terms',termlists')
+ | NtnTypeBinderList -> error_invalid_pattern_notation ?loc ())
+ (subst,(Id.Map.empty,Id.Map.empty)) ids in
+ assert (terms = [] && termlists = []);
+ subst
let make_subst ids l =
let fold accu (id, scl) a = Id.Map.add id (a, scl) accu in
List.fold_left2 fold Id.Map.empty 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
- 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
- let termlists = make_subst idsl argslist in
- let binders = make_subst idsbl bll in
- instantiate_notation_constr loc intern lvar
- (terms, termlists, binders) (Id.Map.empty, env) c
+let intern_notation intern env ntnvars loc ntn fullargs =
+ (* Adjust to parsing of { } *)
+ let ntn,fullargs = contract_curly_brackets ntn fullargs in
+ (* Recover interpretation { } *)
+ 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;
+ (* Dispatch parsing substitution to an interpretation substitution *)
+ let subst = split_by_type ids fullargs in
+ (* Instantiate the notation *)
+ instantiate_notation_constr loc intern intern_cases_pattern_as_binder ntnvars subst (Id.Map.empty, env) c
(**********************************************************************)
(* Discriminating between bound variables and global references *)
@@ -762,38 +885,41 @@ let string_of_ty = function
| Variable -> "var"
let gvar (loc, id) us = match us with
-| None -> GVar (loc, id)
+| None -> DAst.make ?loc @@ GVar id
| Some _ ->
- user_err_loc (loc, "", str "Variable " ++ pr_id id ++
+ user_err ?loc (str "Variable " ++ Id.print id ++
str " cannot have a universe instance")
-let intern_var genv (ltacvars,ntnvars) namedctx loc id us =
- (* Is [id] an inductive type potentially with implicit *)
+let intern_var env (ltacvars,ntnvars) namedctx loc id us =
+ (* Is [id] a notation variable *)
+ if Id.Map.mem id ntnvars then
+ begin
+ if not (Id.Map.mem id env.impls) then set_var_scope ?loc id true (env.tmp_scope,env.scopes) ntnvars;
+ gvar (loc,id) us, [], [], []
+ end
+ else
+ (* Is [id] registered with implicit arguments *)
try
- let ty,expl_impls,impls,argsc = Id.Map.find id genv.impls in
+ let ty,expl_impls,impls,argsc = Id.Map.find id env.impls in
let expl_impls = List.map
- (fun id -> CRef (Ident (loc,id),None), Some (loc,ExplByName id)) expl_impls in
+ (fun id -> CAst.make ?loc @@ CRef (make ?loc @@ Ident id,None), Some (make ?loc @@ ExplByName id)) expl_impls in
let tys = string_of_ty ty in
- Dumpglob.dump_reference loc "<>" (Id.to_string id) tys;
+ Dumpglob.dump_reference ?loc "<>" (Id.to_string id) tys;
gvar (loc,id) us, make_implicits_list impls, argsc, expl_impls
with Not_found ->
(* Is [id] bound in current term or is an ltac var bound to constr *)
- if Id.Set.mem id genv.ids || Id.Set.mem id ltacvars.ltac_vars
+ if Id.Set.mem id env.ids || Id.Set.mem id ltacvars.ltac_vars
then
gvar (loc,id) us, [], [], []
- (* Is [id] a notation variable *)
- else if Id.Map.mem id ntnvars
- then
- (set_var_scope loc id true genv ntnvars; gvar (loc,id) us, [], [], [])
- (* Is [id] the special variable for recursive notations *)
else if Id.equal id ldots_var
+ (* Is [id] the special variable for recursive notations? *)
then if Id.Map.is_empty ntnvars
- then error_ldots_var loc
+ then error_ldots_var ?loc
else gvar (loc,id) us, [], [], []
else if Id.Set.mem id ltacvars.ltac_bound then
(* Is [id] bound to a free name in ltac (this is an ltac error message) *)
- user_err_loc (loc,"intern_var",
- str "variable " ++ pr_id id ++ str " should be bound to a term.")
+ user_err ?loc ~hdr:"intern_var"
+ (str "variable " ++ Id.print id ++ str " should be bound to a term.")
else
(* Is [id] a goal or section variable *)
let _ = Context.Named.lookup id namedctx in
@@ -803,29 +929,32 @@ let intern_var genv (ltacvars,ntnvars) namedctx loc id us =
let ref = VarRef id in
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, []
+ Dumpglob.dump_reference ?loc "<>" (string_of_qualid (Decls.variable_secpath id)) "var";
+ DAst.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 c with
- | GRef (loc,ref,_) as x ->
+ match DAst.get 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
- when l != [] && Flags.version_strictly_greater Flags.V8_2 ->
+ c, impls, scopes, []
+ | GApp (r, l) ->
+ begin match DAst.get r with
+ | GRef (ref,_) when l != [] ->
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,[],[],[]
+ end
+ | _ -> c,[],[],[]
-let error_not_enough_arguments loc =
- user_err_loc (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
@@ -833,82 +962,106 @@ let check_no_explicitation l =
match l with
| [] -> ()
| (_, None) :: _ -> assert false
- | (_, Some (loc, _)) :: _ ->
- user_err_loc (loc,"",str"Unexpected explicitation of the argument of an abbreviation.")
+ | (_, Some {loc}) :: _ ->
+ user_err ?loc (str"Unexpected explicitation of the argument of an abbreviation.")
let dump_extended_global loc = function
- | TrueGlobal ref -> (*feedback_global loc ref;*) Dumpglob.add_glob loc ref
- | SynDef sp -> Dumpglob.add_glob_kn loc sp
+ | TrueGlobal ref -> (*feedback_global loc ref;*) Dumpglob.add_glob ?loc ref
+ | SynDef sp -> Dumpglob.add_glob_kn ?loc sp
-let intern_extended_global_of_qualid (loc,qid) =
+let intern_extended_global_of_qualid {loc;v=qid} =
let r = Nametab.locate_extended qid in dump_extended_global loc r; r
let intern_reference ref =
let qid = qualid_of_reference ref in
let r =
try intern_extended_global_of_qualid qid
- with Not_found -> error_global_not_found_loc (fst qid) (snd qid)
+ with Not_found -> error_global_not_found qid
in
Smartlocate.global_of_extended_global r
+let sort_info_of_level_info (info: Misctypes.level_info) : (Libnames.reference * int) option =
+ match info with
+ | Misctypes.UAnonymous -> None
+ | Misctypes.UUnknown -> None
+ | Misctypes.UNamed id -> Some (id, 0)
+
+let glob_sort_of_level (level: Misctypes.glob_level) : Misctypes.glob_sort =
+ match level with
+ | Misctypes.GProp -> Misctypes.GProp
+ | Misctypes.GSet -> Misctypes.GSet
+ | Misctypes.GType info -> Misctypes.GType [sort_info_of_level_info info]
+
(* 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
+let intern_qualid ?(no_secvar=false) qid intern env ntnvars us args =
+ let loc = qid.loc in
+ match intern_extended_global_of_qualid qid with
+ | TrueGlobal (VarRef _) when no_secvar ->
+ (* Rule out section vars since these should have been found by intern_var *)
+ raise Not_found
+ | TrueGlobal ref -> (DAst.make ?loc @@ GRef (ref, us)), true, args
| SynDef sp ->
- let (ids,c) = Syntax_def.search_syntactic_definition sp in
+ let (ids,c) = Syntax_def.search_syntactic_definition ?loc 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
- let subst = (terms, Id.Map.empty, Id.Map.empty) in
+ let subst = (terms, Id.Map.empty, Id.Map.empty, Id.Map.empty) in
let infos = (Id.Map.empty, env) in
let projapp = match c with NRef _ -> true | _ -> false in
- let c = instantiate_notation_constr loc intern lvar subst infos c in
- let c = match us, c with
+ let c = instantiate_notation_constr loc intern intern_cases_pattern_as_binder ntnvars subst infos c in
+ let loc = c.loc in
+ let err () =
+ user_err ?loc (str "Notation " ++ pr_qualid qid.v
+ ++ str " cannot have a universe instance,"
+ ++ str " its expanded head does not start with a reference")
+ in
+ let c = match us, DAst.get 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 _, _ ->
- user_err_loc (loc, "", str "Notation " ++ pr_qualid qid ++
- str " cannot have a universe instance, its expanded head
- does not start with a reference")
+ | Some _, GRef (ref, None) -> DAst.make ?loc @@ GRef (ref, us)
+ | Some _, GApp (r, arg) ->
+ let loc' = r.CAst.loc in
+ begin match DAst.get r with
+ | GRef (ref, None) ->
+ DAst.make ?loc @@ GApp (DAst.make ?loc:loc' @@ GRef (ref, us), arg)
+ | _ -> err ()
+ end
+ | Some [s], GSort (Misctypes.GType []) -> DAst.make ?loc @@ GSort (glob_sort_of_level s)
+ | Some [_old_level], GSort _new_sort ->
+ (* TODO: add old_level and new_sort to the error message *)
+ user_err ?loc (str "Cannot change universe level of notation " ++ pr_qualid qid.v)
+ | Some _, _ -> err ()
in
c, projapp, args2
-(* 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
- | r -> r
-
-let intern_applied_reference intern env namedctx (_, ntnvars as lvar) us args = function
- | Qualid (loc, qid) ->
+let intern_applied_reference intern env namedctx (_, ntnvars as lvar) us args =
+function
+ | {loc; v=Qualid qid} ->
+ let qid = make ?loc qid in
let r,projapp,args2 =
- try intern_qualid loc qid intern env ntnvars us args
- with Not_found -> error_global_not_found_loc loc qid
+ try intern_qualid qid intern env ntnvars us args
+ with Not_found -> error_global_not_found qid
in
let x, imp, scopes, l = find_appl_head_data r in
(x,imp,scopes,l), args2
- | Ident (loc, id) ->
+ | {loc; v=Ident id} ->
try intern_var env lvar namedctx loc id us, args
with Not_found ->
- let qid = qualid_of_ident id in
+ let qid = make ?loc @@ qualid_of_ident id in
try
- let r, projapp, args2 = intern_non_secvar_qualid loc qid intern env ntnvars us args in
+ let r, projapp, args2 = intern_qualid ~no_secvar:true qid intern env ntnvars us args in
let x, imp, scopes, l = find_appl_head_data r in
(x,imp,scopes,l), args2
with Not_found ->
(* Extra allowance for non globalizing functions *)
if !interning_grammar || env.unb then
(gvar (loc,id) us, [], [], []), args
- else error_global_not_found_loc loc qid
+ else error_global_not_found qid
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
@@ -917,7 +1070,17 @@ let interp_reference vars r =
(**********************************************************************)
(** {5 Cases } *)
-(** {6 Elemtary bricks } *)
+(** Private internalization patterns *)
+type 'a raw_cases_pattern_expr_r =
+ | RCPatAlias of 'a raw_cases_pattern_expr * Misctypes.lname
+ | RCPatCstr of Globnames.global_reference
+ * 'a raw_cases_pattern_expr list * 'a raw_cases_pattern_expr list
+ (** [RCPatCstr (loc, c, l1, l2)] represents [((@ c l1) l2)] *)
+ | RCPatAtom of (Misctypes.lident * (Notation_term.tmp_scope_name option * Notation_term.scope_name list)) option
+ | RCPatOr of 'a raw_cases_pattern_expr list
+and 'a raw_cases_pattern_expr = ('a raw_cases_pattern_expr_r, 'a) DAst.t
+
+(** {6 Elementary bricks } *)
let apply_scope_env env = function
| [] -> {env with tmp_scope = None}, []
| sc::scl -> {env with tmp_scope = sc}, scl
@@ -947,17 +1110,6 @@ let find_remaining_scopes pl1 pl2 ref =
in ((try List.firstn len_pl1 allscs with Failure _ -> simple_adjust_scopes len_pl1 allscs),
simple_adjust_scopes len_pl2 (aux (impl_list,scope_list)))
-let merge_subst s1 s2 = Id.Map.fold Id.Map.add s1 s2
-
-let product_of_cases_patterns ids idspl =
- List.fold_right (fun (ids,pl) (ids',ptaill) ->
- (ids @ ids',
- (* Cartesian prod of the or-pats for the nth arg and the tail args *)
- List.flatten (
- List.map (fun (subst,p) ->
- List.map (fun (subst',ptail) -> (merge_subst subst subst',p::ptail)) ptaill) pl)))
- idspl (ids,[Id.Map.empty,[]])
-
(* @return the first variable that occurs twice in a pattern
naive n^2 algo *)
@@ -965,8 +1117,11 @@ let rec has_duplicate = function
| [] -> None
| x::l -> if Id.List.mem x l then (Some x) else has_duplicate l
+let loc_of_multiple_pattern pl =
+ Loc.merge_opt (cases_pattern_expr_loc (List.hd pl)) (cases_pattern_expr_loc (List.last pl))
+
let loc_of_lhs lhs =
- Loc.merge (fst (List.hd lhs)) (fst (List.last lhs))
+ Loc.merge_opt (loc_of_multiple_pattern (List.hd lhs)) (loc_of_multiple_pattern (List.last lhs))
let check_linearity lhs ids =
match has_duplicate ids with
@@ -981,9 +1136,18 @@ let check_number_of_pattern loc n l =
if not (Int.equal n p) then raise (InternalizationError (loc,BadPatternsNumber (n,p)))
let check_or_pat_variables loc ids idsl =
- if List.exists (fun ids' -> not (List.eq_set Id.equal ids ids')) idsl then
- user_err_loc (loc, "", str
- "The components of this disjunctive pattern must bind the same variables.")
+ let eq_id {v=id} {v=id'} = Id.equal id id' in
+ (* Collect remaining patterns which do not have the same variables as the first pattern *)
+ let idsl = List.filter (fun ids' -> not (List.eq_set eq_id ids ids')) idsl in
+ match idsl with
+ | ids'::_ ->
+ (* Look for an [id] which is either in [ids] and not in [ids'] or in [ids'] and not in [ids] *)
+ let ids'' = List.subtract eq_id ids ids' in
+ let ids'' = if ids'' = [] then List.subtract eq_id ids' ids else ids'' in
+ user_err ?loc
+ (strbrk "The components of this disjunctive pattern must bind the same variables (" ++
+ Id.print (List.hd ids'').v ++ strbrk " is not bound in all patterns).")
+ | [] -> ()
(** Use only when params were NOT asked to the user.
@return if letin are included *)
@@ -991,9 +1155,48 @@ let check_constructor_length env loc cstr len_pl pl0 =
let n = len_pl + List.length pl0 in
if Int.equal n (Inductiveops.constructor_nallargs cstr) then false else
(Int.equal n (Inductiveops.constructor_nalldecls cstr) ||
- (error_wrong_numarg_constructor_loc loc env cstr
+ (error_wrong_numarg_constructor ?loc env cstr
(Inductiveops.constructor_nrealargs cstr)))
+open Term
+open Declarations
+
+(* Similar to Cases.adjust_local_defs but on RCPat *)
+let insert_local_defs_in_pattern (ind,j) l =
+ let (mib,mip) = Global.lookup_inductive ind in
+ if mip.mind_consnrealdecls.(j-1) = mip.mind_consnrealargs.(j-1) then
+ (* Optimisation *) l
+ else
+ let typi = mip.mind_nf_lc.(j-1) in
+ let (_,typi) = decompose_prod_n_assum (Context.Rel.length mib.mind_params_ctxt) typi in
+ let (decls,_) = decompose_prod_assum typi in
+ let rec aux decls args =
+ match decls, args with
+ | Context.Rel.Declaration.LocalDef _ :: decls, args -> (DAst.make @@ RCPatAtom None) :: aux decls args
+ | _, [] -> [] (* In particular, if there were trailing local defs, they have been inserted *)
+ | Context.Rel.Declaration.LocalAssum _ :: decls, a :: args -> a :: aux decls args
+ | _ -> assert false in
+ aux (List.rev decls) l
+
+let add_local_defs_and_check_length loc env g pl args = match g with
+ | ConstructRef cstr ->
+ (* We consider that no variables corresponding to local binders
+ have been given in the "explicit" arguments, which come from a
+ "@C args" notation or from a custom user notation *)
+ let pl' = insert_local_defs_in_pattern cstr pl in
+ let maxargs = Inductiveops.constructor_nalldecls cstr in
+ if List.length pl' + List.length args > maxargs then
+ error_wrong_numarg_constructor ?loc env cstr (Inductiveops.constructor_nrealargs cstr);
+ (* Two possibilities: either the args are given with explict
+ variables for local definitions, then we give the explicit args
+ extended with local defs, so that there is nothing more to be
+ added later on; or the args are not enough to have all arguments,
+ which a priori means local defs to add in the [args] part, so we
+ postpone the insertion of local defs in the explicit args *)
+ (* Note: further checks done later by check_constructor_length *)
+ if List.length pl' + List.length args = maxargs then pl' else pl
+ | _ -> pl
+
let add_implicits_check_length fail nargs nargs_with_letin impls_st len_pl1 pl2 =
let impl_list = if Int.equal len_pl1 0
then select_impargs_size (List.length pl2) impls_st
@@ -1005,10 +1208,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,RCPatAtom(Loc.ghost,None)::out)
+ then let (b,out) = aux i (q,[]) in (b,(DAst.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,RCPatAtom(Loc.ghost,None)::out)
+ then let (b,out) = aux i (q,l) in (b,(DAst.make @@ RCPatAtom None)::out)
else let (b,out) = aux (succ i) (q,tt) in (b,hh::out)
in aux 0 (impl_list,pl2)
@@ -1016,14 +1219,14 @@ let add_implicits_check_constructor_length env loc c len_pl1 pl2 =
let nargs = Inductiveops.constructor_nallargs c in
let nargs' = Inductiveops.constructor_nalldecls c in
let impls_st = implicits_of_global (ConstructRef c) in
- add_implicits_check_length (error_wrong_numarg_constructor_loc loc env c)
+ add_implicits_check_length (error_wrong_numarg_constructor ?loc env c)
nargs nargs' impls_st len_pl1 pl2
let add_implicits_check_ind_length env loc c len_pl1 pl2 =
let nallargs = inductive_nallargs_env env c in
let nalldecls = inductive_nalldecls_env env c in
let impls_st = implicits_of_global (IndRef c) in
- add_implicits_check_length (error_wrong_numarg_inductive_loc loc env c)
+ add_implicits_check_length (error_wrong_numarg_inductive ?loc env c)
nallargs nalldecls impls_st len_pl1 pl2
(** Do not raise NotEnoughArguments thanks to preconditions*)
@@ -1033,8 +1236,9 @@ 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') params;
+ List.iter (fun c -> match DAst.get c with
+ | PatVar Anonymous -> ()
+ | PatVar _ | PatCstr(_,_,_) -> error_parameter_not_implicit ?loc:c.CAst.loc) params;
args
let find_constructor loc add_params ref =
@@ -1042,10 +1246,10 @@ let find_constructor loc add_params ref =
| ConstructRef cstr -> cstr
| IndRef _ ->
let error = str "There is an inductive name deep in a \"in\" clause." in
- user_err_loc (loc, "find_constructor", error)
+ user_err ?loc ~hdr:"find_constructor" error
| ConstRef _ | VarRef _ ->
let error = str "This reference is not a constructor." in
- user_err_loc (loc, "find_constructor", error)
+ user_err ?loc ~hdr:"find_constructor" error
in
cstr, match add_params with
| Some nb_args ->
@@ -1054,12 +1258,12 @@ 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, DAst.make @@ PatVar Anonymous)])
| None -> []
let find_pattern_variable = function
- | Ident (loc,id) -> id
- | Qualid (loc,_) as x -> raise (InternalizationError(loc,NotAConstructor x))
+ | {v=Ident id} -> id
+ | {loc;v=Qualid _} as x -> raise (InternalizationError(loc,NotAConstructor x))
let check_duplicate loc fields =
let eq (ref1, _) (ref2, _) = eq_reference ref1 ref2 in
@@ -1067,7 +1271,7 @@ let check_duplicate loc fields =
match dups with
| [] -> ()
| (r, _) :: _ ->
- user_err_loc (loc, "", str "This record defines several times the field " ++
+ user_err ?loc (str "This record defines several times the field " ++
pr_reference r ++ str ".")
(** [sort_fields ~complete loc fields completer] expects a list
@@ -1092,17 +1296,17 @@ let sort_fields ~complete loc fields completer =
let gr = global_reference_of_reference first_field_ref in
(gr, Recordops.find_projection gr)
with Not_found ->
- user_err_loc (loc_of_reference first_field_ref, "intern",
- pr_reference first_field_ref ++ str": Not a projection")
+ user_err ?loc ~hdr:"intern"
+ (pr_reference first_field_ref ++ str": Not a projection")
in
(* the number of parameters *)
let nparams = record.Recordops.s_EXPECTEDPARAM in
(* the reference constructor of the record *)
let base_constructor =
let global_record_id = ConstructRef record.Recordops.s_CONST in
- try Qualid (loc, shortest_qualid_of_global Id.Set.empty global_record_id)
+ try make ?loc @@ Qualid (shortest_qualid_of_global Id.Set.empty global_record_id)
with Not_found ->
- anomaly (str "Environment corruption for records") in
+ anomaly (str "Environment corruption for records.") in
let () = check_duplicate loc fields in
let (end_index, (* one past the last field index *)
first_field_index, (* index of the first field of the record *)
@@ -1113,17 +1317,17 @@ let sort_fields ~complete loc fields completer =
let rec build_proj_list projs proj_kinds idx ~acc_first_idx acc =
match projs with
| [] -> (idx, acc_first_idx, acc)
- | (Some name) :: projs ->
- let field_glob_ref = ConstRef name in
+ | (Some field_glob_id) :: projs ->
+ let field_glob_ref = ConstRef field_glob_id in
let first_field = eq_gr field_glob_ref first_field_glob_ref in
begin match proj_kinds with
- | [] -> anomaly (Pp.str "Number of projections mismatch")
+ | [] -> anomaly (Pp.str "Number of projections mismatch.")
| (_, regular) :: proj_kinds ->
(* "regular" is false when the field is defined
by a let-in in the record declaration
(its value is fixed from other fields). *)
if first_field && not regular && complete then
- user_err_loc (loc, "", str "No local fields allowed in a record construction.")
+ user_err ?loc (str "No local fields allowed in a record construction.")
else if first_field then
build_proj_list projs proj_kinds (idx+1) ~acc_first_idx:idx acc
else if not regular && complete then
@@ -1131,12 +1335,12 @@ let sort_fields ~complete loc fields completer =
build_proj_list projs proj_kinds idx ~acc_first_idx acc
else
build_proj_list projs proj_kinds (idx+1) ~acc_first_idx
- ((idx, field_glob_ref) :: acc)
+ ((idx, field_glob_id) :: acc)
end
| None :: projs ->
if complete then
(* we don't want anonymous fields *)
- user_err_loc (loc, "", str "This record contains anonymous fields.")
+ user_err ?loc (str "This record contains anonymous fields.")
else
(* anonymous arguments don't appear in proj_kinds *)
build_proj_list projs proj_kinds (idx+1) ~acc_first_idx acc
@@ -1150,15 +1354,14 @@ let sort_fields ~complete loc fields completer =
| (field_ref, field_value) :: fields ->
let field_glob_ref = try global_reference_of_reference field_ref
with Not_found ->
- user_err_loc (loc_of_reference field_ref, "intern",
- str "The field \"" ++ pr_reference field_ref ++ str "\" does not exist.") in
+ user_err ?loc ~hdr:"intern"
+ (str "The field \"" ++ pr_reference field_ref ++ str "\" does not exist.") in
let remaining_projs, (field_index, _) =
- let the_proj (idx, glob_ref) = eq_gr field_glob_ref glob_ref in
+ let the_proj (idx, glob_id) = eq_gr field_glob_ref (ConstRef glob_id) in
try CList.extract_first the_proj remaining_projs
with Not_found ->
- user_err_loc
- (loc, "",
- str "This record contains fields of different records.")
+ user_err ?loc
+ (str "This record contains fields of different records.")
in
index_fields fields remaining_projs ((field_index, field_value) :: acc)
| [] ->
@@ -1181,7 +1384,7 @@ let sort_fields ~complete loc fields completer =
(** {6 Manage multiple aliases} *)
type alias = {
- alias_ids : Id.t list;
+ alias_ids : Misctypes.lident list;
alias_map : Id.t Id.Map.t;
}
@@ -1192,17 +1395,20 @@ let empty_alias = {
(* [merge_aliases] returns the sets of all aliases encountered at this
point and a substitution mapping extra aliases to the first one *)
-let merge_aliases aliases id =
- let alias_ids = aliases.alias_ids @ [id] in
+let merge_aliases aliases {loc;v=na} =
+ match na with
+ | Anonymous -> aliases
+ | Name id ->
+ let alias_ids = aliases.alias_ids @ [make ?loc id] in
let alias_map = match aliases.alias_ids with
| [] -> aliases.alias_map
- | id' :: _ -> Id.Map.add id id' aliases.alias_map
+ | {v=id'} :: _ -> Id.Map.add id id' aliases.alias_map
in
{ alias_ids; alias_map; }
let alias_of als = match als.alias_ids with
| [] -> Anonymous
-| id :: _ -> Name id
+| {v=id} :: _ -> Name id
(** {6 Expanding notations }
@@ -1212,16 +1418,42 @@ let alias_of als = match als.alias_ids with
*)
-let rec subst_pat_iterator y t p = match p with
- | RCPatAtom (_,id) ->
- begin match id with Some x when Id.equal x y -> t | _ -> p end
- | RCPatCstr (loc,id,l1,l2) ->
- RCPatCstr (loc,id,List.map (subst_pat_iterator y t) l1,
- List.map (subst_pat_iterator y t) l2)
- | RCPatAlias (l,p,a) -> RCPatAlias (l,subst_pat_iterator y t p,a)
- | RCPatOr (l,pl) -> RCPatOr (l,List.map (subst_pat_iterator y t) pl)
+let is_zero s =
+ let rec aux i =
+ Int.equal (String.length s) i || (s.[i] == '0' && aux (i+1))
+ in aux 0
-let drop_notations_pattern looked_for =
+let merge_subst s1 s2 = Id.Map.fold Id.Map.add s1 s2
+
+let product_of_cases_patterns aliases idspl =
+ (* each [pl] is a disjunction of patterns over common identifiers [ids] *)
+ (* We stepwise build a disjunction of patterns [ptaill] over common [ids'] *)
+ List.fold_right (fun (ids,pl) (ids',ptaill) ->
+ (ids @ ids',
+ (* Cartesian prod of the or-pats for the nth arg and the tail args *)
+ List.flatten (
+ List.map (fun (subst,p) ->
+ List.map (fun (subst',ptail) -> (merge_subst subst subst',p::ptail)) ptaill) pl)))
+ idspl (aliases.alias_ids,[aliases.alias_map,[]])
+
+let rec subst_pat_iterator y t = DAst.(map (function
+ | RCPatAtom id as p ->
+ begin match id with Some ({v=x},_) when Id.equal x y -> DAst.get t | _ -> p end
+ | RCPatCstr (id,l1,l2) ->
+ 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 is_non_zero c = match c with
+| { CAst.v = CPrim (Numeral (p, true)) } -> not (is_zero p)
+| _ -> false
+
+let is_non_zero_pat c = match c with
+| { CAst.v = CPatPrim (Numeral (p, true)) } -> not (is_zero p)
+| _ -> false
+
+let drop_notations_pattern looked_for genv =
(* At toplevel, Constructors and Inductives are accepted, in recursive calls
only constructor are allowed *)
let ensure_kind top loc g =
@@ -1229,15 +1461,31 @@ let drop_notations_pattern looked_for =
if top then looked_for g else
match g with ConstructRef _ -> () | _ -> raise Not_found
with Not_found ->
- error_invalid_pattern_notation loc
+ error_invalid_pattern_notation ?loc ()
in
let test_kind top =
if top then looked_for else function ConstructRef _ -> () | _ -> raise Not_found
in
+ (** [rcp_of_glob] : from [glob_constr] to [raw_cases_pattern_expr] *)
+ let rec rcp_of_glob scopes x = DAst.(map (function
+ | GVar id -> RCPatAtom (Some (CAst.make ?loc:x.loc id,scopes))
+ | GHole (_,_,_) -> RCPatAtom (None)
+ | GRef (g,_) -> RCPatCstr (g,[],[])
+ | GApp (r, l) ->
+ begin match DAst.get r with
+ | GRef (g,_) ->
+ let allscs = find_arguments_scope g in
+ let allscs = simple_adjust_scopes (List.length l) allscs in (* TO CHECK *)
+ RCPatCstr (g, List.map2 (fun sc a -> rcp_of_glob (sc,snd scopes) a) allscs l,[])
+ | _ ->
+ CErrors.anomaly Pp.(str "Invalid return pattern from Notation.interp_prim_token_cases_pattern_expr.")
+ end
+ | _ -> CErrors.anomaly Pp.(str "Invalid return pattern from Notation.interp_prim_token_cases_pattern_expr."))) x
+ in
let rec drop_syndef top scopes re pats =
- let (loc,qid) = qualid_of_reference re in
+ let qid = qualid_of_reference re in
try
- match locate_extended qid with
+ match locate_extended qid.v with
| SynDef sp ->
let (vars,a) = Syntax_def.search_syntactic_definition sp in
(match a with
@@ -1255,80 +1503,96 @@ 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:qid.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
+ let idspl1 = List.map (in_not false qid.loc scopes (subst, Id.Map.empty) []) args in
let (_,argscs) = find_remaining_scopes pats1 pats2 g in
Some (g, idspl1, List.map2 (in_pat_sc scopes) argscs pats2)
| _ -> raise Not_found)
| TrueGlobal g ->
test_kind top g;
- Dumpglob.add_glob loc g;
+ Dumpglob.add_glob ?loc:qid.loc g;
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 = function
- | CPatAlias (loc, p, id) -> RCPatAlias (loc, in_pat top scopes p, id)
- | CPatRecord (loc, l) ->
+ and in_pat top scopes pt =
+ let open CAst in
+ let loc = pt.loc in
+ match pt.v with
+ | CPatAlias (p, id) -> DAst.make ?loc @@ RCPatAlias (in_pat top scopes p, id)
+ | CPatRecord l ->
let sorted_fields =
- sort_fields ~complete:false loc l (fun _idx -> (CPatAtom (loc, None))) in
+ sort_fields ~complete:false loc l (fun _idx -> CAst.make ?loc @@ CPatAtom None) in
begin match sorted_fields with
- | None -> RCPatAtom (loc, None)
+ | None -> DAst.make ?loc @@ RCPatAtom None
| Some (n, head, pl) ->
let pl =
if !asymmetric_patterns then pl else
- let pars = List.make n (CPatAtom (loc, 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) -> RCPatCstr(loc, a, b, c)
- |None -> raise (InternalizationError (loc,NotAConstructor head))
+ | Some (a,b,c) -> DAst.make ?loc @@ RCPatCstr(a, b, c)
+ | None -> raise (InternalizationError (loc,NotAConstructor head))
end
- | CPatCstr (loc, head, None, pl) ->
+ | CPatCstr (head, None, pl) ->
begin
match drop_syndef top scopes head pl with
- | Some (a,b,c) -> RCPatCstr(loc, a, b, c)
+ | Some (a,b,c) -> DAst.make ?loc @@ RCPatCstr(a, b, c)
| None -> raise (InternalizationError (loc,NotAConstructor head))
end
- | CPatCstr (loc, r, Some expl_pl, pl) ->
- let g = try locate (snd (qualid_of_reference r))
+ | CPatCstr (r, Some expl_pl, pl) ->
+ let g = try locate (qualid_of_reference r).v
with Not_found ->
raise (InternalizationError (loc,NotAConstructor r)) in
if expl_pl == [] then
(* Convention: (@r) deactivates all further implicit arguments and scopes *)
- RCPatCstr (loc, g, List.map (in_pat false scopes) pl, [])
+ DAst.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
- RCPatCstr (loc, 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)
- | CPatNotation (_,"( _ )",([a],[]),[]) ->
+ DAst.make ?loc @@ RCPatCstr (g, List.map2 (in_pat_sc scopes) argscs1 expl_pl @ List.map (in_pat false scopes) pl, [])
+ | CPatNotation ("- _",([a],[]),[]) when is_non_zero_pat a ->
+ let p = match a.CAst.v with CPatPrim (Numeral (p, _)) -> p | _ -> assert false in
+ let pat, _df = Notation.interp_prim_token_cases_pattern_expr ?loc (ensure_kind false loc) (Numeral (p,false)) scopes in
+ rcp_of_glob scopes pat
+ | CPatNotation ("( _ )",([a],[]),[]) ->
in_pat top scopes a
- | CPatNotation (loc, 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',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 (loc, key, e) ->
- in_pat top (None,find_delimiters_scope loc key::snd scopes) e
- | CPatPrim (loc,p) -> fst (Notation.interp_prim_token_cases_pattern_expr loc (test_kind false) p scopes)
- | CPatAtom (loc, Some id) ->
+ | CPatNotation (ntn,fullargs,extrargs) ->
+ let ntn,(terms,termlists) = contract_curly_brackets_pat ntn fullargs in
+ let ((ids',c),df) = Notation.interp_notation ?loc ntn scopes in
+ let (terms,termlists) = split_by_type_pat ?loc ids' (terms,termlists) in
+ Dumpglob.dump_notation_location (patntn_loc ?loc fullargs ntn) ntn df;
+ in_not top loc scopes (terms,termlists) extrargs c
+ | CPatDelimiters (key, e) ->
+ in_pat top (None,find_delimiters_scope ?loc key::snd scopes) e
+ | CPatPrim p ->
+ let pat, _df = Notation.interp_prim_token_cases_pattern_expr ?loc (test_kind false) p scopes in
+ rcp_of_glob scopes pat
+ | CPatAtom (Some id) ->
begin
match drop_syndef top scopes id [] with
- |Some (a,b,c) -> RCPatCstr (loc, a, b, c)
- |None -> RCPatAtom (loc, Some (find_pattern_variable id))
+ | Some (a,b,c) -> DAst.make ?loc @@ RCPatCstr (a, b, c)
+ | None -> DAst.make ?loc @@ RCPatAtom (Some ((make ?loc @@ find_pattern_variable id),scopes))
end
- | CPatAtom (loc,None) -> RCPatAtom (loc,None)
- | CPatOr (loc, pl) ->
- RCPatOr (loc,List.map (in_pat top scopes) pl)
- | CPatCast _ ->
- assert false
+ | CPatAtom None -> DAst.make ?loc @@ RCPatAtom None
+ | CPatOr pl -> DAst.make ?loc @@ RCPatOr (List.map (in_pat top scopes) pl)
+ | CPatCast (_,_) ->
+ (* We raise an error if the pattern contains a cast, due to
+ current restrictions on casts in patterns. Cast in patterns
+ are supportted only in local binders and only at top
+ level. In fact, they are currently eliminated by the
+ parser. The only reason why they are in the
+ [cases_pattern_expr] type is that the parser needs to factor
+ the "(c : t)" notation with user defined notations (such as
+ the pair). In the long term, we will try to support such
+ casts everywhere, and use them to print the domains of
+ lambdas in the encoding of match in constr. This check is
+ here and not in the parser because it would require
+ duplicating the levels of the [pattern] rule. *)
+ CErrors.user_err ?loc ~hdr:"drop_notations_pattern"
+ (Pp.strbrk "Casts are not supported in this pattern.")
and in_pat_sc scopes x = in_pat false (x,snd scopes)
and in_not top loc scopes (subst,substlist as fullsubst) args = function
| NVar id ->
@@ -1340,22 +1604,22 @@ 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 RCPatAtom (loc,Some id) else
- anomaly (str "Unbound pattern notation variable: " ++ Id.print id)
+ if Id.equal id ldots_var then DAst.make ?loc @@ RCPatAtom (Some ((make ?loc id),scopes)) else
+ anomaly (str "Unbound pattern notation variable: " ++ Id.print id ++ str ".")
end
| NRef g ->
ensure_kind top loc g;
let (_,argscs) = find_remaining_scopes [] args g in
- RCPatCstr (loc, g, [], List.map2 (in_pat_sc scopes) argscs args)
+ DAst.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
- RCPatCstr (loc, 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) ->
- if not (List.is_empty args) then user_err_loc
- (loc,"",strbrk "Application of arguments to a recursive notation not supported in patterns.");
+ let pl = List.map2 (fun x -> in_not false loc (x,snd scopes) fullsubst []) argscs1 pl in
+ let pl = add_local_defs_and_check_length loc genv g pl args in
+ DAst.make ?loc @@ RCPatCstr (g, pl @ List.map (in_pat false scopes) args, [])
+ | NList (x,y,iter,terminator,revert) ->
+ if not (List.is_empty args) then user_err ?loc
+ (strbrk "Application of arguments to a recursive notation not supported in patterns.");
(try
(* All elements of the list are in scopes (scopt,subscopes) *)
let (l,(scopt,subscopes)) = Id.Map.find x substlist in
@@ -1364,27 +1628,28 @@ let drop_notations_pattern looked_for =
let nsubst = Id.Map.add y (a, (scopt, subscopes)) subst in
let u = in_not false loc scopes (nsubst, substlist) [] iter in
subst_pat_iterator ldots_var t u)
- (if lassoc then List.rev l else l) termin
+ (if revert then List.rev l else l) termin
with Not_found ->
- anomaly (Pp.str "Inconsistent substitution of recursive notation"))
+ anomaly (Pp.str "Inconsistent substitution of recursive notation."))
| NHole _ ->
let () = assert (List.is_empty args) in
- RCPatAtom (loc, None)
- | t -> error_invalid_pattern_notation loc
+ DAst.make ?loc @@ RCPatAtom None
+ | t -> error_invalid_pattern_notation ?loc ()
in in_pat true
-let rec intern_pat genv aliases pat =
+let rec intern_pat genv ntnvars aliases pat =
let intern_cstr_with_all_args loc c with_letin idslpl1 pl2 =
- 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 idslpl2 = List.map (intern_pat genv ntnvars empty_alias) pl2 in
+ let (ids',pll) = product_of_cases_patterns aliases (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, DAst.make ?loc @@ PatCstr (c,chop_params_pattern loc (fst c) pl with_letin,alias_of aliases))) pll in
ids',pl' in
- match pat with
- | RCPatAlias (loc, p, id) ->
+ let loc = pat.loc in
+ match DAst.get pat with
+ | RCPatAlias (p, id) ->
let aliases' = merge_aliases aliases id in
- intern_pat genv aliases' p
- | RCPatCstr (loc, head, expl_pl, pl) ->
+ intern_pat genv ntnvars aliases' p
+ | 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
@@ -1396,92 +1661,59 @@ 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)
- | RCPatAtom (loc, Some id) ->
- let aliases = merge_aliases aliases id in
- (aliases.alias_ids,[aliases.alias_map, PatVar (loc, alias_of aliases)])
- | RCPatAtom (loc, None) ->
+ | RCPatAtom (Some ({loc;v=id},scopes)) ->
+ let aliases = merge_aliases aliases (make ?loc @@ Name id) in
+ set_var_scope ?loc id false scopes ntnvars;
+ (aliases.alias_ids,[aliases.alias_map, DAst.make ?loc @@ PatVar (alias_of aliases)]) (* TO CHECK: aura-t-on id? *)
+ | RCPatAtom (None) ->
let { alias_ids = ids; alias_map = asubst; } = aliases in
- (ids, [asubst, PatVar (loc, alias_of aliases)])
- | RCPatOr (loc, pl) ->
+ (ids, [asubst, DAst.make ?loc @@ PatVar (alias_of aliases)])
+ | RCPatOr pl ->
assert (not (List.is_empty pl));
- let pl' = List.map (intern_pat genv aliases) pl in
+ let pl' = List.map (intern_pat genv ntnvars aliases) pl in
let (idsl,pl') = List.split pl' in
let ids = List.hd idsl in
check_or_pat_variables loc ids (List.tl idsl);
(ids,List.flatten pl')
-(* [check_no_patcast p] raises an error if [p] contains a cast. This code is a
- bit ad-hoc, and is due to current restrictions on casts in patterns. We
- support them only in local binders and only at top level. In fact, they are
- currently eliminated by the parser. The only reason why they are in the
- [cases_pattern_expr] type is that the parser needs to factor the "(c : t)"
- notation with user defined notations (such as the pair). In the long term, we
- will try to support such casts everywhere, and use them to print the domains
- 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 = function
- | CPatCast (loc,_,_) ->
- CErrors.user_err_loc (loc, "check_no_patcast",
- Pp.strbrk "Casts are not supported here.")
- | CPatDelimiters(_,_,p)
- | CPatAlias(_,p,_) -> check_no_patcast p
- | CPatCstr(_,_,opl,pl) ->
- Option.iter (List.iter check_no_patcast) opl;
- List.iter check_no_patcast pl
- | CPatOr(_,pl) ->
- List.iter check_no_patcast pl
- | CPatNotation(_,_,subst,pl) ->
- check_no_patcast_subst subst;
- List.iter check_no_patcast pl
- | CPatRecord(_,prl) ->
- List.iter (fun (_,p) -> check_no_patcast p) prl
- | CPatAtom _ | CPatPrim _ -> ()
-
-and check_no_patcast_subst (pl,pll) =
- List.iter check_no_patcast pl;
- List.iter (List.iter check_no_patcast) pll
-
-let intern_cases_pattern genv scopes aliases pat =
- check_no_patcast pat;
- intern_pat genv aliases
- (drop_notations_pattern (function ConstructRef _ -> () | _ -> raise Not_found) scopes pat)
+let intern_cases_pattern genv ntnvars scopes aliases pat =
+ intern_pat genv ntnvars aliases
+ (drop_notations_pattern (function ConstructRef _ -> () | _ -> raise Not_found) genv scopes pat)
let _ =
intern_cases_pattern_fwd :=
- fun scopes p -> intern_cases_pattern (Global.env ()) scopes empty_alias p
+ fun ntnvars scopes p -> intern_cases_pattern (Global.env ()) ntnvars scopes empty_alias p
-let intern_ind_pattern genv scopes pat =
- check_no_patcast pat;
+let intern_ind_pattern genv ntnvars scopes pat =
let no_not =
try
- drop_notations_pattern (function (IndRef _ | ConstructRef _) -> () | _ -> raise Not_found) scopes pat
- with InternalizationError(loc,NotAConstructor _) -> error_bad_inductive_type loc
+ drop_notations_pattern (function (IndRef _ | ConstructRef _) -> () | _ -> raise Not_found) genv scopes pat
+ with InternalizationError(loc,NotAConstructor _) -> error_bad_inductive_type ?loc
in
- match no_not with
- | RCPatCstr (loc, head, expl_pl, pl) ->
- let c = (function IndRef ind -> ind | _ -> error_bad_inductive_type loc) head in
+ let loc = no_not.CAst.loc in
+ match DAst.get no_not 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
- let idslpl1 = List.rev_map (intern_pat genv empty_alias) expl_pl in
- let idslpl2 = List.map (intern_pat genv empty_alias) pl2 in
+ let idslpl = List.map (intern_pat genv ntnvars empty_alias) (expl_pl@pl2) in
(with_letin,
- match product_of_cases_patterns [] (List.rev_append idslpl1 idslpl2) with
+ match product_of_cases_patterns empty_alias idslpl with
| _,[_,pl] -> (c,chop_params_pattern loc c pl with_letin)
- | _ -> error_bad_inductive_type loc)
- | x -> error_bad_inductive_type (raw_cases_pattern_expr_loc x)
+ | _ -> error_bad_inductive_type ?loc)
+ | x -> error_bad_inductive_type ?loc
(**********************************************************************)
(* Utilities for application *)
let merge_impargs l args =
let test x = function
- | (_, Some (_, y)) -> explicitation_eq x y
+ | (_, Some {v=y}) -> explicitation_eq x y
| _ -> false
in
List.fold_right (fun a l ->
match a with
- | (_,Some (_,(ExplByName id as x))) when
+ | (_, Some {v=ExplByName id as x}) when
List.exists (test x) args -> l
| _ -> a::l)
l args
@@ -1489,10 +1721,19 @@ let merge_impargs l args =
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)
- | _ -> anomaly (Pp.str "Only refs have implicits")
+let set_hole_implicit i b c =
+ let loc = c.CAst.loc in
+ match DAst.get c with
+ | GRef (r, _) -> Loc.tag ?loc (Evar_kinds.ImplicitArg (r,i,b),Misctypes.IntroAnonymous,None)
+ | GApp (r, _) ->
+ let loc = r.CAst.loc in
+ begin match DAst.get r with
+ | GRef (r, _) ->
+ Loc.tag ?loc (Evar_kinds.ImplicitArg (r,i,b),Misctypes.IntroAnonymous,None)
+ | _ -> anomaly (Pp.str "Only refs have implicits.")
+ end
+ | 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 =
List.exists (fun imp -> is_status_implicit imp && Id.equal id (name_of_implicit imp))
@@ -1504,14 +1745,14 @@ let extract_explicit_arg imps args =
let (eargs,rargs) = aux l in
match e with
| None -> (eargs,a::rargs)
- | Some (loc,pos) ->
+ | Some {loc;v=pos} ->
let id = match pos with
| ExplByName id ->
if not (exists_implicit_name id imps) then
- user_err_loc
- (loc,"",str "Wrong argument name: " ++ pr_id id ++ str ".");
+ user_err ?loc
+ (str "Wrong argument name: " ++ Id.print id ++ str ".");
if Id.Map.mem id eargs then
- user_err_loc (loc,"",str "Argument name " ++ pr_id id
+ user_err ?loc (str "Argument name " ++ Id.print id
++ str " occurs more than once.");
id
| ExplByPos (p,_id) ->
@@ -1521,11 +1762,11 @@ let extract_explicit_arg imps args =
if not (is_status_implicit imp) then failwith "imp";
name_of_implicit imp
with Failure _ (* "nth" | "imp" *) ->
- user_err_loc
- (loc,"",str"Wrong argument position: " ++ int p ++ str ".")
+ user_err ?loc
+ (str"Wrong argument position: " ++ int p ++ str ".")
in
if Id.Map.mem id eargs then
- user_err_loc (loc,"",str"Argument at position " ++ int p ++
+ user_err ?loc (str"Argument at position " ++ int p ++
str " is mentioned more than once.");
id in
(Id.Map.add id (loc, a) eargs, rargs)
@@ -1534,17 +1775,17 @@ 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 internalize globalenv env pattern_mode (_, ntnvars as lvar) c =
+ 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)
- 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) ->
- let lf = List.map (fun ((_, id),_,_,_,_) -> id) dl in
+ | CFix ({ CAst.loc = locid; v = iddef}, dl) ->
+ let lf = List.map (fun ({CAst.v = id},_,_,_,_) -> id) dl in
let dl = Array.of_list dl in
let n =
try List.index0 Id.equal iddef lf
@@ -1556,10 +1797,11 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
let intern_ro_arg f =
let before, after = split_at_annot bl n in
let (env',rbefore) = List.fold_left intern_local_binder (env,[]) before in
- let rbefore = List.map (function BDRawDef a -> a | BDPattern _ -> assert false) rbefore in
let ro = f (intern env') in
- let n' = Option.map (fun _ -> List.count (fun (_,(_,_,b,_)) -> (* remove let-ins *) b = None) rbefore) n in
- let rbefore = List.map (fun a -> BDRawDef a) rbefore in
+ let n' = Option.map (fun _ -> List.count (fun c -> match DAst.get c with
+ | 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) =
@@ -1571,28 +1813,24 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
| CMeasureRec (m,r) ->
intern_ro_arg (fun f -> GMeasureRec (f m, Option.map f r))
in
- let bl =
- List.rev_map
- (function
- | BDRawDef a -> a
- | BDPattern (loc,_,_,_,_) ->
- Loc.raise loc (Stream.Error "pattern with quote not allowed after fix")) rbl in
- ((n, ro), bl, intern_type env' ty, env')) dl in
+ let bl = List.rev (List.map glob_local_binder_of_extended rbl) in
+ ((n, ro), bl, intern_type env' ty, env')) dl in
let idl = Array.map2 (fun (_,_,_,_,bd) (a,b,c,env') ->
let env'' = List.fold_left_i (fun i en name ->
let (_,bli,tyi,_) = idl_temp.(i) in
- let fix_args = (List.map (fun (_,(na, bk, _, _)) -> (build_impls bk na)) bli) 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 (CAst.make @@ Name name)) 0 env' lf in
(a,b,c,intern {env'' with tmp_scope = None} bd)) dl idl_temp in
- GRec (loc,GFix
+ DAst.make ?loc @@
+ GRec (GFix
(Array.map (fun (ro,_,_,_) -> ro) idl,n),
Array.of_list lf,
- Array.map (fun (_,bl,_,_) -> List.map snd bl) idl,
+ Array.map (fun (_,bl,_,_) -> bl) idl,
Array.map (fun (_,_,ty,_) -> ty) idl,
Array.map (fun (_,_,_,bd) -> bd) idl)
- | CCoFix (loc, (locid,iddef), dl) ->
- let lf = List.map (fun ((_, id),_,_,_) -> id) dl in
+ | CCoFix ({ CAst.loc = locid; v = iddef }, dl) ->
+ let lf = List.map (fun ({CAst.v = id},_,_,_) -> id) dl in
let dl = Array.of_list dl in
let n =
try List.index0 Id.equal iddef lf
@@ -1600,96 +1838,100 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
raise (InternalizationError (locid,UnboundFixName (true,iddef)))
in
let idl_tmp = Array.map
- (fun ((loc,id),bl,ty,_) ->
+ (fun ({ CAst.loc; v = id },bl,ty,_) ->
let (env',rbl) = List.fold_left intern_local_binder (env,[]) bl in
- let rbl = List.map (function BDRawDef a -> a | BDPattern _ -> assert false) rbl in
- (List.rev rbl,
+ (List.rev (List.map glob_local_binder_of_extended rbl),
intern_type env' ty,env')) dl in
let idl = Array.map2 (fun (_,_,_,bd) (b,c,env') ->
let env'' = List.fold_left_i (fun i en name ->
let (bli,tyi,_) = idl_tmp.(i) in
- let cofix_args = List.map (fun (_, (na, bk, _, _)) -> (build_impls bk na)) bli 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 (CAst.make @@ Name name)) 0 env' lf in
(b,c,intern {env'' with tmp_scope = None} bd)) dl idl_tmp in
- GRec (loc,GCoFix n,
+ DAst.make ?loc @@
+ GRec (GCoFix n,
Array.of_list lf,
- Array.map (fun (bl,_,_) -> List.map snd bl) idl,
+ Array.map (fun (bl,_,_) -> bl) idl,
Array.map (fun (_,ty,_) -> ty) idl,
Array.map (fun (_,_,bd) -> bd) idl)
- | CProdN (loc,[],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 (bl,c2) ->
+ let (env',bl) = List.fold_left intern_local_binder (env,[]) bl in
+ expand_binders ?loc mkGProd bl (intern_type env' c2)
+ | CLambdaN ([],c2) ->
+ (* Such a term is built sometimes: it should not change scope *)
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,c2) ->
+ | CLambdaN (bl,c2) ->
+ let (env',bl) = List.fold_left intern_local_binder (reset_tmp_scope env,[]) bl in
+ expand_binders ?loc mkGLambda bl (intern env' c2)
+ | CLetIn (na,c1,t,c2) ->
let inc1 = intern (reset_tmp_scope env) c1 in
- GLetIn (loc, snd na, inc1,
+ let int = Option.map (intern_type env) t in
+ DAst.make ?loc @@
+ GLetIn (na.CAst.v, inc1, int,
intern (push_name_env ntnvars (impls_term_list inc1) env na) c2)
- | CNotation (loc,"- _",([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) ->
+ | CNotation ("- _", ([a],[],[],[])) when is_non_zero a ->
+ let p = match a.CAst.v with CPrim (Numeral (p, _)) -> p | _ -> assert false in
+ intern env (CAst.make ?loc @@ CPrim (Numeral (p,false)))
+ | 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) ->
- fst (Notation.interp_prim_token loc p (env.tmp_scope,env.scopes))
- | CDelimiters (loc, key, e) ->
+ | CPrim p ->
+ 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
- | CAppExpl (loc, (isproj,ref,us), args) ->
+ 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
intern_applied_reference intern env (Environ.named_context globalenv)
lvar us args ref
in
(* Rem: GApp(_,f,[]) stands for @f *)
- GApp (loc, f, intern_args env args_scopes (List.map fst args))
+ DAst.make ?loc @@
+ GApp (f, intern_args env args_scopes (List.map fst args))
- | CApp (loc, (isproj,f), args) ->
- let f,args = match f with
+ | CApp ((isproj,f), args) ->
+ let f,args = match f.CAst.v with
(* Compact notations like "t.(f args') args" *)
- | CApp (_,(Some _,f), args') when not (Option.has_some isproj) ->
+ | 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
+ match f.CAst.v with
| CRef (ref,us) ->
intern_applied_reference intern env
(Environ.named_context globalenv) lvar us args ref
- | CNotation (loc,ntn,([],[],[])) ->
- let c = intern_notation intern env ntnvars loc ntn ([],[],[]) in
+ | 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
+ | _ -> (intern env f,[],[],[]), args in
+ 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 -> CAst.make ?loc @@ CHole (Some (Evar_kinds.QuestionMark (st,Anonymous)),
+ Misctypes.IntroAnonymous, None))
in
begin
match fields with
- | None -> user_err_loc (loc, "intern", str"No constructor inference.")
+ | 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 (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 (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)
+ (Option.fold_left (fun acc { CAst.v = y } -> Name.fold_right Id.Set.add y acc) acc na)
inb) Id.Set.empty tms in
(* as, in & return vars *)
let forbidden_vars = Option.cata free_vars_of_constr_expr as_in_vars rtnpo in
@@ -1699,103 +1941,123 @@ 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 (CAst.make @@ 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 =
+ let is_patvar c = match DAst.get c with
+ | PatVar _ -> true
+ | _ -> false
+ in
let rec aux = function
| [] -> []
- | (_,PatVar _) :: q -> aux q
+ | (_, c) :: q when is_patvar c -> 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 =
- (Loc.ghost,[],thepats, (* "|p1,..,pn" *)
+ let sub_tms = List.map (fun id -> (DAst.make @@ GVar id),(Name id,None)) thevars (* "match v1,..,vn" *) in
+ let main_sub_eqn = CAst.make @@
+ ([],thepats, (* "|p1,..,pn" *)
Option.cata (intern_type env')
- (GHole(Loc.ghost,Evar_kinds.CasesType false,Misctypes.IntroAnonymous,None))
+ (DAst.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.ghost,[],List.make (List.length thepats) (PatVar(Loc.ghost,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))
+ [CAst.make @@ ([],List.make (List.length thepats) (DAst.make @@ PatVar Anonymous), (* "|_,..,_" *)
+ DAst.make @@ GHole(Evar_kinds.ImpossibleCase,Misctypes.IntroAnonymous,None))] (* "=> _" *) in
+ Some (DAst.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
- GCases (loc, sty, rtnpo, tms, List.flatten eqns')
- | CLetTuple (loc, nal, (na,po), b, c) ->
+ DAst.make ?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 *)
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
+ (CAst.make na') in
intern_type env'' u) po in
- GLetTuple (loc, List.map snd nal, (na', p'), b',
+ DAst.make ?loc @@
+ GLetTuple (List.map (fun { CAst.v } -> v) 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 ->
let env'' = push_name_env ntnvars (Variable,[],[],[]) (reset_hidden_inductive_implicit_test env)
- (Loc.ghost,na') in
+ (CAst.make na') in
intern_type env'' p) po in
- GIf (loc, c', (na', p'), intern env b1, intern env b2)
- | CHole (loc, k, naming, solve) ->
+ DAst.make ?loc @@
+ GIf (c', (na', p'), intern env b1, intern env b2)
+ | CHole (k, naming, solve) ->
let k = match k with
| None ->
let st = Evar_kinds.Define (not (Program.get_proofs_transparency ())) in
- Evar_kinds.QuestionMark st
+ (match naming with
+ | Misctypes.IntroIdentifier id -> Evar_kinds.NamedHole id
+ | _ -> Evar_kinds.QuestionMark (st,Anonymous))
| Some k -> k
in
let solve = match solve with
| None -> None
| Some gen ->
let (ltacvars, ntnvars) = lvar in
+ (* Preventively declare notation variables in ltac as non-bindings *)
+ Id.Map.iter (fun x (used_as_binder,_,_) -> used_as_binder := false) ntnvars;
let ntnvars = Id.Map.domain ntnvars in
+ let extra = ltacvars.ltac_extra in
let lvars = Id.Set.union ltacvars.ltac_bound ltacvars.ltac_vars in
let lvars = Id.Set.union lvars ntnvars in
- let lvars = Id.Set.union lvars env.ids in
+ let ltacvars = Id.Set.union lvars env.ids in
let ist = {
- Genintern.ltacvars = lvars;
- genv = globalenv;
+ Genintern.genv = globalenv;
+ ltacvars;
+ extra;
} in
let (_, glb) = Genintern.generic_intern ist gen in
Some glb
in
- GHole (loc, k, naming, solve)
+ DAst.make ?loc @@
+ GHole (k, naming, solve)
(* Parsing pattern variables *)
- | CPatVar (loc, n) when allow_patvar ->
- GPatVar (loc, (true,n))
- | CEvar (loc, n, []) when allow_patvar ->
- GPatVar (loc, (false,n))
+ | CPatVar n when pattern_mode ->
+ DAst.make ?loc @@
+ GPatVar (Evar_kinds.SecondOrderPatVar n)
+ | CEvar (n, []) when pattern_mode ->
+ DAst.make ?loc @@
+ GPatVar (Evar_kinds.FirstOrderPatVar n)
(* end *)
(* Parsing existential variables *)
- | CEvar (loc, n, l) ->
- GEvar (loc, n, List.map (on_snd (intern env)) l)
- | CPatVar (loc, _) ->
+ | CEvar (n, l) ->
+ DAst.make ?loc @@
+ GEvar (n, List.map (on_snd (intern env)) l)
+ | CPatVar _ ->
raise (InternalizationError (loc,IllegalMetavariable))
(* end *)
- | CSort (loc, s) ->
- GSort(loc,s)
- | CCast (loc, c1, c2) ->
- GCast (loc,intern env c1, Miscops.map_cast_type (intern_type env) c2)
-
+ | CSort s ->
+ DAst.make ?loc @@
+ GSort s
+ | CCast (c1, c2) ->
+ DAst.make ?loc @@
+ GCast (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 =
+ 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 *)
- and intern_multiple_pattern env n (loc,pl) =
- let idsl_pll = List.map (intern_cases_pattern globalenv (None,env.scopes) empty_alias) pl in
+ and intern_multiple_pattern env n pl =
+ let idsl_pll = List.map (intern_cases_pattern globalenv ntnvars (None,env.scopes) empty_alias) pl in
+ let loc = loc_of_multiple_pattern pl in
check_number_of_pattern loc n pl;
- product_of_cases_patterns [] idsl_pll
+ product_of_cases_patterns empty_alias idsl_pll
(* Expands a disjunction of multiple pattern *)
and intern_disjunctive_multiple_pattern env loc n mpl =
@@ -1807,29 +2069,32 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
(ids,List.flatten mpl')
(* Expands a pattern-matching clause [lhs => rhs] *)
- and intern_eqn n env (loc,lhs,rhs) =
+ and intern_eqn n env {loc;v=(lhs,rhs)} =
let eqn_ids,pll = intern_disjunctive_multiple_pattern env loc n lhs in
(* Linearity implies the order in ids is irrelevant *)
+ let eqn_ids = List.map (fun x -> x.v) eqn_ids in
check_linearity lhs eqn_ids;
let env_ids = List.fold_right Id.Set.add eqn_ids env.ids in
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
+ CAst.make ?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)
- | _, None -> None,(Loc.ghost,Anonymous)
- | _, Some (loc,na) -> None,(loc,na) in
+ let extra_id,na =
+ let loc = tm'.CAst.loc in
+ match DAst.get tm', na with
+ | GVar id, None when not (Id.Map.mem id (snd lvar)) -> Some id, CAst.make ?loc @@ Name id
+ | GRef (VarRef id, _), None -> Some id, CAst.make ?loc @@ Name id
+ | _, None -> None, CAst.make Anonymous
+ | _, Some ({ CAst.loc; v = na } as lna) -> None, lna in
(* the "in" part *)
let match_td,typ = match t with
| Some t ->
- let with_letin,(ind,l) = intern_ind_pattern globalenv (None,env.scopes) t in
+ let with_letin,(ind,l) = intern_ind_pattern globalenv ntnvars (None,env.scopes) t in
let (mib,mip) = Inductive.lookup_mind_specif globalenv ind in
let nparams = (List.length (mib.Declarations.mind_params_ctxt)) in
(* for "in Vect n", we answer (["n","n"],[(loc,"n")])
@@ -1841,44 +2106,40 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
let (match_to_do,nal) =
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
+ | { CAst.v = Anonymous } -> l
+ | { CAst.loc; v = (Name y as x) } -> (y, DAst.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.ghost,Anonymous)::var_acc)
+ canonize_args t l forbidden_names match_acc ((CAst.make Anonymous)::var_acc)
| [],[] ->
(add_name match_acc na, var_acc)
- | _::t,PatVar (loc,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 ->
- let fresh =
- Namegen.next_name_away_with_default_using_types "iV" cano_name forbidden_names ty in
- canonize_args t tt (fresh::forbidden_names)
- ((fresh,c)::match_acc) ((cases_pattern_loc c,Name fresh)::var_acc)
+ begin match DAst.get c with
+ | PatVar x ->
+ let loc = c.CAst.loc in
+ canonize_args t tt forbidden_names
+ (add_name match_acc CAst.(make ?loc x)) (CAst.make ?loc x::var_acc)
+ | _ ->
+ let fresh =
+ Namegen.next_name_away_with_default_using_types "iV" cano_name forbidden_names (EConstr.of_constr ty) in
+ canonize_args t tt (Id.Set.add fresh forbidden_names)
+ ((fresh,c)::match_acc) ((CAst.make ?loc:(cases_pattern_loc c) @@ Name fresh)::var_acc)
+ end
| _ -> assert false in
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)
+ canonize_args args_rel l forbidden_names_for_gen [] [] in
+ match_to_do, Some (CAst.make ?loc:(cases_pattern_expr_loc t) (ind,List.rev_map (fun x -> x.v) nal))
| None ->
[], None in
- (tm',(snd na,typ)), extra_id, match_td
-
- and iterate_prod loc2 env bk ty body nal =
- let env, bl = intern_assumption intern ntnvars env nal bk ty in
- it_mkGProd loc2 bl (intern_type env body)
-
- and iterate_lam loc2 env bk ty body nal =
- let env, bl = intern_assumption intern ntnvars env nal bk ty in
- it_mkGLambda loc2 bl (intern env body)
+ (tm',(na.CAst.v, typ)), extra_id, match_td
and intern_impargs c env l subscopes args =
let eargs, rargs = extract_explicit_arg l args in
if !parsing_explicit then
if Id.Map.is_empty eargs then intern_args env subscopes rargs
- else error "Arguments given by name or position not supported in explicit mode."
+ else user_err Pp.(str "Arguments given by name or position not supported in explicit mode.")
else
let rec aux n impl subscopes eargs rargs =
let (enva,subscopes') = apply_scope_env env subscopes in
@@ -1895,34 +2156,37 @@ 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
+ (DAst.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
| (imp::impl', a::rargs') ->
intern enva a :: aux (n+1) impl' subscopes' eargs rargs'
| (imp::impl', []) ->
if not (Id.Map.is_empty eargs) then
(let (id,(loc,_)) = Id.Map.choose eargs in
- user_err_loc (loc,"",str "Not enough non implicit \
+ user_err ?loc (str "Not enough non implicit \
arguments to accept the argument bound to " ++
- pr_id id ++ str"."));
+ Id.print id ++ str"."));
[]
| ([], rargs) ->
assert (Id.Map.is_empty eargs);
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
- | GApp (loc', g, args) -> GApp (Loc.merge loc' loc, g, args@l)
- | _ -> GApp (Loc.merge (loc_of_glob_constr f) loc, f, l)
-
+ | l ->
+ let loc' = f.CAst.loc in
+ match DAst.get f with
+ | GApp (g, args) -> DAst.make ?loc:(Loc.merge_opt loc' loc) @@ GApp (g, args@l)
+ | _ -> DAst.make ?loc:(Loc.merge_opt (loc_of_glob_constr f) loc) @@ GApp (f, l)
+
and intern_args env subscopes = function
| [] -> []
| a::args ->
@@ -1934,8 +2198,8 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
intern env c
with
InternalizationError (loc,e) ->
- user_err_loc (loc,"internalize",
- explain_internalization_error e)
+ user_err ?loc ~hdr:"internalize"
+ (explain_internalization_error e)
(**************************************************************************)
(* Functions to translate constr_expr into glob_constr *)
@@ -1946,35 +2210,34 @@ let extract_ids env =
(Termops.ids_of_rel_context (Environ.rel_context env))
Id.Set.empty
-let scope_of_type_kind = function
+let scope_of_type_kind sigma = function
| IsType -> Notation.current_type_scope_name ()
- | OfType typ -> compute_type_scope typ
+ | OfType typ -> compute_type_scope sigma typ
| WithoutTypeConstraint -> None
let empty_ltac_sign = {
ltac_vars = Id.Set.empty;
ltac_bound = Id.Set.empty;
+ ltac_extra = Genintern.Store.empty;
}
-let intern_gen kind env
- ?(impls=empty_internalization_env) ?(allow_patvar=false) ?(ltacvars=empty_ltac_sign)
+let intern_gen kind env sigma
+ ?(impls=empty_internalization_env) ?(pattern_mode=false) ?(ltacvars=empty_ltac_sign)
c =
- let tmp_scope = scope_of_type_kind kind in
+ let tmp_scope = scope_of_type_kind sigma kind in
internalize env {ids = extract_ids env; unb = false;
tmp_scope = tmp_scope; scopes = [];
impls = impls}
- allow_patvar (ltacvars, Id.Map.empty) c
-
-let intern_constr env c = intern_gen WithoutTypeConstraint env c
-
-let intern_type env c = intern_gen IsType env c
+ pattern_mode (ltacvars, Id.Map.empty) c
+let intern_constr env sigma c = intern_gen WithoutTypeConstraint env sigma c
+let intern_type env sigma c = intern_gen IsType env sigma c
let intern_pattern globalenv patt =
try
- intern_cases_pattern globalenv (None,[]) empty_alias patt
+ intern_cases_pattern globalenv Id.Map.empty (None,[]) empty_alias patt
with
InternalizationError (loc,e) ->
- user_err_loc (loc,"internalize",explain_internalization_error e)
+ user_err ?loc ~hdr:"internalize" (explain_internalization_error e)
(*********************************************************************)
@@ -1983,7 +2246,7 @@ let intern_pattern globalenv patt =
(* All evars resolved *)
let interp_gen kind env sigma ?(impls=empty_internalization_env) c =
- let c = intern_gen kind ~impls env c in
+ let c = intern_gen kind ~impls env sigma c in
understand ~expected_type:kind env sigma c
let interp_constr env sigma ?(impls=empty_internalization_env) c =
@@ -1998,51 +2261,52 @@ let interp_casted_constr env sigma ?(impls=empty_internalization_env) c typ =
(* Not all evars expected to be resolved *)
let interp_open_constr env sigma c =
- understand_tcc env sigma (intern_constr env c)
+ understand_tcc env sigma (intern_constr env sigma c)
(* Not all evars expected to be resolved and computation of implicit args *)
-let interp_constr_evars_gen_impls env evdref
+let interp_constr_evars_gen_impls env sigma
?(impls=empty_internalization_env) expected_type c =
- let c = intern_gen expected_type ~impls env c in
+ let c = intern_gen expected_type ~impls env sigma c in
let imps = Implicit_quantifiers.implicits_of_glob_constr ~with_products:(expected_type == IsType) c in
- understand_tcc_evars env evdref ~expected_type c, imps
+ let sigma, c = understand_tcc env sigma ~expected_type c in
+ sigma, (c, imps)
-let interp_constr_evars_impls env evdref ?(impls=empty_internalization_env) c =
- interp_constr_evars_gen_impls env evdref ~impls WithoutTypeConstraint c
+let interp_constr_evars_impls env sigma ?(impls=empty_internalization_env) c =
+ interp_constr_evars_gen_impls env sigma ~impls WithoutTypeConstraint c
let interp_casted_constr_evars_impls env evdref ?(impls=empty_internalization_env) c typ =
interp_constr_evars_gen_impls env evdref ~impls (OfType typ) c
-let interp_type_evars_impls env evdref ?(impls=empty_internalization_env) c =
- interp_constr_evars_gen_impls env evdref ~impls IsType c
+let interp_type_evars_impls env sigma ?(impls=empty_internalization_env) c =
+ interp_constr_evars_gen_impls env sigma ~impls IsType c
(* Not all evars expected to be resolved, with side-effect on evars *)
-let interp_constr_evars_gen env evdref ?(impls=empty_internalization_env) expected_type c =
- let c = intern_gen expected_type ~impls env c in
- understand_tcc_evars env evdref ~expected_type c
+let interp_constr_evars_gen env sigma ?(impls=empty_internalization_env) expected_type c =
+ let c = intern_gen expected_type ~impls env sigma c in
+ understand_tcc env sigma ~expected_type c
let interp_constr_evars env evdref ?(impls=empty_internalization_env) c =
interp_constr_evars_gen env evdref WithoutTypeConstraint ~impls c
-let interp_casted_constr_evars env evdref ?(impls=empty_internalization_env) c typ =
- interp_constr_evars_gen env evdref ~impls (OfType typ) c
+let interp_casted_constr_evars env sigma ?(impls=empty_internalization_env) c typ =
+ interp_constr_evars_gen env sigma ~impls (OfType typ) c
-let interp_type_evars env evdref ?(impls=empty_internalization_env) c =
- interp_constr_evars_gen env evdref IsType ~impls c
+let interp_type_evars env sigma ?(impls=empty_internalization_env) c =
+ interp_constr_evars_gen env sigma IsType ~impls c
(* Miscellaneous *)
-let intern_constr_pattern env ?(as_type=false) ?(ltacvars=empty_ltac_sign) c =
+let intern_constr_pattern env sigma ?(as_type=false) ?(ltacvars=empty_ltac_sign) c =
let c = intern_gen (if as_type then IsType else WithoutTypeConstraint)
- ~allow_patvar:true ~ltacvars env c in
+ ~pattern_mode:true ~ltacvars env sigma c in
pattern_of_glob_constr c
-let interp_notation_constr ?(impls=empty_internalization_env) nenv a =
- let env = Global.env () in
+let interp_notation_constr env ?(impls=empty_internalization_env) nenv a =
(* [vl] is intended to remember the scope of the free variables of [a] *)
- let vl = Id.Map.map (fun typ -> (ref true, ref None, typ)) nenv.ninterp_var_type in
+ let vl = Id.Map.map (fun typ -> (ref false, ref None, typ)) nenv.ninterp_var_type in
+ let impls = Id.Map.fold (fun id _ impls -> Id.Map.remove id impls) nenv.ninterp_var_type impls in
let c = internalize (Global.env()) {ids = extract_ids env; unb = false;
tmp_scope = None; scopes = []; impls = impls}
false (empty_ltac_sign, vl) a in
@@ -2051,24 +2315,23 @@ let interp_notation_constr ?(impls=empty_internalization_env) nenv a =
(* Splits variables into those that are binding, bound, or both *)
(* binding and bound *)
let out_scope = function None -> None,[] | Some (a,l) -> a,l in
- let vars = Id.Map.map (fun (isonlybinding, sc, typ) ->
- (!isonlybinding, out_scope !sc, typ)) vl in
+ let unused = match reversible with NonInjective ids -> ids | _ -> [] in
+ let vars = Id.Map.mapi (fun id (used_as_binder, sc, typ) ->
+ (!used_as_binder && not (List.mem_f Id.equal id unused), out_scope !sc)) vl in
(* Returns [a] and the ordered list of variables with their scopes *)
vars, a, reversible
(* Interpret binders and contexts *)
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 = intern_gen IsType env sigma 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
- understand_tcc_evars env evdref ~expected_type:IsType t'
-
-open Environ
+let interp_binder_evars env sigma na t =
+ let t = intern_gen IsType env sigma t in
+ let t' = locate_if_hole ?loc:(loc_of_glob_constr t) na t in
+ understand_tcc env sigma ~expected_type:IsType t'
let my_intern_constr env lvar acc c =
internalize env acc false lvar c
@@ -2078,30 +2341,24 @@ let intern_context global_level env impl_env binders =
let lvar = (empty_ltac_sign, Id.Map.empty) in
let lenv, bl = List.fold_left
(fun (lenv, bl) b ->
- let bl = List.map (fun a -> BDRawDef a) bl in
let (env, bl) = intern_local_binder_aux ~global_level (my_intern_constr env lvar) Id.Map.empty (lenv, bl) b in
- let bl =
- List.map
- (function
- | BDRawDef a -> a
- | BDPattern (loc,_,_,_,_) ->
- Loc.raise loc (Stream.Error "pattern with quote not allowed here")) bl in
(env, bl))
({ids = extract_ids env; unb = false;
tmp_scope = None; scopes = []; impls = impl_env}, []) binders in
- (lenv.impls, List.map snd bl)
+ (lenv.impls, List.map glob_local_binder_of_extended bl)
with InternalizationError (loc,e) ->
- user_err_loc (loc,"internalize", explain_internalization_error e)
+ user_err ?loc ~hdr:"internalize" (explain_internalization_error e)
-let interp_rawcontext_evars env evdref k bl =
- let (env, par, _, impls) =
+let interp_glob_context_evars env sigma k bl =
+ let open EConstr in
+ let env, sigma, par, _, impls =
List.fold_left
- (fun (env,params,n,impls) (na, k, b, t) ->
+ (fun (env,sigma,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
+ let sigma, t = understand_tcc env sigma ~expected_type:IsType t' in
match b with
None ->
let d = LocalAssum (na,t) in
@@ -2111,16 +2368,15 @@ let interp_rawcontext_evars env evdref k bl =
(ExplByPos (n, na), (true, true, true)) :: impls
else impls
in
- (push_rel d env, d::params, succ n, impls)
+ (push_rel d env, sigma, d::params, succ n, impls)
| Some b ->
- let c = understand_tcc_evars env evdref ~expected_type:(OfType t) b in
+ let sigma, c = understand_tcc env sigma ~expected_type:(OfType t) b in
let d = LocalDef (na, c, t) in
- (push_rel d env, d::params, n, impls))
- (env,[],k+1,[]) (List.rev bl)
- in (env, par), impls
+ (push_rel d env, sigma, d::params, n, impls))
+ (env,sigma,[],k+1,[]) (List.rev bl)
+ in sigma, ((env, par), impls)
-let interp_context_evars ?(global_level=false) ?(impl_env=empty_internalization_env) ?(shift=0) env evdref params =
+let interp_context_evars ?(global_level=false) ?(impl_env=empty_internalization_env) ?(shift=0) env sigma params =
let int_env,bl = intern_context global_level env impl_env params in
- let x = interp_rawcontext_evars env evdref shift bl in
- int_env, x
-
+ let sigma, x = interp_glob_context_evars env sigma shift bl in
+ sigma, (int_env, x)