aboutsummaryrefslogtreecommitdiff
path: root/interp
diff options
context:
space:
mode:
authorMaxime Dénès2019-04-05 12:18:49 +0200
committerMaxime Dénès2019-04-10 15:41:45 +0200
commit5af486406e366bf23558311a7367e573c617e078 (patch)
treee2996582ca8eab104141dd75b82ac1777e53cb72 /interp
parentf6dc8b19760aaf0cc14e6d9eb2d581ba1a05a762 (diff)
Remove calls to global env in Inductiveops
Diffstat (limited to 'interp')
-rw-r--r--interp/constrextern.ml12
-rw-r--r--interp/constrintern.ml31
-rw-r--r--interp/notation_ops.ml5
3 files changed, 25 insertions, 23 deletions
diff --git a/interp/constrextern.ml b/interp/constrextern.ml
index 24b1362e6d..b89b6b5765 100644
--- a/interp/constrextern.ml
+++ b/interp/constrextern.ml
@@ -212,7 +212,7 @@ let encode_record r =
module PrintingRecordRecord =
PrintingInductiveMake (struct
- let encode = encode_record
+ let encode _env = encode_record
let field = "Record"
let title = "Types leading to pretty-printing using record notation: "
let member_message s b =
@@ -224,7 +224,7 @@ module PrintingRecordRecord =
module PrintingRecordConstructor =
PrintingInductiveMake (struct
- let encode = encode_record
+ let encode _env = encode_record
let field = "Constructor"
let title = "Types leading to pretty-printing using constructor form: "
let member_message s b =
@@ -289,11 +289,11 @@ let extern_reference ?loc vars l = !my_extern_reference vars l
let add_patt_for_params ind l =
if !Flags.in_debugger then l else
- Util.List.addn (Inductiveops.inductive_nparamdecls ind) (CAst.make @@ CPatAtom None) l
+ Util.List.addn (Inductiveops.inductive_nparamdecls (Global.env()) ind) (CAst.make @@ CPatAtom None) l
let add_cpatt_for_params ind l =
if !Flags.in_debugger then l else
- Util.List.addn (Inductiveops.inductive_nparamdecls ind) (DAst.make @@ PatVar Anonymous) l
+ Util.List.addn (Inductiveops.inductive_nparamdecls (Global.env()) ind) (DAst.make @@ PatVar Anonymous) l
let drop_implicits_in_patt cst nb_expl args =
let impl_st = (implicits_of_global cst) in
@@ -364,7 +364,7 @@ let mkPat ?loc qid l = CAst.make ?loc @@
let pattern_printable_in_both_syntax (ind,_ as c) =
let impl_st = extract_impargs_data (implicits_of_global (ConstructRef c)) in
- let nb_params = Inductiveops.inductive_nparams ind in
+ let nb_params = Inductiveops.inductive_nparams (Global.env()) ind in
List.exists (fun (_,impls) ->
(List.length impls >= nb_params) &&
let params,args = Util.List.chop nb_params impls in
@@ -526,7 +526,7 @@ let rec extern_notation_ind_pattern allscopes lonely_seen vars ind args = functi
let extern_ind_pattern_in_scope (custom,scopes as allscopes) vars ind args =
(* pboutill: There are letins in pat which is incompatible with notations and
not explicit application. *)
- if !Flags.in_debugger||Inductiveops.inductive_has_local_defs ind then
+ if !Flags.in_debugger||Inductiveops.inductive_has_local_defs (Global.env()) ind then
let c = extern_reference vars (IndRef ind) in
let args = List.map (extern_cases_pattern_in_scope allscopes vars) args in
CAst.make @@ CPatCstr (c, Some (add_patt_for_params ind args), [])
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index d4555707c4..808dd225af 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -658,7 +658,7 @@ let terms_of_binders bl =
| PatCstr (c,l,_) ->
let qid = qualid_of_path ?loc (Nametab.path_of_global (ConstructRef c)) in
let hole = CAst.make ?loc @@ CHole (None,IntroAnonymous,None) in
- let params = List.make (Inductiveops.inductive_nparams (fst c)) hole in
+ let params = List.make (Inductiveops.inductive_nparams (Global.env()) (fst c)) hole in
CAppExpl ((None,qid,None),params @ List.map term_of_pat l)) pt in
let rec extract_variables l = match l with
| bnd :: l ->
@@ -1212,10 +1212,10 @@ let check_or_pat_variables loc ids idsl =
@return if letin are included *)
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) ||
+ if Int.equal n (Inductiveops.constructor_nallargs env cstr) then false else
+ (Int.equal n (Inductiveops.constructor_nalldecls env cstr) ||
(error_wrong_numarg_constructor ?loc env cstr
- (Inductiveops.constructor_nrealargs cstr)))
+ (Inductiveops.constructor_nrealargs env cstr)))
open Declarations
@@ -1241,9 +1241,9 @@ let add_local_defs_and_check_length loc env g pl args = match g with
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
+ let maxargs = Inductiveops.constructor_nalldecls env cstr in
if List.length pl' + List.length args > maxargs then
- error_wrong_numarg_constructor ?loc env cstr (Inductiveops.constructor_nrealargs cstr);
+ error_wrong_numarg_constructor ?loc env cstr (Inductiveops.constructor_nrealargs env 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
@@ -1273,15 +1273,15 @@ let add_implicits_check_length fail nargs nargs_with_letin impls_st len_pl1 pl2
in aux 0 (impl_list,pl2)
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 nargs = Inductiveops.constructor_nallargs env c in
+ let nargs' = Inductiveops.constructor_nalldecls env c in
let impls_st = implicits_of_global (ConstructRef c) in
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 nallargs = inductive_nallargs env c in
+ let nalldecls = inductive_nalldecls env c in
let impls_st = implicits_of_global (IndRef c) in
add_implicits_check_length (error_wrong_numarg_inductive ?loc env c)
nallargs nalldecls impls_st len_pl1 pl2
@@ -1289,8 +1289,8 @@ let add_implicits_check_ind_length env loc c len_pl1 pl2 =
(** Do not raise NotEnoughArguments thanks to preconditions*)
let chop_params_pattern loc ind args with_letin =
let nparams = if with_letin
- then Inductiveops.inductive_nparamdecls ind
- else Inductiveops.inductive_nparams ind in
+ then Inductiveops.inductive_nparamdecls (Global.env()) ind
+ else Inductiveops.inductive_nparams (Global.env()) ind in
assert (nparams <= List.length args);
let params,args = List.chop nparams args in
List.iter (fun c -> match DAst.get c with
@@ -1310,10 +1310,11 @@ let find_constructor loc add_params ref =
in
cstr, match add_params with
| Some nb_args ->
+ let env = Global.env () in
let nb =
- if Int.equal nb_args (Inductiveops.constructor_nrealdecls cstr)
- then Inductiveops.inductive_nparamdecls ind
- else Inductiveops.inductive_nparams ind
+ if Int.equal nb_args (Inductiveops.constructor_nrealdecls env cstr)
+ then Inductiveops.inductive_nparamdecls env ind
+ else Inductiveops.inductive_nparams env ind
in
List.make nb ([], [(Id.Map.empty, DAst.make @@ PatVar Anonymous)])
| None -> []
diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml
index 1f1cf4506f..9801e56ca1 100644
--- a/interp/notation_ops.ml
+++ b/interp/notation_ops.ml
@@ -782,7 +782,7 @@ let rec pat_binder_of_term t = DAst.map (function
| GApp (t, l) ->
begin match DAst.get t with
| GRef (ConstructRef cstr,_) ->
- let nparams = Inductiveops.inductive_nparams (fst cstr) in
+ let nparams = Inductiveops.inductive_nparams (Global.env()) (fst cstr) in
let _,l = List.chop nparams l in
PatCstr (cstr, List.map pat_binder_of_term l, Anonymous)
| _ -> raise No_match
@@ -909,7 +909,8 @@ let bind_term_as_binding_env alp (terms,termlists,binders,binderlists as sigma)
alp, add_env alp sigma var (DAst.make @@ GVar id)
let bind_binding_as_term_env alp (terms,termlists,binders,binderlists as sigma) var c =
- let pat = try cases_pattern_of_glob_constr Anonymous c with Not_found -> raise No_match in
+ let env = Global.env () in
+ let pat = try cases_pattern_of_glob_constr env Anonymous c with Not_found -> raise No_match in
try
(* If already bound to a binder, unify the term and the binder *)
let patl' = Id.List.assoc var binders in