aboutsummaryrefslogtreecommitdiffhomepage
path: root/interp
diff options
context:
space:
mode:
authorGravatar Emilio Jesus Gallego Arias <e+git@x80.org>2018-05-21 19:21:26 +0200
committerGravatar Emilio Jesus Gallego Arias <e+git@x80.org>2018-05-31 11:16:45 +0200
commit2a69be9e8243fa67d5c7ef5f10e623b02a0a3e2f (patch)
tree493d780d22515a60716845109d12690caf0b1f8a /interp
parentac8a84e3b4dc530b000e17b72c7e26f7a957420f (diff)
[api] Move `Constrexpr` to the `interp` module.
Continuing the interface cleanup we place `Constrexpr` in the internalization module, which is the one that eliminates it. This slims down `pretyping` considerably, including removing the `Univdecls` module which existed only due to bad dependency ordering in the first place. Thanks to @ Skyskimmer we also remove a duplicate `univ_decl` definition among `Misctypes` and `UState`. This is mostly a proof of concept yet as it depends on quite a few patches of the tree. For sure some tweaks will be necessary, but it should be good for review now. IMO the tree is now in a state where we can could easy eliminate more than 10 modules without any impact, IMHO this is a net saving API-wise and would help people to understand the structure of the code better.
Diffstat (limited to 'interp')
-rw-r--r--interp/constrexpr.ml152
-rw-r--r--interp/constrexpr_ops.ml32
-rw-r--r--interp/constrexpr_ops.mli7
-rw-r--r--interp/genredexpr.ml66
-rw-r--r--interp/implicit_quantifiers.ml6
-rw-r--r--interp/implicit_quantifiers.mli4
-rw-r--r--interp/interp.mllib3
-rw-r--r--interp/modintern.ml2
-rw-r--r--interp/redops.ml64
-rw-r--r--interp/redops.mli20
10 files changed, 353 insertions, 3 deletions
diff --git a/interp/constrexpr.ml b/interp/constrexpr.ml
new file mode 100644
index 000000000..ca6ea94f0
--- /dev/null
+++ b/interp/constrexpr.ml
@@ -0,0 +1,152 @@
+(************************************************************************)
+(* * 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 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+open Names
+open Libnames
+open Misctypes
+open Decl_kinds
+
+(** {6 Concrete syntax for terms } *)
+
+(** [constr_expr] is the abstract syntax tree produced by the parser *)
+type universe_decl_expr = (lident list, Glob_term.glob_constraint list) UState.gen_universe_decl
+
+type ident_decl = lident * universe_decl_expr option
+type name_decl = lname * universe_decl_expr option
+
+type notation = string
+
+type explicitation =
+ | ExplByPos of int * Id.t option (* a reference to the n-th product starting from left *)
+ | ExplByName of Id.t
+
+type binder_kind =
+ | Default of binding_kind
+ | Generalized of binding_kind * binding_kind * bool
+ (** Inner binding, outer bindings, typeclass-specific flag
+ for implicit generalization of superclasses *)
+
+type abstraction_kind = AbsLambda | AbsPi
+
+type proj_flag = int option (** [Some n] = proj of the n-th visible argument *)
+
+(** Representation of integer literals that appear in Coq scripts.
+ We now use raw strings of digits in base 10 (big-endian), and a separate
+ sign flag. Note that this representation is not unique, due to possible
+ multiple leading zeros, and -0 = +0 *)
+
+type sign = bool
+type raw_natural_number = string
+
+type prim_token =
+ | Numeral of raw_natural_number * sign
+ | String of string
+
+type instance_expr = Glob_term.glob_level list
+
+type cases_pattern_expr_r =
+ | CPatAlias of cases_pattern_expr * lname
+ | CPatCstr of reference
+ * cases_pattern_expr list option * cases_pattern_expr list
+ (** [CPatCstr (_, c, Some l1, l2)] represents [(@ c l1) l2] *)
+ | CPatAtom of reference option
+ | CPatOr of cases_pattern_expr list
+ | CPatNotation of notation * cases_pattern_notation_substitution
+ * cases_pattern_expr list (** CPatNotation (_, n, l1 ,l2) represents
+ (notation n applied with substitution l1)
+ applied to arguments l2 *)
+ | CPatPrim of prim_token
+ | CPatRecord of (reference * cases_pattern_expr) list
+ | CPatDelimiters of string * cases_pattern_expr
+ | CPatCast of cases_pattern_expr * constr_expr
+and cases_pattern_expr = cases_pattern_expr_r CAst.t
+
+and cases_pattern_notation_substitution =
+ cases_pattern_expr list * (** for constr subterms *)
+ cases_pattern_expr list list (** for recursive notations *)
+
+and constr_expr_r =
+ | CRef of reference * instance_expr option
+ | CFix of lident * fix_expr list
+ | CCoFix of lident * cofix_expr list
+ | CProdN of local_binder_expr list * constr_expr
+ | CLambdaN of local_binder_expr list * constr_expr
+ | CLetIn of lname * constr_expr * constr_expr option * constr_expr
+ | CAppExpl of (proj_flag * reference * instance_expr option) * constr_expr list
+ | CApp of (proj_flag * constr_expr) *
+ (constr_expr * explicitation CAst.t option) list
+ | CRecord of (reference * constr_expr) list
+
+ (* representation of the "let" and "match" constructs *)
+ | CCases of Constr.case_style (* determines whether this value represents "let" or "match" construct *)
+ * constr_expr option (* return-clause *)
+ * case_expr list
+ * branch_expr list (* branches *)
+
+ | CLetTuple of lname list * (lname option * constr_expr option) *
+ constr_expr * constr_expr
+ | CIf of constr_expr * (lname option * constr_expr option)
+ * constr_expr * constr_expr
+ | CHole of Evar_kinds.t option * intro_pattern_naming_expr * Genarg.raw_generic_argument option
+ | CPatVar of patvar
+ | CEvar of Glob_term.existential_name * (Id.t * constr_expr) list
+ | CSort of Glob_term.glob_sort
+ | CCast of constr_expr * constr_expr cast_type
+ | CNotation of notation * constr_notation_substitution
+ | CGeneralization of binding_kind * abstraction_kind option * constr_expr
+ | CPrim of prim_token
+ | CDelimiters of string * constr_expr
+ | CProj of reference * constr_expr
+and constr_expr = constr_expr_r CAst.t
+
+and case_expr = constr_expr (* expression that is being matched *)
+ * lname option (* as-clause *)
+ * cases_pattern_expr option (* in-clause *)
+
+and branch_expr =
+ (cases_pattern_expr list list * constr_expr) CAst.t
+
+and fix_expr =
+ lident * (lident option * recursion_order_expr) *
+ local_binder_expr list * constr_expr * constr_expr
+
+and cofix_expr =
+ lident * local_binder_expr list * constr_expr * constr_expr
+
+and recursion_order_expr =
+ | CStructRec
+ | CWfRec of constr_expr
+ | CMeasureRec of constr_expr * constr_expr option (** measure, relation *)
+
+(* Anonymous defs allowed ?? *)
+and local_binder_expr =
+ | CLocalAssum of lname list * binder_kind * constr_expr
+ | CLocalDef of lname * constr_expr * constr_expr option
+ | CLocalPattern of (cases_pattern_expr * constr_expr option) CAst.t
+
+and constr_notation_substitution =
+ constr_expr list * (** for constr subterms *)
+ constr_expr list list * (** for recursive notations *)
+ cases_pattern_expr list * (** for binders *)
+ local_binder_expr list list (** for binder lists (recursive notations) *)
+
+type constr_pattern_expr = constr_expr
+
+(** Concrete syntax for modules and module types *)
+
+type with_declaration_ast =
+ | CWith_Module of Id.t list CAst.t * qualid CAst.t
+ | CWith_Definition of Id.t list CAst.t * universe_decl_expr option * constr_expr
+
+type module_ast_r =
+ | CMident of qualid
+ | CMapply of module_ast * module_ast
+ | CMwith of module_ast * with_declaration_ast
+and module_ast = module_ast_r CAst.t
diff --git a/interp/constrexpr_ops.ml b/interp/constrexpr_ops.ml
index 3eb5acfc5..1be1dd96c 100644
--- a/interp/constrexpr_ops.ml
+++ b/interp/constrexpr_ops.ml
@@ -600,3 +600,35 @@ let _ = Goptions.declare_bool_option {
Goptions.optread = (fun () -> !asymmetric_patterns);
Goptions.optwrite = (fun a -> asymmetric_patterns:=a);
}
+
+(** Local universe and constraint declarations. *)
+
+let interp_univ_constraints env evd cstrs =
+ let interp (evd,cstrs) (u, d, u') =
+ let ul = Pretyping.interp_known_glob_level evd u in
+ let u'l = Pretyping.interp_known_glob_level evd u' in
+ let cstr = (ul,d,u'l) in
+ let cstrs' = Univ.Constraint.add cstr cstrs in
+ try let evd = Evd.add_constraints evd (Univ.Constraint.singleton cstr) in
+ evd, cstrs'
+ with Univ.UniverseInconsistency e ->
+ CErrors.user_err ~hdr:"interp_constraint"
+ (Univ.explain_universe_inconsistency (Termops.pr_evd_level evd) e)
+ in
+ List.fold_left interp (evd,Univ.Constraint.empty) cstrs
+
+let interp_univ_decl env decl =
+ let open UState in
+ let pl : lident list = decl.univdecl_instance in
+ let evd = Evd.from_ctx (UState.make_with_initial_binders (Environ.universes env) pl) in
+ let evd, cstrs = interp_univ_constraints env evd decl.univdecl_constraints in
+ let decl = { univdecl_instance = pl;
+ univdecl_extensible_instance = decl.univdecl_extensible_instance;
+ univdecl_constraints = cstrs;
+ univdecl_extensible_constraints = decl.univdecl_extensible_constraints }
+ in evd, decl
+
+let interp_univ_decl_opt env l =
+ match l with
+ | None -> Evd.from_env env, UState.default_univ_decl
+ | Some decl -> interp_univ_decl env decl
diff --git a/interp/constrexpr_ops.mli b/interp/constrexpr_ops.mli
index 2df9e5a8c..b4f0886ac 100644
--- a/interp/constrexpr_ops.mli
+++ b/interp/constrexpr_ops.mli
@@ -116,3 +116,10 @@ val error_invalid_pattern_notation : ?loc:Loc.t -> unit -> 'a
(** Placeholder for global option, should be moved to a parameter *)
val asymmetric_patterns : bool ref
+
+(** Local universe and constraint declarations. *)
+val interp_univ_decl : Environ.env -> universe_decl_expr ->
+ Evd.evar_map * UState.universe_decl
+
+val interp_univ_decl_opt : Environ.env -> universe_decl_expr option ->
+ Evd.evar_map * UState.universe_decl
diff --git a/interp/genredexpr.ml b/interp/genredexpr.ml
new file mode 100644
index 000000000..80697461a
--- /dev/null
+++ b/interp/genredexpr.ml
@@ -0,0 +1,66 @@
+(************************************************************************)
+(* * 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 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+(** Reduction expressions *)
+
+(** The parsing produces initially a list of [red_atom] *)
+
+type 'a red_atom =
+ | FBeta
+ | FMatch
+ | FFix
+ | FCofix
+ | FZeta
+ | FConst of 'a list
+ | FDeltaBut of 'a list
+
+(** This list of atoms is immediately converted to a [glob_red_flag] *)
+
+type 'a glob_red_flag = {
+ rBeta : bool;
+ rMatch : bool;
+ rFix : bool;
+ rCofix : bool;
+ rZeta : bool;
+ rDelta : bool; (** true = delta all but rConst; false = delta only on rConst*)
+ rConst : 'a list
+}
+
+(** Generic kinds of reductions *)
+
+type ('a,'b,'c) red_expr_gen =
+ | Red of bool
+ | Hnf
+ | Simpl of 'b glob_red_flag*('b,'c) Util.union Locus.with_occurrences option
+ | Cbv of 'b glob_red_flag
+ | Cbn of 'b glob_red_flag
+ | Lazy of 'b glob_red_flag
+ | Unfold of 'b Locus.with_occurrences list
+ | Fold of 'a list
+ | Pattern of 'a Locus.with_occurrences list
+ | ExtraRedExpr of string
+ | CbvVm of ('b,'c) Util.union Locus.with_occurrences option
+ | CbvNative of ('b,'c) Util.union Locus.with_occurrences option
+
+type ('a,'b,'c) may_eval =
+ | ConstrTerm of 'a
+ | ConstrEval of ('a,'b,'c) red_expr_gen * 'a
+ | ConstrContext of Misctypes.lident * 'a
+ | ConstrTypeOf of 'a
+
+open Libnames
+open Constrexpr
+open Misctypes
+
+type r_trm = constr_expr
+type r_pat = constr_pattern_expr
+type r_cst = reference or_by_notation
+
+type raw_red_expr = (r_trm, r_cst, r_pat) red_expr_gen
diff --git a/interp/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml
index 289890544..b48db9ac5 100644
--- a/interp/implicit_quantifiers.ml
+++ b/interp/implicit_quantifiers.ml
@@ -17,12 +17,14 @@ open Glob_term
open Constrexpr
open Libnames
open Typeclasses
-open Typeclasses_errors
open Pp
open Libobject
open Nameops
open Context.Rel.Declaration
+exception MismatchedContextInstance of Environ.env * Typeclasses_errors.contexts * constr_expr list * Context.Rel.t (* found, expected *)
+let mismatched_ctx_inst_err env c n m = raise (MismatchedContextInstance (env, c, n, m))
+
module RelDecl = Context.Rel.Declaration
(*i*)
@@ -238,7 +240,7 @@ let implicit_application env ?(allow_partial=true) f ty =
let applen = List.fold_left (fun acc (x, y) -> opt_succ y acc) 0 par in
let needlen = List.fold_left (fun acc x -> opt_succ x acc) 0 ci in
if not (Int.equal needlen applen) then
- Typeclasses_errors.mismatched_ctx_inst (Global.env ()) Parameters (List.map fst par) rd
+ mismatched_ctx_inst_err (Global.env ()) Typeclasses_errors.Parameters (List.map fst par) rd
end;
let pars = List.rev (List.combine ci rd) in
let args, avoid = combine_params avoid f par pars in
diff --git a/interp/implicit_quantifiers.mli b/interp/implicit_quantifiers.mli
index 39d0174f9..e64c5c542 100644
--- a/interp/implicit_quantifiers.mli
+++ b/interp/implicit_quantifiers.mli
@@ -45,3 +45,7 @@ val implicit_application : Id.Set.t -> ?allow_partial:bool ->
(Id.Set.t -> GlobRef.t option * Context.Rel.Declaration.t ->
Constrexpr.constr_expr * Id.Set.t) ->
constr_expr -> constr_expr * Id.Set.t
+
+(* Should be likely located elsewhere *)
+exception MismatchedContextInstance of Environ.env * Typeclasses_errors.contexts * constr_expr list * Context.Rel.t (* found, expected *)
+val mismatched_ctx_inst_err : Environ.env -> Typeclasses_errors.contexts -> constr_expr list -> Context.Rel.t -> 'a
diff --git a/interp/interp.mllib b/interp/interp.mllib
index 37a327a7d..8ed2b9fad 100644
--- a/interp/interp.mllib
+++ b/interp/interp.mllib
@@ -1,3 +1,6 @@
+Constrexpr
+Genredexpr
+Redops
Tactypes
Stdarg
Genintern
diff --git a/interp/modintern.ml b/interp/modintern.ml
index dc93d8dc4..fefd2ab6f 100644
--- a/interp/modintern.ml
+++ b/interp/modintern.ml
@@ -63,7 +63,7 @@ let transl_with_decl env = function
| CWith_Module ({CAst.v=fqid},qid) ->
WithMod (fqid,lookup_module qid), Univ.ContextSet.empty
| CWith_Definition ({CAst.v=fqid},udecl,c) ->
- let sigma, udecl = Univdecls.interp_univ_decl_opt env udecl in
+ let sigma, udecl = Constrexpr_ops.interp_univ_decl_opt env udecl in
let c, ectx = interp_constr env sigma c in
begin match UState.check_univ_decl ~poly:(Flags.is_universe_polymorphism()) ectx udecl with
| Entries.Polymorphic_const_entry ctx ->
diff --git a/interp/redops.ml b/interp/redops.ml
new file mode 100644
index 000000000..b9a74136e
--- /dev/null
+++ b/interp/redops.ml
@@ -0,0 +1,64 @@
+(************************************************************************)
+(* * 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 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+open Genredexpr
+
+let union_consts l1 l2 = Util.List.union Pervasives.(=) l1 l2 (* FIXME *)
+
+let make_red_flag l =
+ let rec add_flag red = function
+ | [] -> red
+ | FBeta :: lf -> add_flag { red with rBeta = true } lf
+ | FMatch :: lf -> add_flag { red with rMatch = true } lf
+ | FFix :: lf -> add_flag { red with rFix = true } lf
+ | FCofix :: lf -> add_flag { red with rCofix = true } lf
+ | FZeta :: lf -> add_flag { red with rZeta = true } lf
+ | FConst l :: lf ->
+ if red.rDelta then
+ CErrors.user_err Pp.(str
+ "Cannot set both constants to unfold and constants not to unfold");
+ add_flag { red with rConst = union_consts red.rConst l } lf
+ | FDeltaBut l :: lf ->
+ if red.rConst <> [] && not red.rDelta then
+ CErrors.user_err Pp.(str
+ "Cannot set both constants to unfold and constants not to unfold");
+ add_flag
+ { red with rConst = union_consts red.rConst l; rDelta = true }
+ lf
+ in
+ add_flag
+ {rBeta = false; rMatch = false; rFix = false; rCofix = false;
+ rZeta = false; rDelta = false; rConst = []}
+ l
+
+
+let all_flags =
+ {rBeta = true; rMatch = true; rFix = true; rCofix = true;
+ rZeta = true; rDelta = true; rConst = []}
+
+(** Mapping [red_expr_gen] *)
+
+let map_flags f flags =
+ { flags with rConst = List.map f flags.rConst }
+
+let map_occs f (occ,e) = (occ,f e)
+
+let map_red_expr_gen f g h = function
+ | Fold l -> Fold (List.map f l)
+ | Pattern occs_l -> Pattern (List.map (map_occs f) occs_l)
+ | Simpl (flags,occs_o) ->
+ Simpl (map_flags g flags, Option.map (map_occs (Util.map_union g h)) occs_o)
+ | Unfold occs_l -> Unfold (List.map (map_occs g) occs_l)
+ | Cbv flags -> Cbv (map_flags g flags)
+ | Lazy flags -> Lazy (map_flags g flags)
+ | CbvVm occs_o -> CbvVm (Option.map (map_occs (Util.map_union g h)) occs_o)
+ | CbvNative occs_o -> CbvNative (Option.map (map_occs (Util.map_union g h)) occs_o)
+ | Cbn flags -> Cbn (map_flags g flags)
+ | ExtraRedExpr _ | Red _ | Hnf as x -> x
diff --git a/interp/redops.mli b/interp/redops.mli
new file mode 100644
index 000000000..7254f29b2
--- /dev/null
+++ b/interp/redops.mli
@@ -0,0 +1,20 @@
+(************************************************************************)
+(* * 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 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+open Genredexpr
+
+val make_red_flag : 'a red_atom list -> 'a glob_red_flag
+
+val all_flags : 'a glob_red_flag
+
+(** Mapping [red_expr_gen] *)
+
+val map_red_expr_gen : ('a -> 'd) -> ('b -> 'e) -> ('c -> 'f) ->
+ ('a,'b,'c) red_expr_gen -> ('d,'e,'f) red_expr_gen