aboutsummaryrefslogtreecommitdiffhomepage
path: root/vernac
diff options
context:
space:
mode:
authorGravatar Emilio Jesus Gallego Arias <e+git@x80.org>2016-12-18 18:14:58 +0100
committerGravatar Emilio Jesus Gallego Arias <e+git@x80.org>2017-02-15 20:45:28 +0100
commitfa9df2efe5666789bf91bb7761567cd53e6c451f (patch)
treedfabeded9b4060114e0f9d7f4d3760efc982ae0c /vernac
parent0df095ec0715f548180bbff70a6feb673c6726a6 (diff)
[stm] Break stm/toplevel dependency loop.
Currently, the STM, vernac interpretation, and the toplevel are intertwined in a mutual dependency that needs to be resolved using imperative callbacks. This is problematic for a few reasons, in particular it makes the interpretation of commands that affect the document quite intricate. As a first step, we split the `toplevel/` directory into two: "pure" vernac interpretation is moved to the `vernac/` directory, on which the STM relies. Test suite passes, and only one command seems to be disabled with this approach, "Show Script" which is to my understanding obsolete. Subsequent commits will fix this and refine some of the invariants that are not needed anymore.
Diffstat (limited to 'vernac')
-rw-r--r--vernac/assumptions.ml320
-rw-r--r--vernac/assumptions.mli31
-rw-r--r--vernac/auto_ind_decl.ml969
-rw-r--r--vernac/auto_ind_decl.mli42
-rw-r--r--vernac/class.ml328
-rw-r--r--vernac/class.mli48
-rw-r--r--vernac/classes.ml410
-rw-r--r--vernac/classes.mli66
-rw-r--r--vernac/command.ml1357
-rw-r--r--vernac/command.mli176
-rw-r--r--vernac/discharge.ml120
-rw-r--r--vernac/discharge.mli14
-rw-r--r--vernac/doc.tex10
-rw-r--r--vernac/explainErr.ml129
-rw-r--r--vernac/explainErr.mli21
-rw-r--r--vernac/himsg.ml1268
-rw-r--r--vernac/himsg.mli42
-rw-r--r--vernac/ind_tables.ml203
-rw-r--r--vernac/ind_tables.mli51
-rw-r--r--vernac/indschemes.ml530
-rw-r--r--vernac/indschemes.mli49
-rw-r--r--vernac/lemmas.ml557
-rw-r--r--vernac/lemmas.mli69
-rw-r--r--vernac/locality.ml107
-rw-r--r--vernac/locality.mli51
-rw-r--r--vernac/metasyntax.ml1469
-rw-r--r--vernac/metasyntax.mli63
-rw-r--r--vernac/mltop.ml447
-rw-r--r--vernac/mltop.mli88
-rw-r--r--vernac/obligations.ml1179
-rw-r--r--vernac/obligations.mli113
-rw-r--r--vernac/record.ml583
-rw-r--r--vernac/record.mli46
-rw-r--r--vernac/search.ml381
-rw-r--r--vernac/search.mli84
-rw-r--r--vernac/vernac.mllib17
-rw-r--r--vernac/vernacentries.ml2271
-rw-r--r--vernac/vernacentries.mli66
-rw-r--r--vernac/vernacinterp.ml77
-rw-r--r--vernac/vernacinterp.mli20
40 files changed, 13872 insertions, 0 deletions
diff --git a/vernac/assumptions.ml b/vernac/assumptions.ml
new file mode 100644
index 000000000..8865cd646
--- /dev/null
+++ b/vernac/assumptions.ml
@@ -0,0 +1,320 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* The following definitions are used by the function
+ [assumptions] which gives as an output the set of all
+ axioms and sections variables on which a given term depends
+ in a context (expectingly the Global context) *)
+
+(* Initial author: Arnaud Spiwack
+ Module-traversing code: Pierre Letouzey *)
+
+open Pp
+open CErrors
+open Util
+open Names
+open Term
+open Declarations
+open Mod_subst
+open Globnames
+open Printer
+open Context.Named.Declaration
+
+module NamedDecl = Context.Named.Declaration
+
+(** For a constant c in a module sealed by an interface (M:T and
+ not M<:T), [Global.lookup_constant] may return a [constant_body]
+ without body. We fix this by looking in the implementation
+ of the module *)
+
+let modcache = ref (MPmap.empty : structure_body MPmap.t)
+
+let rec search_mod_label lab = function
+ | [] -> raise Not_found
+ | (l, SFBmodule mb) :: _ when Label.equal l lab -> mb
+ | _ :: fields -> search_mod_label lab fields
+
+let rec search_cst_label lab = function
+ | [] -> raise Not_found
+ | (l, SFBconst cb) :: _ when Label.equal l lab -> cb
+ | _ :: fields -> search_cst_label lab fields
+
+(* TODO: using [empty_delta_resolver] below is probably slightly incorrect. But:
+ a) I don't see currently what should be used instead
+ b) this shouldn't be critical for Print Assumption. At worse some
+ constants will have a canonical name which is non-canonical,
+ leading to failures in [Global.lookup_constant], but our own
+ [lookup_constant] should work.
+*)
+
+let rec fields_of_functor f subs mp0 args = function
+ |NoFunctor a -> f subs mp0 args a
+ |MoreFunctor (mbid,_,e) ->
+ match args with
+ | [] -> assert false (* we should only encounter applied functors *)
+ | mpa :: args ->
+ let subs = join (map_mbid mbid mpa empty_delta_resolver (*TODO*)) subs in
+ fields_of_functor f subs mp0 args e
+
+let rec lookup_module_in_impl mp =
+ try Global.lookup_module mp
+ with Not_found ->
+ (* The module we search might not be exported by its englobing module(s).
+ We access the upper layer, and then do a manual search *)
+ match mp with
+ | MPfile _ -> raise Not_found (* can happen if mp is an open module *)
+ | MPbound _ -> assert false
+ | MPdot (mp',lab') ->
+ let fields = memoize_fields_of_mp mp' in
+ search_mod_label lab' fields
+
+and memoize_fields_of_mp mp =
+ try MPmap.find mp !modcache
+ with Not_found ->
+ let l = fields_of_mp mp in
+ modcache := MPmap.add mp l !modcache;
+ l
+
+and fields_of_mp mp =
+ let mb = lookup_module_in_impl mp in
+ let fields,inner_mp,subs = fields_of_mb empty_subst mb [] in
+ let subs =
+ if mp_eq inner_mp mp then subs
+ else add_mp inner_mp mp mb.mod_delta subs
+ in
+ Modops.subst_structure subs fields
+
+and fields_of_mb subs mb args = match mb.mod_expr with
+ |Algebraic expr -> fields_of_expression subs mb.mod_mp args expr
+ |Struct sign -> fields_of_signature subs mb.mod_mp args sign
+ |Abstract|FullStruct -> fields_of_signature subs mb.mod_mp args mb.mod_type
+
+(** The Abstract case above corresponds to [Declare Module] *)
+
+and fields_of_signature x =
+ fields_of_functor
+ (fun subs mp0 args struc ->
+ assert (List.is_empty args);
+ (struc, mp0, subs)) x
+
+and fields_of_expr subs mp0 args = function
+ |MEident mp ->
+ let mb = lookup_module_in_impl (subst_mp subs mp) in
+ fields_of_mb subs mb args
+ |MEapply (me1,mp2) -> fields_of_expr subs mp0 (mp2::args) me1
+ |MEwith _ -> assert false (* no 'with' in [mod_expr] *)
+
+and fields_of_expression x = fields_of_functor fields_of_expr x
+
+let lookup_constant_in_impl cst fallback =
+ try
+ let mp,dp,lab = repr_kn (canonical_con cst) in
+ let fields = memoize_fields_of_mp mp in
+ (* A module found this way is necessarily closed, in particular
+ our constant cannot be in an opened section : *)
+ search_cst_label lab fields
+ with Not_found ->
+ (* Either:
+ - The module part of the constant isn't registered yet :
+ we're still in it, so the [constant_body] found earlier
+ (if any) was a true axiom.
+ - The label has not been found in the structure. This is an error *)
+ match fallback with
+ | Some cb -> cb
+ | None -> anomaly (str "Print Assumption: unknown constant " ++ pr_con cst)
+
+let lookup_constant cst =
+ try
+ let cb = Global.lookup_constant cst in
+ if Declareops.constant_has_body cb then cb
+ else lookup_constant_in_impl cst (Some cb)
+ with Not_found -> lookup_constant_in_impl cst None
+
+(** Graph traversal of an object, collecting on the way the dependencies of
+ traversed objects *)
+
+let label_of = function
+ | ConstRef kn -> pi3 (repr_con kn)
+ | IndRef (kn,_)
+ | ConstructRef ((kn,_),_) -> pi3 (repr_mind kn)
+ | VarRef id -> Label.of_id id
+
+let rec traverse current ctx accu t = match kind_of_term t with
+| Var id ->
+ let body () = id |> Global.lookup_named |> NamedDecl.get_value in
+ traverse_object accu body (VarRef id)
+| Const (kn, _) ->
+ let body () = Global.body_of_constant_body (lookup_constant kn) in
+ traverse_object accu body (ConstRef kn)
+| Ind ((mind, _) as ind, _) ->
+ traverse_inductive accu mind (IndRef ind)
+| Construct (((mind, _), _) as cst, _) ->
+ traverse_inductive accu mind (ConstructRef cst)
+| Meta _ | Evar _ -> assert false
+| Case (_,oty,c,[||]) ->
+ (* non dependent match on an inductive with no constructors *)
+ begin match Constr.(kind oty, kind c) with
+ | Lambda(_,_,oty), Const (kn, _)
+ when Vars.noccurn 1 oty &&
+ not (Declareops.constant_has_body (lookup_constant kn)) ->
+ let body () = Global.body_of_constant_body (lookup_constant kn) in
+ traverse_object
+ ~inhabits:(current,ctx,Vars.subst1 mkProp oty) accu body (ConstRef kn)
+ | _ ->
+ Termops.fold_constr_with_full_binders
+ Context.Rel.add (traverse current) ctx accu t
+ end
+| _ -> Termops.fold_constr_with_full_binders
+ Context.Rel.add (traverse current) ctx accu t
+
+and traverse_object ?inhabits (curr, data, ax2ty) body obj =
+ let data, ax2ty =
+ let already_in = Refmap_env.mem obj data in
+ match body () with
+ | None ->
+ let data =
+ if not already_in then Refmap_env.add obj Refset_env.empty data else data in
+ let ax2ty =
+ if Option.is_empty inhabits then ax2ty else
+ let ty = Option.get inhabits in
+ try let l = Refmap_env.find obj ax2ty in Refmap_env.add obj (ty::l) ax2ty
+ with Not_found -> Refmap_env.add obj [ty] ax2ty in
+ data, ax2ty
+ | Some body ->
+ if already_in then data, ax2ty else
+ let contents,data,ax2ty =
+ traverse (label_of obj) Context.Rel.empty
+ (Refset_env.empty,data,ax2ty) body in
+ Refmap_env.add obj contents data, ax2ty
+ in
+ (Refset_env.add obj curr, data, ax2ty)
+
+(** Collects the references occurring in the declaration of mutual inductive
+ definitions. All the constructors and names of a mutual inductive
+ definition share exactly the same dependencies. Also, there is no explicit
+ dependency between mutually defined inductives and constructors. *)
+and traverse_inductive (curr, data, ax2ty) mind obj =
+ let firstind_ref = (IndRef (mind, 0)) in
+ let label = label_of obj in
+ let data, ax2ty =
+ (* Invariant : I_0 \in data iff I_i \in data iff c_ij \in data
+ where I_0, I_1, ... are in the same mutual definition and c_ij
+ are all their constructors. *)
+ if Refmap_env.mem firstind_ref data then data, ax2ty else
+ let mib = Global.lookup_mind mind in
+ (* Collects references of parameters *)
+ let param_ctx = mib.mind_params_ctxt in
+ let nparam = List.length param_ctx in
+ let accu =
+ traverse_context label Context.Rel.empty
+ (Refset_env.empty, data, ax2ty) param_ctx
+ in
+ (* Build the context of all arities *)
+ let arities_ctx =
+ let global_env = Global.env () in
+ Array.fold_left (fun accu oib ->
+ let pspecif = Univ.in_punivs (mib, oib) in
+ let ind_type = Inductive.type_of_inductive global_env pspecif in
+ let ind_name = Name oib.mind_typename in
+ Context.Rel.add (Context.Rel.Declaration.LocalAssum (ind_name, ind_type)) accu)
+ Context.Rel.empty mib.mind_packets
+ in
+ (* For each inductive, collects references in their arity and in the type
+ of constructors*)
+ let (contents, data, ax2ty) = Array.fold_left (fun accu oib ->
+ let arity_wo_param =
+ List.rev (List.skipn nparam (List.rev oib.mind_arity_ctxt))
+ in
+ let accu =
+ traverse_context
+ label param_ctx accu arity_wo_param
+ in
+ Array.fold_left (fun accu cst_typ ->
+ let param_ctx, cst_typ_wo_param = Term.decompose_prod_n_assum nparam cst_typ in
+ let ctx = Context.(Rel.fold_outside Context.Rel.add ~init:arities_ctx param_ctx) in
+ traverse label ctx accu cst_typ_wo_param)
+ accu oib.mind_user_lc)
+ accu mib.mind_packets
+ in
+ (* Maps all these dependencies to inductives and constructors*)
+ let data = Array.fold_left_i (fun n data oib ->
+ let ind = (mind, n) in
+ let data = Refmap_env.add (IndRef ind) contents data in
+ Array.fold_left_i (fun k data _ ->
+ Refmap_env.add (ConstructRef (ind, k+1)) contents data
+ ) data oib.mind_consnames) data mib.mind_packets
+ in
+ data, ax2ty
+ in
+ (Refset_env.add obj curr, data, ax2ty)
+
+(** Collects references in a rel_context. *)
+and traverse_context current ctx accu ctxt =
+ snd (Context.Rel.fold_outside (fun decl (ctx, accu) ->
+ match decl with
+ | Context.Rel.Declaration.LocalDef (_,c,t) ->
+ let accu = traverse current ctx (traverse current ctx accu t) c in
+ let ctx = Context.Rel.add decl ctx in
+ ctx, accu
+ | Context.Rel.Declaration.LocalAssum (_,t) ->
+ let accu = traverse current ctx accu t in
+ let ctx = Context.Rel.add decl ctx in
+ ctx, accu) ctxt ~init:(ctx, accu))
+
+let traverse current t =
+ let () = modcache := MPmap.empty in
+ traverse current Context.Rel.empty (Refset_env.empty, Refmap_env.empty, Refmap_env.empty) t
+
+(** Hopefully bullet-proof function to recover the type of a constant. It just
+ ignores all the universe stuff. There are many issues that can arise when
+ considering terms out of any valid environment, so use with caution. *)
+let type_of_constant cb = match cb.Declarations.const_type with
+| Declarations.RegularArity ty -> ty
+| Declarations.TemplateArity (ctx, arity) ->
+ Term.mkArity (ctx, Sorts.sort_of_univ arity.Declarations.template_level)
+
+let assumptions ?(add_opaque=false) ?(add_transparent=false) st gr t =
+ let (idts, knst) = st in
+ (** Only keep the transitive dependencies *)
+ let (_, graph, ax2ty) = traverse (label_of gr) t in
+ let fold obj _ accu = match obj with
+ | VarRef id ->
+ let decl = Global.lookup_named id in
+ if is_local_assum decl then
+ let t = Context.Named.Declaration.get_type decl in
+ ContextObjectMap.add (Variable id) t accu
+ else accu
+ | ConstRef kn ->
+ let cb = lookup_constant kn in
+ let accu =
+ if cb.const_typing_flags.check_guarded then accu
+ else
+ let l = try Refmap_env.find obj ax2ty with Not_found -> [] in
+ ContextObjectMap.add (Axiom (Guarded kn, l)) Constr.mkProp accu
+ in
+ if not (Declareops.constant_has_body cb) || not cb.const_typing_flags.check_universes then
+ let t = type_of_constant cb in
+ let l = try Refmap_env.find obj ax2ty with Not_found -> [] in
+ ContextObjectMap.add (Axiom (Constant kn,l)) t accu
+ else if add_opaque && (Declareops.is_opaque cb || not (Cpred.mem kn knst)) then
+ let t = type_of_constant cb in
+ ContextObjectMap.add (Opaque kn) t accu
+ else if add_transparent then
+ let t = type_of_constant cb in
+ ContextObjectMap.add (Transparent kn) t accu
+ else
+ accu
+ | IndRef (m,_) | ConstructRef ((m,_),_) ->
+ let mind = Global.lookup_mind m in
+ if mind.mind_typing_flags.check_guarded then
+ accu
+ else
+ let l = try Refmap_env.find obj ax2ty with Not_found -> [] in
+ ContextObjectMap.add (Axiom (Positive m, l)) Constr.mkProp accu
+ in
+ Refmap_env.fold fold graph ContextObjectMap.empty
diff --git a/vernac/assumptions.mli b/vernac/assumptions.mli
new file mode 100644
index 000000000..072675783
--- /dev/null
+++ b/vernac/assumptions.mli
@@ -0,0 +1,31 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open Term
+open Globnames
+open Printer
+
+(** Collects all the objects on which a term directly relies, bypassing kernel
+ opacity, together with the recursive dependence DAG of objects.
+
+ WARNING: some terms may not make sense in the environment, because they are
+ sealed inside opaque modules. Do not try to do anything fancy with those
+ terms apart from printing them, otherwise demons may fly out of your nose.
+*)
+val traverse :
+ Label.t -> constr ->
+ (Refset_env.t * Refset_env.t Refmap_env.t *
+ (label * Context.Rel.t * types) list Refmap_env.t)
+
+(** Collects all the assumptions (optionally including opaque definitions)
+ on which a term relies (together with their type). The above warning of
+ {!traverse} also applies. *)
+val assumptions :
+ ?add_opaque:bool -> ?add_transparent:bool -> transparent_state ->
+ global_reference -> constr -> Term.types ContextObjectMap.t
diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml
new file mode 100644
index 000000000..594f2e944
--- /dev/null
+++ b/vernac/auto_ind_decl.ml
@@ -0,0 +1,969 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* This file is about the automatic generation of schemes about
+ decidable equality, created by Vincent Siles, Oct 2007 *)
+
+open Tacmach
+open CErrors
+open Util
+open Pp
+open Term
+open Vars
+open Termops
+open Declarations
+open Names
+open Globnames
+open Nameops
+open Inductiveops
+open Tactics
+open Ind_tables
+open Misctypes
+open Proofview.Notations
+
+module RelDecl = Context.Rel.Declaration
+
+let out_punivs = Univ.out_punivs
+
+(**********************************************************************)
+(* Generic synthesis of boolean equality *)
+
+let quick_chop n l =
+ let rec kick_last = function
+ | t::[] -> []
+ | t::q -> t::(kick_last q)
+ | [] -> failwith "kick_last"
+and aux = function
+ | (0,l') -> l'
+ | (n,h::t) -> aux (n-1,t)
+ | _ -> failwith "quick_chop"
+ in
+ if n > (List.length l) then failwith "quick_chop args"
+ else kick_last (aux (n,l) )
+
+let deconstruct_type t =
+ let l,r = decompose_prod t in
+ (List.rev_map snd l)@[r]
+
+exception EqNotFound of inductive * inductive
+exception EqUnknown of string
+exception UndefinedCst of string
+exception InductiveWithProduct
+exception InductiveWithSort
+exception ParameterWithoutEquality of global_reference
+exception NonSingletonProp of inductive
+exception DecidabilityMutualNotSupported
+exception NoDecidabilityCoInductive
+
+let dl = Loc.ghost
+
+let constr_of_global g = lazy (Universes.constr_of_global g)
+
+(* Some pre declaration of constant we are going to use *)
+let bb = constr_of_global Coqlib.glob_bool
+
+let andb_prop = fun _ -> (Coqlib.build_bool_type()).Coqlib.andb_prop
+
+let andb_true_intro = fun _ ->
+ (Coqlib.build_bool_type()).Coqlib.andb_true_intro
+
+let tt = constr_of_global Coqlib.glob_true
+
+let ff = constr_of_global Coqlib.glob_false
+
+let eq = constr_of_global Coqlib.glob_eq
+
+let sumbool = Coqlib.build_coq_sumbool
+
+let andb = fun _ -> (Coqlib.build_bool_type()).Coqlib.andb
+
+let induct_on c = induction false None c None None
+
+let destruct_on c = destruct false None c None None
+
+let destruct_on_using c id =
+ destruct false None c
+ (Some (dl,IntroOrPattern [[dl,IntroNaming IntroAnonymous];
+ [dl,IntroNaming (IntroIdentifier id)]]))
+ None
+
+let destruct_on_as c l =
+ destruct false None c (Some (dl,l)) None
+
+(* reconstruct the inductive with the correct deBruijn indexes *)
+let mkFullInd (ind,u) n =
+ let mib = Global.lookup_mind (fst ind) in
+ let nparams = mib.mind_nparams in
+ let nparrec = mib.mind_nparams_rec in
+ (* params context divided *)
+ let lnonparrec,lnamesparrec =
+ context_chop (nparams-nparrec) mib.mind_params_ctxt in
+ if nparrec > 0
+ then mkApp (mkIndU (ind,u),
+ Array.of_list(Context.Rel.to_extended_list (nparrec+n) lnamesparrec))
+ else mkIndU (ind,u)
+
+let check_bool_is_defined () =
+ try let _ = Global.type_of_global_unsafe Coqlib.glob_bool in ()
+ with e when CErrors.noncritical e -> raise (UndefinedCst "bool")
+
+let beq_scheme_kind_aux = ref (fun _ -> failwith "Undefined")
+
+let build_beq_scheme mode kn =
+ check_bool_is_defined ();
+ (* fetching global env *)
+ let env = Global.env() in
+ (* fetching the mutual inductive body *)
+ let mib = Global.lookup_mind kn in
+ (* number of inductives in the mutual *)
+ let nb_ind = Array.length mib.mind_packets in
+ (* number of params in the type *)
+ let nparams = mib.mind_nparams in
+ let nparrec = mib.mind_nparams_rec in
+ (* params context divided *)
+ let lnonparrec,lnamesparrec =
+ context_chop (nparams-nparrec) mib.mind_params_ctxt in
+ (* predef coq's boolean type *)
+ (* rec name *)
+ let rec_name i =(Id.to_string (Array.get mib.mind_packets i).mind_typename)^
+ "_eqrec"
+ in
+ (* construct the "fun A B ... N, eqA eqB eqC ... N => fixpoint" part *)
+ let create_input c =
+ let myArrow u v = mkArrow u (lift 1 v)
+ and eqName = function
+ | Name s -> Id.of_string ("eq_"^(Id.to_string s))
+ | Anonymous -> Id.of_string "eq_A"
+ in
+ let ext_rel_list = Context.Rel.to_extended_list 0 lnamesparrec in
+ let lift_cnt = ref 0 in
+ let eqs_typ = List.map (fun aa ->
+ let a = lift !lift_cnt aa in
+ incr lift_cnt;
+ myArrow a (myArrow a (Lazy.force bb))
+ ) ext_rel_list in
+
+ let eq_input = List.fold_left2
+ ( fun a b decl -> (* mkLambda(n,b,a) ) *)
+ (* here I leave the Naming thingy so that the type of
+ the function is more readable for the user *)
+ mkNamedLambda (eqName (RelDecl.get_name decl)) b a )
+ c (List.rev eqs_typ) lnamesparrec
+ in
+ List.fold_left (fun a decl ->(* mkLambda(n,t,a)) eq_input rel_list *)
+ (* Same here , hoping the auto renaming will do something good ;) *)
+ mkNamedLambda
+ (match RelDecl.get_name decl with Name s -> s | Anonymous -> Id.of_string "A")
+ (RelDecl.get_type decl) a) eq_input lnamesparrec
+ in
+ let make_one_eq cur =
+ let u = Univ.Instance.empty in
+ let ind = (kn,cur),u (* FIXME *) in
+ (* current inductive we are working on *)
+ let cur_packet = mib.mind_packets.(snd (fst ind)) in
+ (* Inductive toto : [rettyp] := *)
+ let rettyp = Inductive.type_of_inductive env ((mib,cur_packet),u) in
+ (* split rettyp in a list without the non rec params and the last ->
+ e.g. Inductive vec (A:Set) : nat -> Set := ... will do [nat] *)
+ let rettyp_l = quick_chop nparrec (deconstruct_type rettyp) in
+ (* give a type A, this function tries to find the equality on A declared
+ previously *)
+ (* nlist = the number of args (A , B , ... )
+ eqA = the deBruijn index of the first eq param
+ ndx = how much to translate due to the 2nd Case
+ *)
+ let compute_A_equality rel_list nlist eqA ndx t =
+ let lifti = ndx in
+ let rec aux c =
+ let (c,a) = Reductionops.whd_betaiota_stack Evd.empty c in
+ match kind_of_term c with
+ | Rel x -> mkRel (x-nlist+ndx), Safe_typing.empty_private_constants
+ | Var x ->
+ let eid = id_of_string ("eq_"^(string_of_id x)) in
+ let () =
+ try ignore (Environ.lookup_named eid env)
+ with Not_found -> raise (ParameterWithoutEquality (VarRef x))
+ in
+ mkVar eid, Safe_typing.empty_private_constants
+ | Cast (x,_,_) -> aux (applist (x,a))
+ | App _ -> assert false
+ | Ind ((kn',i as ind'),u) (*FIXME: universes *) ->
+ if eq_mind kn kn' then mkRel(eqA-nlist-i+nb_ind-1), Safe_typing.empty_private_constants
+ else begin
+ try
+ let eq, eff =
+ let c, eff = find_scheme ~mode (!beq_scheme_kind_aux()) (kn',i) in
+ mkConst c, eff in
+ let eqa, eff =
+ let eqa, effs = List.split (List.map aux a) in
+ Array.of_list eqa,
+ List.fold_left Safe_typing.concat_private eff (List.rev effs)
+ in
+ let args =
+ Array.append
+ (Array.of_list (List.map (fun x -> lift lifti x) a)) eqa in
+ if Int.equal (Array.length args) 0 then eq, eff
+ else mkApp (eq, args), eff
+ with Not_found -> raise(EqNotFound (ind', fst ind))
+ end
+ | Sort _ -> raise InductiveWithSort
+ | Prod _ -> raise InductiveWithProduct
+ | Lambda _-> raise (EqUnknown "abstraction")
+ | LetIn _ -> raise (EqUnknown "let-in")
+ | Const kn ->
+ (match Environ.constant_opt_value_in env kn with
+ | None -> raise (ParameterWithoutEquality (ConstRef (fst kn)))
+ | Some c -> aux (applist (c,a)))
+ | Proj _ -> raise (EqUnknown "projection")
+ | Construct _ -> raise (EqUnknown "constructor")
+ | Case _ -> raise (EqUnknown "match")
+ | CoFix _ -> raise (EqUnknown "cofix")
+ | Fix _ -> raise (EqUnknown "fix")
+ | Meta _ -> raise (EqUnknown "meta-variable")
+ | Evar _ -> raise (EqUnknown "existential variable")
+ in
+ aux t
+ in
+ (* construct the predicate for the Case part*)
+ let do_predicate rel_list n =
+ List.fold_left (fun a b -> mkLambda(Anonymous,b,a))
+ (mkLambda (Anonymous,
+ mkFullInd ind (n+3+(List.length rettyp_l)+nb_ind-1),
+ (Lazy.force bb)))
+ (List.rev rettyp_l) in
+ (* make_one_eq *)
+ (* do the [| C1 ... => match Y with ... end
+ ...
+ Cn => match Y with ... end |] part *)
+ let ci = make_case_info env (fst ind) MatchStyle in
+ let constrs n = get_constructors env (make_ind_family (ind,
+ Context.Rel.to_extended_list (n+nb_ind-1) mib.mind_params_ctxt)) in
+ let constrsi = constrs (3+nparrec) in
+ let n = Array.length constrsi in
+ let ar = Array.make n (Lazy.force ff) in
+ let eff = ref Safe_typing.empty_private_constants in
+ for i=0 to n-1 do
+ let nb_cstr_args = List.length constrsi.(i).cs_args in
+ let ar2 = Array.make n (Lazy.force ff) in
+ let constrsj = constrs (3+nparrec+nb_cstr_args) in
+ for j=0 to n-1 do
+ if Int.equal i j then
+ ar2.(j) <- let cc = (match nb_cstr_args with
+ | 0 -> Lazy.force tt
+ | _ -> let eqs = Array.make nb_cstr_args (Lazy.force tt) in
+ for ndx = 0 to nb_cstr_args-1 do
+ let cc = RelDecl.get_type (List.nth constrsi.(i).cs_args ndx) in
+ let eqA, eff' = compute_A_equality rel_list
+ nparrec
+ (nparrec+3+2*nb_cstr_args)
+ (nb_cstr_args+ndx+1)
+ cc
+ in
+ eff := Safe_typing.concat_private eff' !eff;
+ Array.set eqs ndx
+ (mkApp (eqA,
+ [|mkRel (ndx+1+nb_cstr_args);mkRel (ndx+1)|]
+ ))
+ done;
+ Array.fold_left
+ (fun a b -> mkApp (andb(),[|b;a|]))
+ (eqs.(0))
+ (Array.sub eqs 1 (nb_cstr_args - 1))
+ )
+ in
+ (List.fold_left (fun a decl -> mkLambda (RelDecl.get_name decl, RelDecl.get_type decl, a)) cc
+ (constrsj.(j).cs_args)
+ )
+ else ar2.(j) <- (List.fold_left (fun a decl ->
+ mkLambda (RelDecl.get_name decl, RelDecl.get_type decl, a)) (Lazy.force ff) (constrsj.(j).cs_args) )
+ done;
+
+ ar.(i) <- (List.fold_left (fun a decl -> mkLambda (RelDecl.get_name decl, RelDecl.get_type decl, a))
+ (mkCase (ci,do_predicate rel_list nb_cstr_args,
+ mkVar (Id.of_string "Y") ,ar2))
+ (constrsi.(i).cs_args))
+ done;
+ mkNamedLambda (Id.of_string "X") (mkFullInd ind (nb_ind-1+1)) (
+ mkNamedLambda (Id.of_string "Y") (mkFullInd ind (nb_ind-1+2)) (
+ mkCase (ci, do_predicate rel_list 0,mkVar (Id.of_string "X"),ar))),
+ !eff
+ in (* build_beq_scheme *)
+ let names = Array.make nb_ind Anonymous and
+ types = Array.make nb_ind mkSet and
+ cores = Array.make nb_ind mkSet in
+ let eff = ref Safe_typing.empty_private_constants in
+ let u = Univ.Instance.empty in
+ for i=0 to (nb_ind-1) do
+ names.(i) <- Name (Id.of_string (rec_name i));
+ types.(i) <- mkArrow (mkFullInd ((kn,i),u) 0)
+ (mkArrow (mkFullInd ((kn,i),u) 1) (Lazy.force bb));
+ let c, eff' = make_one_eq i in
+ cores.(i) <- c;
+ eff := Safe_typing.concat_private eff' !eff
+ done;
+ (Array.init nb_ind (fun i ->
+ let kelim = Inductive.elim_sorts (mib,mib.mind_packets.(i)) in
+ if not (Sorts.List.mem InSet kelim) then
+ raise (NonSingletonProp (kn,i));
+ if mib.mind_finite = Decl_kinds.CoFinite then
+ raise NoDecidabilityCoInductive;
+ let fix = mkFix (((Array.make nb_ind 0),i),(names,types,cores)) in
+ create_input fix),
+ Evd.make_evar_universe_context (Global.env ()) None),
+ !eff
+
+let beq_scheme_kind = declare_mutual_scheme_object "_beq" build_beq_scheme
+
+let _ = beq_scheme_kind_aux := fun () -> beq_scheme_kind
+
+(* This function tryies to get the [inductive] between a constr
+ the constr should be Ind i or App(Ind i,[|args|])
+*)
+let destruct_ind c =
+ try let u,v = destApp c in
+ let indc = destInd u in
+ indc,v
+ with DestKO -> let indc = destInd c in
+ indc,[||]
+
+(*
+ In the following, avoid is the list of names to avoid.
+ If the args of the Inductive type are A1 ... An
+ then avoid should be
+ [| lb_An ... lb _A1 (resp. bl_An ... bl_A1)
+ eq_An .... eq_A1 An ... A1 |]
+so from Ai we can find the the correct eq_Ai bl_ai or lb_ai
+*)
+(* used in the leib -> bool side*)
+let do_replace_lb mode lb_scheme_key aavoid narg p q =
+ let avoid = Array.of_list aavoid in
+ let do_arg v offset =
+ try
+ let x = narg*offset in
+ let s = destVar v in
+ let n = Array.length avoid in
+ let rec find i =
+ if Id.equal avoid.(n-i) s then avoid.(n-i-x)
+ else (if i<n then find (i+1)
+ else user_err ~hdr:"AutoIndDecl.do_replace_lb"
+ (str "Var " ++ pr_id s ++ str " seems unknown.")
+ )
+ in mkVar (find 1)
+ with e when CErrors.noncritical e ->
+ (* if this happen then the args have to be already declared as a
+ Parameter*)
+ (
+ let mp,dir,lbl = repr_con (fst (destConst v)) in
+ mkConst (make_con mp dir (mk_label (
+ if Int.equal offset 1 then ("eq_"^(Label.to_string lbl))
+ else ((Label.to_string lbl)^"_lb")
+ )))
+ )
+ in
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let type_of_pq = Tacmach.New.of_old (fun gl -> pf_unsafe_type_of gl p) gl in
+ let u,v = destruct_ind type_of_pq
+ in let lb_type_of_p =
+ try
+ let c, eff = find_scheme ~mode lb_scheme_key (out_punivs u) (*FIXME*) in
+ Proofview.tclUNIT (mkConst c, eff)
+ with Not_found ->
+ (* spiwack: the format of this error message should probably
+ be improved. *)
+ let err_msg =
+ (str "Leibniz->boolean:" ++
+ str "You have to declare the" ++
+ str "decidability over " ++
+ Printer.pr_constr type_of_pq ++
+ str " first.")
+ in
+ Tacticals.New.tclZEROMSG err_msg
+ in
+ lb_type_of_p >>= fun (lb_type_of_p,eff) ->
+ let lb_args = Array.append (Array.append
+ (Array.map (fun x -> x) v)
+ (Array.map (fun x -> do_arg x 1) v))
+ (Array.map (fun x -> do_arg x 2) v)
+ in let app = if Array.equal eq_constr lb_args [||]
+ then lb_type_of_p else mkApp (lb_type_of_p,lb_args)
+ in
+ Tacticals.New.tclTHENLIST [
+ Proofview.tclEFFECTS eff;
+ Equality.replace p q ; apply app ; Auto.default_auto]
+ end }
+
+(* used in the bool -> leib side *)
+let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt =
+ let avoid = Array.of_list aavoid in
+ let do_arg v offset =
+ try
+ let x = narg*offset in
+ let s = destVar v in
+ let n = Array.length avoid in
+ let rec find i =
+ if Id.equal avoid.(n-i) s then avoid.(n-i-x)
+ else (if i<n then find (i+1)
+ else user_err ~hdr:"AutoIndDecl.do_replace_bl"
+ (str "Var " ++ pr_id s ++ str " seems unknown.")
+ )
+ in mkVar (find 1)
+ with e when CErrors.noncritical e ->
+ (* if this happen then the args have to be already declared as a
+ Parameter*)
+ (
+ let mp,dir,lbl = repr_con (fst (destConst v)) in
+ mkConst (make_con mp dir (mk_label (
+ if Int.equal offset 1 then ("eq_"^(Label.to_string lbl))
+ else ((Label.to_string lbl)^"_bl")
+ )))
+ )
+ in
+
+ let rec aux l1 l2 =
+ match (l1,l2) with
+ | (t1::q1,t2::q2) ->
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let tt1 = Tacmach.New.pf_unsafe_type_of gl t1 in
+ if eq_constr t1 t2 then aux q1 q2
+ else (
+ let u,v = try destruct_ind tt1
+ (* trick so that the good sequence is returned*)
+ with e when CErrors.noncritical e -> indu,[||]
+ in if eq_ind (fst u) ind
+ then Tacticals.New.tclTHENLIST [Equality.replace t1 t2; Auto.default_auto ; aux q1 q2 ]
+ else (
+ let bl_t1, eff =
+ try
+ let c, eff = find_scheme bl_scheme_key (out_punivs u) (*FIXME*) in
+ mkConst c, eff
+ with Not_found ->
+ (* spiwack: the format of this error message should probably
+ be improved. *)
+ let err_msg = string_of_ppcmds
+ (str "boolean->Leibniz:" ++
+ str "You have to declare the" ++
+ str "decidability over " ++
+ Printer.pr_constr tt1 ++
+ str " first.")
+ in
+ error err_msg
+ in let bl_args =
+ Array.append (Array.append
+ (Array.map (fun x -> x) v)
+ (Array.map (fun x -> do_arg x 1) v))
+ (Array.map (fun x -> do_arg x 2) v )
+ in
+ let app = if Array.equal eq_constr bl_args [||]
+ then bl_t1 else mkApp (bl_t1,bl_args)
+ in
+ Tacticals.New.tclTHENLIST [
+ Proofview.tclEFFECTS eff;
+ Equality.replace_by t1 t2
+ (Tacticals.New.tclTHEN (apply app) (Auto.default_auto)) ;
+ aux q1 q2 ]
+ )
+ )
+ end }
+ | ([],[]) -> Proofview.tclUNIT ()
+ | _ -> Tacticals.New.tclZEROMSG (str "Both side of the equality must have the same arity.")
+ in
+ begin try Proofview.tclUNIT (destApp lft)
+ with DestKO -> Tacticals.New.tclZEROMSG (str "replace failed.")
+ end >>= fun (ind1,ca1) ->
+ begin try Proofview.tclUNIT (destApp rgt)
+ with DestKO -> Tacticals.New.tclZEROMSG (str "replace failed.")
+ end >>= fun (ind2,ca2) ->
+ begin try Proofview.tclUNIT (out_punivs (destInd ind1))
+ with DestKO ->
+ begin try Proofview.tclUNIT (fst (fst (destConstruct ind1)))
+ with DestKO -> Tacticals.New.tclZEROMSG (str "The expected type is an inductive one.")
+ end
+ end >>= fun (sp1,i1) ->
+ begin try Proofview.tclUNIT (out_punivs (destInd ind2))
+ with DestKO ->
+ begin try Proofview.tclUNIT (fst (fst (destConstruct ind2)))
+ with DestKO -> Tacticals.New.tclZEROMSG (str "The expected type is an inductive one.")
+ end
+ end >>= fun (sp2,i2) ->
+ if not (eq_mind sp1 sp2) || not (Int.equal i1 i2)
+ then Tacticals.New.tclZEROMSG (str "Eq should be on the same type")
+ else aux (Array.to_list ca1) (Array.to_list ca2)
+
+(*
+ create, from a list of ids [i1,i2,...,in] the list
+ [(in,eq_in,in_bl,in_al),,...,(i1,eq_i1,i1_bl_i1_al )]
+*)
+let list_id l = List.fold_left ( fun a decl -> let s' =
+ match RelDecl.get_name decl with
+ Name s -> Id.to_string s
+ | Anonymous -> "A" in
+ (Id.of_string s',Id.of_string ("eq_"^s'),
+ Id.of_string (s'^"_bl"),
+ Id.of_string (s'^"_lb"))
+ ::a
+ ) [] l
+(*
+ build the right eq_I A B.. N eq_A .. eq_N
+*)
+let eqI ind l =
+ let list_id = list_id l in
+ let eA = Array.of_list((List.map (fun (s,_,_,_) -> mkVar s) list_id)@
+ (List.map (fun (_,seq,_,_)-> mkVar seq) list_id ))
+ and e, eff =
+ try let c, eff = find_scheme beq_scheme_kind ind in mkConst c, eff
+ with Not_found -> user_err ~hdr:"AutoIndDecl.eqI"
+ (str "The boolean equality on " ++ pr_mind (fst ind) ++ str " is needed.");
+ in (if Array.equal eq_constr eA [||] then e else mkApp(e,eA)), eff
+
+(**********************************************************************)
+(* Boolean->Leibniz *)
+
+open Namegen
+
+let compute_bl_goal ind lnamesparrec nparrec =
+ let eqI, eff = eqI ind lnamesparrec in
+ let list_id = list_id lnamesparrec in
+ let avoid = List.fold_right (Nameops.name_fold (fun id l -> id::l)) (List.map RelDecl.get_name lnamesparrec) [] in
+ let create_input c =
+ let x = next_ident_away (Id.of_string "x") avoid and
+ y = next_ident_away (Id.of_string "y") avoid in
+ let bl_typ = List.map (fun (s,seq,_,_) ->
+ mkNamedProd x (mkVar s) (
+ mkNamedProd y (mkVar s) (
+ mkArrow
+ ( mkApp(Lazy.force eq,[|(Lazy.force bb);mkApp(mkVar seq,[|mkVar x;mkVar y|]);(Lazy.force tt)|]))
+ ( mkApp(Lazy.force eq,[|mkVar s;mkVar x;mkVar y|]))
+ ))
+ ) list_id in
+ let bl_input = List.fold_left2 ( fun a (s,_,sbl,_) b ->
+ mkNamedProd sbl b a
+ ) c (List.rev list_id) (List.rev bl_typ) in
+ let eqs_typ = List.map (fun (s,_,_,_) ->
+ mkProd(Anonymous,mkVar s,mkProd(Anonymous,mkVar s,(Lazy.force bb)))
+ ) list_id in
+ let eq_input = List.fold_left2 ( fun a (s,seq,_,_) b ->
+ mkNamedProd seq b a
+ ) bl_input (List.rev list_id) (List.rev eqs_typ) in
+ List.fold_left (fun a decl -> mkNamedProd
+ (match RelDecl.get_name decl with Name s -> s | Anonymous -> next_ident_away (Id.of_string "A") avoid)
+ (RelDecl.get_type decl) a) eq_input lnamesparrec
+ in
+ let n = next_ident_away (Id.of_string "x") avoid and
+ m = next_ident_away (Id.of_string "y") avoid in
+ let u = Univ.Instance.empty in
+ create_input (
+ mkNamedProd n (mkFullInd (ind,u) nparrec) (
+ mkNamedProd m (mkFullInd (ind,u) (nparrec+1)) (
+ mkArrow
+ (mkApp(Lazy.force eq,[|(Lazy.force bb);mkApp(eqI,[|mkVar n;mkVar m|]);(Lazy.force tt)|]))
+ (mkApp(Lazy.force eq,[|mkFullInd (ind,u) (nparrec+3);mkVar n;mkVar m|]))
+ ))), eff
+
+let compute_bl_tact mode bl_scheme_key ind lnamesparrec nparrec =
+ let list_id = list_id lnamesparrec in
+ let avoid = ref [] in
+ let first_intros =
+ ( List.map (fun (s,_,_,_) -> s ) list_id ) @
+ ( List.map (fun (_,seq,_,_ ) -> seq) list_id ) @
+ ( List.map (fun (_,_,sbl,_ ) -> sbl) list_id )
+ in
+ let fresh_id s gl =
+ Tacmach.New.of_old begin fun gsig ->
+ let fresh = fresh_id (!avoid) s gsig in
+ avoid := fresh::(!avoid); fresh
+ end gl
+ in
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let fresh_first_intros = List.map (fun id -> fresh_id id gl) first_intros in
+ let freshn = fresh_id (Id.of_string "x") gl in
+ let freshm = fresh_id (Id.of_string "y") gl in
+ let freshz = fresh_id (Id.of_string "Z") gl in
+ (* try with *)
+ Tacticals.New.tclTHENLIST [ intros_using fresh_first_intros;
+ intro_using freshn ;
+ induct_on (mkVar freshn);
+ intro_using freshm;
+ destruct_on (mkVar freshm);
+ intro_using freshz;
+ intros;
+ Tacticals.New.tclTRY (
+ Tacticals.New.tclORELSE reflexivity (Equality.discr_tac false None)
+ );
+ simpl_in_hyp (freshz,Locus.InHyp);
+(*
+repeat ( apply andb_prop in z;let z1:= fresh "Z" in destruct z as [z1 z]).
+*)
+ Tacticals.New.tclREPEAT (
+ Tacticals.New.tclTHENLIST [
+ Simple.apply_in freshz (andb_prop());
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let fresht = fresh_id (Id.of_string "Z") gl in
+ destruct_on_as (mkVar freshz)
+ (IntroOrPattern [[dl,IntroNaming (IntroIdentifier fresht);
+ dl,IntroNaming (IntroIdentifier freshz)]])
+ end }
+ ]);
+(*
+ Ci a1 ... an = Ci b1 ... bn
+ replace bi with ai; auto || replace bi with ai by apply typeofbi_prod ; auto
+*)
+ Proofview.Goal.nf_enter { enter = begin fun gls ->
+ let gl = Proofview.Goal.concl gls in
+ match (kind_of_term gl) with
+ | App (c,ca) -> (
+ match (kind_of_term c) with
+ | Ind (indeq, u) ->
+ if eq_gr (IndRef indeq) Coqlib.glob_eq
+ then
+ Tacticals.New.tclTHEN
+ (do_replace_bl mode bl_scheme_key ind
+ (!avoid)
+ nparrec (ca.(2))
+ (ca.(1)))
+ Auto.default_auto
+ else
+ Tacticals.New.tclZEROMSG (str "Failure while solving Boolean->Leibniz.")
+ | _ -> Tacticals.New.tclZEROMSG (str" Failure while solving Boolean->Leibniz.")
+ )
+ | _ -> Tacticals.New.tclZEROMSG (str "Failure while solving Boolean->Leibniz.")
+ end }
+
+ ]
+ end }
+
+let bl_scheme_kind_aux = ref (fun _ -> failwith "Undefined")
+
+let side_effect_of_mode = function
+ | Declare.UserAutomaticRequest -> false
+ | Declare.InternalTacticRequest -> true
+ | Declare.UserIndividualRequest -> false
+
+let make_bl_scheme mode mind =
+ let mib = Global.lookup_mind mind in
+ if not (Int.equal (Array.length mib.mind_packets) 1) then
+ user_err
+ (str "Automatic building of boolean->Leibniz lemmas not supported");
+ let ind = (mind,0) in
+ let nparams = mib.mind_nparams in
+ let nparrec = mib.mind_nparams_rec in
+ let lnonparrec,lnamesparrec = (* TODO subst *)
+ context_chop (nparams-nparrec) mib.mind_params_ctxt in
+ let bl_goal, eff = compute_bl_goal ind lnamesparrec nparrec in
+ let ctx = Evd.make_evar_universe_context (Global.env ()) None in
+ let side_eff = side_effect_of_mode mode in
+ let (ans, _, ctx) = Pfedit.build_by_tactic ~side_eff (Global.env()) ctx bl_goal
+ (compute_bl_tact mode (!bl_scheme_kind_aux()) (ind, Univ.Instance.empty) lnamesparrec nparrec)
+ in
+ ([|ans|], ctx), eff
+
+let bl_scheme_kind = declare_mutual_scheme_object "_dec_bl" make_bl_scheme
+
+let _ = bl_scheme_kind_aux := fun () -> bl_scheme_kind
+
+(**********************************************************************)
+(* Leibniz->Boolean *)
+
+let compute_lb_goal ind lnamesparrec nparrec =
+ let list_id = list_id lnamesparrec in
+ let eq = Lazy.force eq and tt = Lazy.force tt and bb = Lazy.force bb in
+ let avoid = List.fold_right (Nameops.name_fold (fun id l -> id::l)) (List.map RelDecl.get_name lnamesparrec) [] in
+ let eqI, eff = eqI ind lnamesparrec in
+ let create_input c =
+ let x = next_ident_away (Id.of_string "x") avoid and
+ y = next_ident_away (Id.of_string "y") avoid in
+ let lb_typ = List.map (fun (s,seq,_,_) ->
+ mkNamedProd x (mkVar s) (
+ mkNamedProd y (mkVar s) (
+ mkArrow
+ ( mkApp(eq,[|mkVar s;mkVar x;mkVar y|]))
+ ( mkApp(eq,[|bb;mkApp(mkVar seq,[|mkVar x;mkVar y|]);tt|]))
+ ))
+ ) list_id in
+ let lb_input = List.fold_left2 ( fun a (s,_,_,slb) b ->
+ mkNamedProd slb b a
+ ) c (List.rev list_id) (List.rev lb_typ) in
+ let eqs_typ = List.map (fun (s,_,_,_) ->
+ mkProd(Anonymous,mkVar s,mkProd(Anonymous,mkVar s,bb))
+ ) list_id in
+ let eq_input = List.fold_left2 ( fun a (s,seq,_,_) b ->
+ mkNamedProd seq b a
+ ) lb_input (List.rev list_id) (List.rev eqs_typ) in
+ List.fold_left (fun a decl -> mkNamedProd
+ (match (RelDecl.get_name decl) with Name s -> s | Anonymous -> Id.of_string "A")
+ (RelDecl.get_type decl) a) eq_input lnamesparrec
+ in
+ let n = next_ident_away (Id.of_string "x") avoid and
+ m = next_ident_away (Id.of_string "y") avoid in
+ let u = Univ.Instance.empty in
+ create_input (
+ mkNamedProd n (mkFullInd (ind,u) nparrec) (
+ mkNamedProd m (mkFullInd (ind,u) (nparrec+1)) (
+ mkArrow
+ (mkApp(eq,[|mkFullInd (ind,u) (nparrec+2);mkVar n;mkVar m|]))
+ (mkApp(eq,[|bb;mkApp(eqI,[|mkVar n;mkVar m|]);tt|]))
+ ))), eff
+
+let compute_lb_tact mode lb_scheme_key ind lnamesparrec nparrec =
+ let list_id = list_id lnamesparrec in
+ let avoid = ref [] in
+ let first_intros =
+ ( List.map (fun (s,_,_,_) -> s ) list_id ) @
+ ( List.map (fun (_,seq,_,_) -> seq) list_id ) @
+ ( List.map (fun (_,_,_,slb) -> slb) list_id )
+ in
+ let fresh_id s gl =
+ Tacmach.New.of_old begin fun gsig ->
+ let fresh = fresh_id (!avoid) s gsig in
+ avoid := fresh::(!avoid); fresh
+ end gl
+ in
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let fresh_first_intros = List.map (fun id -> fresh_id id gl) first_intros in
+ let freshn = fresh_id (Id.of_string "x") gl in
+ let freshm = fresh_id (Id.of_string "y") gl in
+ let freshz = fresh_id (Id.of_string "Z") gl in
+ (* try with *)
+ Tacticals.New.tclTHENLIST [ intros_using fresh_first_intros;
+ intro_using freshn ;
+ induct_on (mkVar freshn);
+ intro_using freshm;
+ destruct_on (mkVar freshm);
+ intro_using freshz;
+ intros;
+ Tacticals.New.tclTRY (
+ Tacticals.New.tclORELSE reflexivity (Equality.discr_tac false None)
+ );
+ Equality.inj None false None (mkVar freshz,NoBindings);
+ intros; simpl_in_concl;
+ Auto.default_auto;
+ Tacticals.New.tclREPEAT (
+ Tacticals.New.tclTHENLIST [apply (andb_true_intro());
+ simplest_split ;Auto.default_auto ]
+ );
+ Proofview.Goal.nf_enter { enter = begin fun gls ->
+ let gl = Proofview.Goal.concl gls in
+ (* assume the goal to be eq (eq_type ...) = true *)
+ match (kind_of_term gl) with
+ | App(c,ca) -> (match (kind_of_term ca.(1)) with
+ | App(c',ca') ->
+ let n = Array.length ca' in
+ do_replace_lb mode lb_scheme_key
+ (!avoid)
+ nparrec
+ ca'.(n-2) ca'.(n-1)
+ | _ ->
+ Tacticals.New.tclZEROMSG (str "Failure while solving Leibniz->Boolean.")
+ )
+ | _ ->
+ Tacticals.New.tclZEROMSG (str "Failure while solving Leibniz->Boolean.")
+ end }
+ ]
+ end }
+
+let lb_scheme_kind_aux = ref (fun () -> failwith "Undefined")
+
+let make_lb_scheme mode mind =
+ let mib = Global.lookup_mind mind in
+ if not (Int.equal (Array.length mib.mind_packets) 1) then
+ user_err
+ (str "Automatic building of Leibniz->boolean lemmas not supported");
+ let ind = (mind,0) in
+ let nparams = mib.mind_nparams in
+ let nparrec = mib.mind_nparams_rec in
+ let lnonparrec,lnamesparrec =
+ context_chop (nparams-nparrec) mib.mind_params_ctxt in
+ let lb_goal, eff = compute_lb_goal ind lnamesparrec nparrec in
+ let ctx = Evd.make_evar_universe_context (Global.env ()) None in
+ let side_eff = side_effect_of_mode mode in
+ let (ans, _, ctx) = Pfedit.build_by_tactic ~side_eff (Global.env()) ctx lb_goal
+ (compute_lb_tact mode (!lb_scheme_kind_aux()) ind lnamesparrec nparrec)
+ in
+ ([|ans|], ctx), eff
+
+let lb_scheme_kind = declare_mutual_scheme_object "_dec_lb" make_lb_scheme
+
+let _ = lb_scheme_kind_aux := fun () -> lb_scheme_kind
+
+(**********************************************************************)
+(* Decidable equality *)
+
+let check_not_is_defined () =
+ try ignore (Coqlib.build_coq_not ())
+ with e when CErrors.noncritical e -> raise (UndefinedCst "not")
+
+(* {n=m}+{n<>m} part *)
+let compute_dec_goal ind lnamesparrec nparrec =
+ check_not_is_defined ();
+ let eq = Lazy.force eq and tt = Lazy.force tt and bb = Lazy.force bb in
+ let list_id = list_id lnamesparrec in
+ let avoid = List.fold_right (Nameops.name_fold (fun id l -> id::l)) (List.map RelDecl.get_name lnamesparrec) [] in
+ let create_input c =
+ let x = next_ident_away (Id.of_string "x") avoid and
+ y = next_ident_away (Id.of_string "y") avoid in
+ let lb_typ = List.map (fun (s,seq,_,_) ->
+ mkNamedProd x (mkVar s) (
+ mkNamedProd y (mkVar s) (
+ mkArrow
+ ( mkApp(eq,[|mkVar s;mkVar x;mkVar y|]))
+ ( mkApp(eq,[|bb;mkApp(mkVar seq,[|mkVar x;mkVar y|]);tt|]))
+ ))
+ ) list_id in
+ let bl_typ = List.map (fun (s,seq,_,_) ->
+ mkNamedProd x (mkVar s) (
+ mkNamedProd y (mkVar s) (
+ mkArrow
+ ( mkApp(eq,[|bb;mkApp(mkVar seq,[|mkVar x;mkVar y|]);tt|]))
+ ( mkApp(eq,[|mkVar s;mkVar x;mkVar y|]))
+ ))
+ ) list_id in
+
+ let lb_input = List.fold_left2 ( fun a (s,_,_,slb) b ->
+ mkNamedProd slb b a
+ ) c (List.rev list_id) (List.rev lb_typ) in
+ let bl_input = List.fold_left2 ( fun a (s,_,sbl,_) b ->
+ mkNamedProd sbl b a
+ ) lb_input (List.rev list_id) (List.rev bl_typ) in
+
+ let eqs_typ = List.map (fun (s,_,_,_) ->
+ mkProd(Anonymous,mkVar s,mkProd(Anonymous,mkVar s,bb))
+ ) list_id in
+ let eq_input = List.fold_left2 ( fun a (s,seq,_,_) b ->
+ mkNamedProd seq b a
+ ) bl_input (List.rev list_id) (List.rev eqs_typ) in
+ List.fold_left (fun a decl -> mkNamedProd
+ (match RelDecl.get_name decl with Name s -> s | Anonymous -> Id.of_string "A")
+ (RelDecl.get_type decl) a) eq_input lnamesparrec
+ in
+ let n = next_ident_away (Id.of_string "x") avoid and
+ m = next_ident_away (Id.of_string "y") avoid in
+ let eqnm = mkApp(eq,[|mkFullInd ind (2*nparrec+2);mkVar n;mkVar m|]) in
+ create_input (
+ mkNamedProd n (mkFullInd ind (2*nparrec)) (
+ mkNamedProd m (mkFullInd ind (2*nparrec+1)) (
+ mkApp(sumbool(),[|eqnm;mkApp (Coqlib.build_coq_not(),[|eqnm|])|])
+ )
+ )
+ )
+
+let compute_dec_tact ind lnamesparrec nparrec =
+ let eq = Lazy.force eq and tt = Lazy.force tt
+ and ff = Lazy.force ff and bb = Lazy.force bb in
+ let list_id = list_id lnamesparrec in
+ let eqI, eff = eqI ind lnamesparrec in
+ let avoid = ref [] in
+ let eqtrue x = mkApp(eq,[|bb;x;tt|]) in
+ let eqfalse x = mkApp(eq,[|bb;x;ff|]) in
+ let first_intros =
+ ( List.map (fun (s,_,_,_) -> s ) list_id ) @
+ ( List.map (fun (_,seq,_,_) -> seq) list_id ) @
+ ( List.map (fun (_,_,sbl,_) -> sbl) list_id ) @
+ ( List.map (fun (_,_,_,slb) -> slb) list_id )
+ in
+ let fresh_id s gl =
+ Tacmach.New.of_old begin fun gsig ->
+ let fresh = fresh_id (!avoid) s gsig in
+ avoid := fresh::(!avoid); fresh
+ end gl
+ in
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let fresh_first_intros = List.map (fun id -> fresh_id id gl) first_intros in
+ let freshn = fresh_id (Id.of_string "x") gl in
+ let freshm = fresh_id (Id.of_string "y") gl in
+ let freshH = fresh_id (Id.of_string "H") gl in
+ let eqbnm = mkApp(eqI,[|mkVar freshn;mkVar freshm|]) in
+ let arfresh = Array.of_list fresh_first_intros in
+ let xargs = Array.sub arfresh 0 (2*nparrec) in
+ begin try
+ let c, eff = find_scheme bl_scheme_kind ind in
+ Proofview.tclUNIT (mkConst c,eff) with
+ Not_found ->
+ Tacticals.New.tclZEROMSG (str "Error during the decidability part, boolean to leibniz equality is required.")
+ end >>= fun (blI,eff') ->
+ begin try
+ let c, eff = find_scheme lb_scheme_kind ind in
+ Proofview.tclUNIT (mkConst c,eff) with
+ Not_found ->
+ Tacticals.New.tclZEROMSG (str "Error during the decidability part, leibniz to boolean equality is required.")
+ end >>= fun (lbI,eff'') ->
+ let eff = (Safe_typing.concat_private eff'' (Safe_typing.concat_private eff' eff)) in
+ Tacticals.New.tclTHENLIST [
+ Proofview.tclEFFECTS eff;
+ intros_using fresh_first_intros;
+ intros_using [freshn;freshm];
+ (*we do this so we don't have to prove the same goal twice *)
+ assert_by (Name freshH) (
+ mkApp(sumbool(),[|eqtrue eqbnm; eqfalse eqbnm|])
+ )
+ (Tacticals.New.tclTHEN (destruct_on eqbnm) Auto.default_auto);
+
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let freshH2 = fresh_id (Id.of_string "H") gl in
+ Tacticals.New.tclTHENS (destruct_on_using (mkVar freshH) freshH2) [
+ (* left *)
+ Tacticals.New.tclTHENLIST [
+ simplest_left;
+ apply (mkApp(blI,Array.map(fun x->mkVar x) xargs));
+ Auto.default_auto
+ ]
+ ;
+
+ (*right *)
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let freshH3 = fresh_id (Id.of_string "H") gl in
+ Tacticals.New.tclTHENLIST [
+ simplest_right ;
+ unfold_constr (Lazy.force Coqlib.coq_not_ref);
+ intro;
+ Equality.subst_all ();
+ assert_by (Name freshH3)
+ (mkApp(eq,[|bb;mkApp(eqI,[|mkVar freshm;mkVar freshm|]);tt|]))
+ (Tacticals.New.tclTHENLIST [
+ apply (mkApp(lbI,Array.map (fun x->mkVar x) xargs));
+ Auto.default_auto
+ ]);
+ Equality.general_rewrite_bindings_in true
+ Locus.AllOccurrences true false
+ (List.hd !avoid)
+ ((mkVar (List.hd (List.tl !avoid))),
+ NoBindings
+ )
+ true;
+ Equality.discr_tac false None
+ ]
+ end }
+ ]
+ end }
+ ]
+ end }
+
+let make_eq_decidability mode mind =
+ let mib = Global.lookup_mind mind in
+ if not (Int.equal (Array.length mib.mind_packets) 1) then
+ raise DecidabilityMutualNotSupported;
+ let ind = (mind,0) in
+ let nparams = mib.mind_nparams in
+ let nparrec = mib.mind_nparams_rec in
+ let u = Univ.Instance.empty in
+ let ctx = Evd.make_evar_universe_context (Global.env ()) None in
+ let lnonparrec,lnamesparrec =
+ context_chop (nparams-nparrec) mib.mind_params_ctxt in
+ let side_eff = side_effect_of_mode mode in
+ let (ans, _, ctx) = Pfedit.build_by_tactic ~side_eff (Global.env()) ctx
+ (compute_dec_goal (ind,u) lnamesparrec nparrec)
+ (compute_dec_tact ind lnamesparrec nparrec)
+ in
+ ([|ans|], ctx), Safe_typing.empty_private_constants
+
+let eq_dec_scheme_kind =
+ declare_mutual_scheme_object "_eq_dec" make_eq_decidability
+
+(* The eq_dec_scheme proofs depend on the equality and discr tactics
+ but the inj tactics, that comes with discr, depends on the
+ eq_dec_scheme... *)
+
+let _ = Equality.set_eq_dec_scheme_kind eq_dec_scheme_kind
diff --git a/vernac/auto_ind_decl.mli b/vernac/auto_ind_decl.mli
new file mode 100644
index 000000000..60232ba8f
--- /dev/null
+++ b/vernac/auto_ind_decl.mli
@@ -0,0 +1,42 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open Ind_tables
+
+(** This file is about the automatic generation of schemes about
+ decidable equality,
+ @author Vincent Siles
+ Oct 2007 *)
+
+(** {6 Build boolean equality of a block of mutual inductive types } *)
+
+exception EqNotFound of inductive * inductive
+exception EqUnknown of string
+exception UndefinedCst of string
+exception InductiveWithProduct
+exception InductiveWithSort
+exception ParameterWithoutEquality of Globnames.global_reference
+exception NonSingletonProp of inductive
+exception DecidabilityMutualNotSupported
+exception NoDecidabilityCoInductive
+
+val beq_scheme_kind : mutual scheme_kind
+val build_beq_scheme : mutual_scheme_object_function
+
+(** {6 Build equivalence between boolean equality and Leibniz equality } *)
+
+val lb_scheme_kind : mutual scheme_kind
+val make_lb_scheme : mutual_scheme_object_function
+val bl_scheme_kind : mutual scheme_kind
+val make_bl_scheme : mutual_scheme_object_function
+
+(** {6 Build decidability of equality } *)
+
+val eq_dec_scheme_kind : mutual scheme_kind
+val make_eq_decidability : mutual_scheme_object_function
diff --git a/vernac/class.ml b/vernac/class.ml
new file mode 100644
index 000000000..0dc799014
--- /dev/null
+++ b/vernac/class.ml
@@ -0,0 +1,328 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open CErrors
+open Util
+open Pp
+open Names
+open Term
+open Vars
+open Termops
+open Entries
+open Environ
+open Classops
+open Declare
+open Globnames
+open Nametab
+open Decl_kinds
+
+let strength_min l = if List.mem `LOCAL l then `LOCAL else `GLOBAL
+
+let loc_of_bool b = if b then `LOCAL else `GLOBAL
+
+(* Errors *)
+
+type coercion_error_kind =
+ | AlreadyExists
+ | NotAFunction
+ | NoSource of cl_typ option
+ | ForbiddenSourceClass of cl_typ
+ | NoTarget
+ | WrongTarget of cl_typ * cl_typ
+ | NotAClass of global_reference
+
+exception CoercionError of coercion_error_kind
+
+let explain_coercion_error g = function
+ | AlreadyExists ->
+ (Printer.pr_global g ++ str" is already a coercion")
+ | NotAFunction ->
+ (Printer.pr_global g ++ str" is not a function")
+ | NoSource (Some cl) ->
+ (str "Cannot recognize " ++ pr_class cl ++ str " as a source class of "
+ ++ Printer.pr_global g)
+ | NoSource None ->
+ (str ": cannot find the source class of " ++ Printer.pr_global g)
+ | ForbiddenSourceClass cl ->
+ pr_class cl ++ str " cannot be a source class"
+ | NoTarget ->
+ (str"Cannot find the target class")
+ | WrongTarget (clt,cl) ->
+ (str"Found target class " ++ pr_class cl ++
+ str " instead of " ++ pr_class clt)
+ | NotAClass ref ->
+ (str "Type of " ++ Printer.pr_global ref ++
+ str " does not end with a sort")
+
+(* Verifications pour l'ajout d'une classe *)
+
+let check_reference_arity ref =
+ if not (Reductionops.is_arity (Global.env()) Evd.empty (Global.type_of_global_unsafe ref)) then
+ raise (CoercionError (NotAClass ref))
+
+let check_arity = function
+ | CL_FUN | CL_SORT -> ()
+ | CL_CONST cst -> check_reference_arity (ConstRef cst)
+ | CL_PROJ cst -> check_reference_arity (ConstRef cst)
+ | CL_SECVAR id -> check_reference_arity (VarRef id)
+ | CL_IND kn -> check_reference_arity (IndRef kn)
+
+(* Coercions *)
+
+(* check that the computed target is the provided one *)
+let check_target clt = function
+ | Some cl when not (cl_typ_eq cl clt) -> raise (CoercionError (WrongTarget(clt,cl)))
+ | _ -> ()
+
+(* condition d'heritage uniforme *)
+
+let uniform_cond nargs lt =
+ let rec aux = function
+ | (0,[]) -> true
+ | (n,t::l) ->
+ let t = strip_outer_cast t in
+ isRel t && Int.equal (destRel t) n && aux ((n-1),l)
+ | _ -> false
+ in
+ aux (nargs,lt)
+
+let class_of_global = function
+ | ConstRef sp ->
+ if Environ.is_projection sp (Global.env ())
+ then CL_PROJ sp else CL_CONST sp
+ | IndRef sp -> CL_IND sp
+ | VarRef id -> CL_SECVAR id
+ | ConstructRef _ as c ->
+ user_err ~hdr:"class_of_global"
+ (str "Constructors, such as " ++ Printer.pr_global c ++
+ str ", cannot be used as a class.")
+
+(*
+lp est la liste (inverse'e) des arguments de la coercion
+ids est le nom de la classe source
+sps_opt est le sp de la classe source dans le cas des structures
+retourne:
+la classe source
+nbre d'arguments de la classe
+le constr de la class
+la liste des variables dont depend la classe source
+l'indice de la classe source dans la liste lp
+*)
+
+let get_source lp source =
+ match source with
+ | None ->
+ let (cl1,u1,lv1) =
+ match lp with
+ | [] -> raise Not_found
+ | t1::_ -> find_class_type Evd.empty t1
+ in
+ (cl1,u1,lv1,1)
+ | Some cl ->
+ let rec aux = function
+ | [] -> raise Not_found
+ | t1::lt ->
+ try
+ let cl1,u1,lv1 = find_class_type Evd.empty t1 in
+ if cl_typ_eq cl cl1 then cl1,u1,lv1,(List.length lt+1)
+ else raise Not_found
+ with Not_found -> aux lt
+ in aux (List.rev lp)
+
+let get_target t ind =
+ if (ind > 1) then
+ CL_FUN
+ else
+ match pi1 (find_class_type Evd.empty t) with
+ | CL_CONST p when Environ.is_projection p (Global.env ()) ->
+ CL_PROJ p
+ | x -> x
+
+
+let prods_of t =
+ let rec aux acc d = match kind_of_term d with
+ | Prod (_,c1,c2) -> aux (c1::acc) c2
+ | Cast (c,_,_) -> aux acc c
+ | _ -> (d,acc)
+ in
+ aux [] t
+
+let strength_of_cl = function
+ | CL_CONST kn -> `GLOBAL
+ | CL_SECVAR id -> `LOCAL
+ | _ -> `GLOBAL
+
+let strength_of_global = function
+ | VarRef _ -> `LOCAL
+ | _ -> `GLOBAL
+
+let get_strength stre ref cls clt =
+ let stres = strength_of_cl cls in
+ let stret = strength_of_cl clt in
+ let stref = strength_of_global ref in
+ strength_min [stre;stres;stret;stref]
+
+let ident_key_of_class = function
+ | CL_FUN -> "Funclass"
+ | CL_SORT -> "Sortclass"
+ | CL_CONST sp | CL_PROJ sp -> Label.to_string (con_label sp)
+ | CL_IND (sp,_) -> Label.to_string (mind_label sp)
+ | CL_SECVAR id -> Id.to_string id
+
+(* Identity coercion *)
+
+let error_not_transparent source =
+ user_err ~hdr:"build_id_coercion"
+ (pr_class source ++ str " must be a transparent constant.")
+
+let build_id_coercion idf_opt source poly =
+ let env = Global.env () in
+ let sigma = Evd.from_env env in
+ let sigma, vs = match source with
+ | CL_CONST sp -> Evd.fresh_global env sigma (ConstRef sp)
+ | _ -> error_not_transparent source in
+ let c = match constant_opt_value_in env (destConst vs) with
+ | Some c -> c
+ | None -> error_not_transparent source in
+ let lams,t = decompose_lam_assum c in
+ let val_f =
+ it_mkLambda_or_LetIn
+ (mkLambda (Name Namegen.default_dependent_ident,
+ applistc vs (Context.Rel.to_extended_list 0 lams),
+ mkRel 1))
+ lams
+ in
+ let typ_f =
+ it_mkProd_wo_LetIn
+ (mkProd (Anonymous, applistc vs (Context.Rel.to_extended_list 0 lams), lift 1 t))
+ lams
+ in
+ (* juste pour verification *)
+ let _ =
+ if not
+ (Reductionops.is_conv_leq env sigma
+ (Typing.unsafe_type_of env sigma val_f) typ_f)
+ then
+ user_err (strbrk
+ "Cannot be defined as coercion (maybe a bad number of arguments).")
+ in
+ let idf =
+ match idf_opt with
+ | Some idf -> idf
+ | None ->
+ let cl,u,_ = find_class_type sigma t in
+ Id.of_string ("Id_"^(ident_key_of_class source)^"_"^
+ (ident_key_of_class cl))
+ in
+ let constr_entry = (* Cast is necessary to express [val_f] is identity *)
+ DefinitionEntry
+ (definition_entry ~types:typ_f ~poly ~univs:(snd (Evd.universe_context sigma))
+ ~inline:true (mkCast (val_f, DEFAULTcast, typ_f)))
+ in
+ let decl = (constr_entry, IsDefinition IdentityCoercion) in
+ let kn = declare_constant idf decl in
+ ConstRef kn
+
+let check_source = function
+| Some (CL_FUN|CL_SORT as s) -> raise (CoercionError (ForbiddenSourceClass s))
+| _ -> ()
+
+(*
+nom de la fonction coercion
+strength de f
+nom de la classe source (optionnel)
+sp de la classe source (dans le cas des structures)
+nom de la classe target (optionnel)
+booleen "coercion identite'?"
+
+lorque source est None alors target est None aussi.
+*)
+
+let warn_uniform_inheritance =
+ CWarnings.create ~name:"uniform-inheritance" ~category:"typechecker"
+ (fun g ->
+ Printer.pr_global g ++
+ strbrk" does not respect the uniform inheritance condition")
+
+let add_new_coercion_core coef stre poly source target isid =
+ check_source source;
+ let t = Global.type_of_global_unsafe coef in
+ if coercion_exists coef then raise (CoercionError AlreadyExists);
+ let tg,lp = prods_of t in
+ let llp = List.length lp in
+ if Int.equal llp 0 then raise (CoercionError NotAFunction);
+ let (cls,us,lvs,ind) =
+ try
+ get_source lp source
+ with Not_found ->
+ raise (CoercionError (NoSource source))
+ in
+ check_source (Some cls);
+ if not (uniform_cond (llp-ind) lvs) then
+ warn_uniform_inheritance coef;
+ let clt =
+ try
+ get_target tg ind
+ with Not_found ->
+ raise (CoercionError NoTarget)
+ in
+ check_target clt target;
+ check_arity cls;
+ check_arity clt;
+ let local = match get_strength stre coef cls clt with
+ | `LOCAL -> true
+ | `GLOBAL -> false
+ in
+ declare_coercion coef ~local ~isid ~src:cls ~target:clt ~params:(List.length lvs)
+
+
+let try_add_new_coercion_core ref ~local c d e f =
+ try add_new_coercion_core ref (loc_of_bool local) c d e f
+ with CoercionError e ->
+ user_err ~hdr:"try_add_new_coercion_core"
+ (explain_coercion_error ref e ++ str ".")
+
+let try_add_new_coercion ref ~local poly =
+ try_add_new_coercion_core ref ~local poly None None false
+
+let try_add_new_coercion_subclass cl ~local poly =
+ let coe_ref = build_id_coercion None cl poly in
+ try_add_new_coercion_core coe_ref ~local poly (Some cl) None true
+
+let try_add_new_coercion_with_target ref ~local poly ~source ~target =
+ try_add_new_coercion_core ref ~local poly (Some source) (Some target) false
+
+let try_add_new_identity_coercion id ~local poly ~source ~target =
+ let ref = build_id_coercion (Some id) source poly in
+ try_add_new_coercion_core ref ~local poly (Some source) (Some target) true
+
+let try_add_new_coercion_with_source ref ~local poly ~source =
+ try_add_new_coercion_core ref ~local poly (Some source) None false
+
+let add_coercion_hook poly local ref =
+ let stre = match local with
+ | Local -> true
+ | Global -> false
+ | Discharge -> assert false
+ in
+ let () = try_add_new_coercion ref stre poly in
+ let msg = pr_global_env Id.Set.empty ref ++ str " is now a coercion" in
+ Flags.if_verbose Feedback.msg_info msg
+
+let add_coercion_hook poly = Lemmas.mk_hook (add_coercion_hook poly)
+
+let add_subclass_hook poly local ref =
+ let stre = match local with
+ | Local -> true
+ | Global -> false
+ | Discharge -> assert false
+ in
+ let cl = class_of_global ref in
+ try_add_new_coercion_subclass cl stre poly
+
+let add_subclass_hook poly = Lemmas.mk_hook (add_subclass_hook poly)
diff --git a/vernac/class.mli b/vernac/class.mli
new file mode 100644
index 000000000..5f9ae28f6
--- /dev/null
+++ b/vernac/class.mli
@@ -0,0 +1,48 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open Classops
+open Globnames
+
+(** Classes and coercions. *)
+
+(** [try_add_new_coercion_with_target ref s src tg] declares [ref] as a coercion
+ from [src] to [tg] *)
+val try_add_new_coercion_with_target : global_reference -> local:bool ->
+ Decl_kinds.polymorphic ->
+ source:cl_typ -> target:cl_typ -> unit
+
+(** [try_add_new_coercion ref s] declares [ref], assumed to be of type
+ [(x1:T1)...(xn:Tn)src->tg], as a coercion from [src] to [tg] *)
+val try_add_new_coercion : global_reference -> local:bool ->
+ Decl_kinds.polymorphic -> unit
+
+(** [try_add_new_coercion_subclass cst s] expects that [cst] denotes a
+ transparent constant which unfolds to some class [tg]; it declares
+ an identity coercion from [cst] to [tg], named something like
+ ["Id_cst_tg"] *)
+val try_add_new_coercion_subclass : cl_typ -> local:bool ->
+ Decl_kinds.polymorphic -> unit
+
+(** [try_add_new_coercion_with_source ref s src] declares [ref] as a coercion
+ from [src] to [tg] where the target is inferred from the type of [ref] *)
+val try_add_new_coercion_with_source : global_reference -> local:bool ->
+ Decl_kinds.polymorphic -> source:cl_typ -> unit
+
+(** [try_add_new_identity_coercion id s src tg] enriches the
+ environment with a new definition of name [id] declared as an
+ identity coercion from [src] to [tg] *)
+val try_add_new_identity_coercion : Id.t -> local:bool ->
+ Decl_kinds.polymorphic -> source:cl_typ -> target:cl_typ -> unit
+
+val add_coercion_hook : Decl_kinds.polymorphic -> unit Lemmas.declaration_hook
+
+val add_subclass_hook : Decl_kinds.polymorphic -> unit Lemmas.declaration_hook
+
+val class_of_global : global_reference -> cl_typ
diff --git a/vernac/classes.ml b/vernac/classes.ml
new file mode 100644
index 000000000..6512f3def
--- /dev/null
+++ b/vernac/classes.ml
@@ -0,0 +1,410 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i*)
+open Names
+open Term
+open Vars
+open Environ
+open Nametab
+open CErrors
+open Util
+open Typeclasses_errors
+open Typeclasses
+open Libnames
+open Globnames
+open Constrintern
+open Constrexpr
+open Sigma.Notations
+open Context.Rel.Declaration
+
+module RelDecl = Context.Rel.Declaration
+(*i*)
+
+open Decl_kinds
+open Entries
+
+let refine_instance = ref true
+
+let _ = Goptions.declare_bool_option {
+ Goptions.optsync = true;
+ Goptions.optdepr = false;
+ Goptions.optname = "definition of instances by refining";
+ Goptions.optkey = ["Refine";"Instance";"Mode"];
+ Goptions.optread = (fun () -> !refine_instance);
+ Goptions.optwrite = (fun b -> refine_instance := b)
+}
+
+let typeclasses_db = "typeclass_instances"
+
+let set_typeclass_transparency c local b =
+ Hints.add_hints local [typeclasses_db]
+ (Hints.HintsTransparencyEntry ([c], b))
+
+let _ =
+ Hook.set Typeclasses.add_instance_hint_hook
+ (fun inst path local info poly ->
+ let inst' = match inst with IsConstr c -> Hints.IsConstr (c, Univ.ContextSet.empty)
+ | IsGlobal gr -> Hints.IsGlobRef gr
+ in
+ let info =
+ let open Vernacexpr in
+ { info with hint_pattern =
+ Option.map
+ (Constrintern.intern_constr_pattern (Global.env()))
+ info.hint_pattern } in
+ Flags.silently (fun () ->
+ Hints.add_hints local [typeclasses_db]
+ (Hints.HintsResolveEntry
+ [info, poly, false, Hints.PathHints path, inst'])) ());
+ Hook.set Typeclasses.set_typeclass_transparency_hook set_typeclass_transparency;
+ Hook.set Typeclasses.classes_transparent_state_hook
+ (fun () -> Hints.Hint_db.transparent_state (Hints.searchtable_map typeclasses_db))
+
+open Vernacexpr
+
+(** TODO: add subinstances *)
+let existing_instance glob g info =
+ let c = global g in
+ let info = Option.default Hints.empty_hint_info info in
+ let instance = Global.type_of_global_unsafe c in
+ let _, r = decompose_prod_assum instance in
+ match class_of_constr r with
+ | Some (_, ((tc,u), _)) -> add_instance (new_instance tc info glob
+ (*FIXME*) (Flags.use_polymorphic_flag ()) c)
+ | None -> user_err ~loc:(loc_of_reference g)
+ ~hdr:"declare_instance"
+ (Pp.str "Constant does not build instances of a declared type class.")
+
+let mismatched_params env n m = mismatched_ctx_inst env Parameters n m
+let mismatched_props env n m = mismatched_ctx_inst env Properties n m
+
+(* Declare everything in the parameters as implicit, and the class instance as well *)
+
+let type_ctx_instance evars env ctx inst subst =
+ let rec aux (subst, instctx) l = function
+ decl :: ctx ->
+ let t' = substl subst (RelDecl.get_type decl) in
+ let c', l =
+ match decl with
+ | LocalAssum _ -> interp_casted_constr_evars env evars (List.hd l) t', List.tl l
+ | LocalDef (_,b,_) -> substl subst b, l
+ in
+ let d = RelDecl.get_name decl, Some c', t' in
+ aux (c' :: subst, d :: instctx) l ctx
+ | [] -> subst
+ in aux (subst, []) inst (List.rev ctx)
+
+let id_of_class cl =
+ match cl.cl_impl with
+ | ConstRef kn -> let _,_,l = repr_con kn in Label.to_id l
+ | IndRef (kn,i) ->
+ let mip = (Environ.lookup_mind kn (Global.env ())).Declarations.mind_packets in
+ mip.(0).Declarations.mind_typename
+ | _ -> assert false
+
+open Pp
+
+let instance_hook k info global imps ?hook cst =
+ Impargs.maybe_declare_manual_implicits false cst ~enriching:false imps;
+ Typeclasses.declare_instance (Some info) (not global) cst;
+ (match hook with Some h -> h cst | None -> ())
+
+let declare_instance_constant k info global imps ?hook id pl poly evm term termtype =
+ let kind = IsDefinition Instance in
+ let evm =
+ let levels = Univ.LSet.union (Universes.universes_of_constr termtype)
+ (Universes.universes_of_constr term) in
+ Evd.restrict_universe_context evm levels
+ in
+ let pl, uctx = Evd.universe_context ?names:pl evm in
+ let entry =
+ Declare.definition_entry ~types:termtype ~poly ~univs:uctx term
+ in
+ let cdecl = (DefinitionEntry entry, kind) in
+ let kn = Declare.declare_constant id cdecl in
+ Declare.definition_message id;
+ Universes.register_universe_binders (ConstRef kn) pl;
+ instance_hook k info global imps ?hook (ConstRef kn);
+ id
+
+let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) poly ctx (instid, bk, cl) props
+ ?(generalize=true)
+ ?(tac:unit Proofview.tactic option) ?hook pri =
+ let env = Global.env() in
+ let ((loc, instid), pl) = instid in
+ let uctx = Evd.make_evar_universe_context env pl in
+ let evars = ref (Evd.from_ctx uctx) in
+ let tclass, ids =
+ match bk with
+ | Decl_kinds.Implicit ->
+ Implicit_quantifiers.implicit_application Id.Set.empty ~allow_partial:false
+ (fun avoid (clname, _) ->
+ match clname with
+ | Some (cl, b) ->
+ let t = CHole (Loc.ghost, None, Misctypes.IntroAnonymous, None) in
+ t, avoid
+ | None -> failwith ("new instance: under-applied typeclass"))
+ cl
+ | Explicit -> cl, Id.Set.empty
+ in
+ let tclass =
+ if generalize then CGeneralization (Loc.ghost, Implicit, Some AbsPi, tclass)
+ else tclass
+ in
+ let k, u, cty, ctx', ctx, len, imps, subst =
+ let impls, ((env', ctx), imps) = interp_context_evars env evars ctx in
+ let c', imps' = interp_type_evars_impls ~impls env' evars tclass in
+ let len = List.length ctx in
+ let imps = imps @ Impargs.lift_implicits len imps' in
+ let ctx', c = decompose_prod_assum c' in
+ let ctx'' = ctx' @ ctx in
+ let k, args = Typeclasses.dest_class_app (push_rel_context ctx'' env) c in
+ let cl, u = Typeclasses.typeclass_univ_instance k in
+ let _, args =
+ List.fold_right (fun decl (args, args') ->
+ match decl with
+ | LocalAssum _ -> (List.tl args, List.hd args :: args')
+ | LocalDef (_,b,_) -> (args, substl args' b :: args'))
+ (snd cl.cl_context) (args, [])
+ in
+ cl, u, c', ctx', ctx, len, imps, args
+ in
+ let id =
+ match instid with
+ Name id ->
+ let sp = Lib.make_path id in
+ if Nametab.exists_cci sp then
+ user_err ~hdr:"new_instance" (Nameops.pr_id id ++ Pp.str " already exists.");
+ id
+ | Anonymous ->
+ let i = Nameops.add_suffix (id_of_class k) "_instance_0" in
+ Namegen.next_global_ident_away i (Termops.ids_of_context env)
+ in
+ let env' = push_rel_context ctx env in
+ evars := Evarutil.nf_evar_map !evars;
+ evars := resolve_typeclasses ~filter:Typeclasses.all_evars ~fail:true env !evars;
+ let subst = List.map (Evarutil.nf_evar !evars) subst in
+ if abstract then
+ begin
+ let subst = List.fold_left2
+ (fun subst' s decl -> if is_local_assum decl then s :: subst' else subst')
+ [] subst (snd k.cl_context)
+ in
+ let (_, ty_constr) = instance_constructor (k,u) subst in
+ let nf, subst = Evarutil.e_nf_evars_and_universes evars in
+ let termtype =
+ let t = it_mkProd_or_LetIn ty_constr (ctx' @ ctx) in
+ nf t
+ in
+ Pretyping.check_evars env Evd.empty !evars termtype;
+ let pl, ctx = Evd.universe_context ?names:pl !evars in
+ let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id
+ (ParameterEntry
+ (None,poly,(termtype,ctx),None), Decl_kinds.IsAssumption Decl_kinds.Logical)
+ in
+ Universes.register_universe_binders (ConstRef cst) pl;
+ instance_hook k pri global imps ?hook (ConstRef cst); id
+ end
+ else (
+ let props =
+ match props with
+ | Some (true, CRecord (loc, fs)) ->
+ if List.length fs > List.length k.cl_props then
+ mismatched_props env' (List.map snd fs) k.cl_props;
+ Some (Inl fs)
+ | Some (_, t) -> Some (Inr t)
+ | None ->
+ if Flags.is_program_mode () then Some (Inl [])
+ else None
+ in
+ let subst =
+ match props with
+ | None -> if List.is_empty k.cl_props then Some (Inl subst) else None
+ | Some (Inr term) ->
+ let c = interp_casted_constr_evars env' evars term cty in
+ Some (Inr (c, subst))
+ | Some (Inl props) ->
+ let get_id =
+ function
+ | Ident id' -> id'
+ | Qualid (loc,id') -> (loc, snd (repr_qualid id'))
+ in
+ let props, rest =
+ List.fold_left
+ (fun (props, rest) decl ->
+ if is_local_assum decl then
+ try
+ let is_id (id', _) = match RelDecl.get_name decl, get_id id' with
+ | Name id, (_, id') -> Id.equal id id'
+ | Anonymous, _ -> false
+ in
+ let (loc_mid, c) =
+ List.find is_id rest
+ in
+ let rest' =
+ List.filter (fun v -> not (is_id v)) rest
+ in
+ let (loc, mid) = get_id loc_mid in
+ List.iter (fun (n, _, x) ->
+ if Name.equal n (Name mid) then
+ Option.iter (fun x -> Dumpglob.add_glob loc (ConstRef x)) x)
+ k.cl_projs;
+ c :: props, rest'
+ with Not_found ->
+ (CHole (Loc.ghost, None(* Some Evar_kinds.GoalEvar *), Misctypes.IntroAnonymous, None) :: props), rest
+ else props, rest)
+ ([], props) k.cl_props
+ in
+ match rest with
+ | (n, _) :: _ ->
+ unbound_method env' k.cl_impl (get_id n)
+ | _ ->
+ Some (Inl (type_ctx_instance evars (push_rel_context ctx' env')
+ k.cl_props props subst))
+ in
+ let term, termtype =
+ match subst with
+ | None -> let termtype = it_mkProd_or_LetIn cty ctx in
+ None, termtype
+ | Some (Inl subst) ->
+ let subst = List.fold_left2
+ (fun subst' s decl -> if is_local_assum decl then s :: subst' else subst')
+ [] subst (k.cl_props @ snd k.cl_context)
+ in
+ let (app, ty_constr) = instance_constructor (k,u) subst in
+ let termtype = it_mkProd_or_LetIn ty_constr (ctx' @ ctx) in
+ let term = Termops.it_mkLambda_or_LetIn (Option.get app) (ctx' @ ctx) in
+ Some term, termtype
+ | Some (Inr (def, subst)) ->
+ let termtype = it_mkProd_or_LetIn cty ctx in
+ let term = Termops.it_mkLambda_or_LetIn def ctx in
+ Some term, termtype
+ in
+ let _ =
+ evars := Evarutil.nf_evar_map !evars;
+ evars := Typeclasses.resolve_typeclasses ~filter:Typeclasses.no_goals_or_obligations ~fail:true
+ env !evars;
+ (* Try resolving fields that are typeclasses automatically. *)
+ evars := Typeclasses.resolve_typeclasses ~filter:Typeclasses.all_evars ~fail:false
+ env !evars
+ in
+ let _ = evars := Evarutil.nf_evar_map_undefined !evars in
+ let evm, nf = Evarutil.nf_evar_map_universes !evars in
+ let termtype = nf termtype in
+ let _ = (* Check that the type is free of evars now. *)
+ Pretyping.check_evars env Evd.empty evm termtype
+ in
+ let term = Option.map nf term in
+ if not (Evd.has_undefined evm) && not (Option.is_empty term) then
+ declare_instance_constant k pri global imps ?hook id pl
+ poly evm (Option.get term) termtype
+ else if Flags.is_program_mode () || refine || Option.is_empty term then begin
+ let kind = Decl_kinds.Global, poly, Decl_kinds.DefinitionBody Decl_kinds.Instance in
+ if Flags.is_program_mode () then
+ let hook vis gr _ =
+ let cst = match gr with ConstRef kn -> kn | _ -> assert false in
+ Impargs.declare_manual_implicits false gr ~enriching:false [imps];
+ Typeclasses.declare_instance (Some pri) (not global) (ConstRef cst)
+ in
+ let obls, constr, typ =
+ match term with
+ | Some t ->
+ let obls, _, constr, typ =
+ Obligations.eterm_obligations env id evm 0 t termtype
+ in obls, Some constr, typ
+ | None -> [||], None, termtype
+ in
+ let hook = Lemmas.mk_hook hook in
+ let ctx = Evd.evar_universe_context evm in
+ ignore (Obligations.add_definition id ?term:constr
+ ?pl typ ctx ~kind:(Global,poly,Instance) ~hook obls);
+ id
+ else
+ (Flags.silently
+ (fun () ->
+ (* spiwack: it is hard to reorder the actions to do
+ the pretyping after the proof has opened. As a
+ consequence, we use the low-level primitives to code
+ the refinement manually.*)
+ let gls = List.rev (Evd.future_goals evm) in
+ let evm = Evd.reset_future_goals evm in
+ Lemmas.start_proof id kind evm termtype
+ (Lemmas.mk_hook
+ (fun _ -> instance_hook k pri global imps ?hook));
+ (* spiwack: I don't know what to do with the status here. *)
+ if not (Option.is_empty term) then
+ let init_refine =
+ Tacticals.New.tclTHENLIST [
+ Refine.refine { run = fun evm -> Sigma (Option.get term, evm, Sigma.refl) };
+ Proofview.Unsafe.tclNEWGOALS gls;
+ Tactics.New.reduce_after_refine;
+ ]
+ in
+ ignore (Pfedit.by init_refine)
+ else if Flags.is_auto_intros () then
+ ignore (Pfedit.by (Tacticals.New.tclDO len Tactics.intro));
+ (match tac with Some tac -> ignore (Pfedit.by tac) | None -> ())) ();
+ id)
+ end
+ else CErrors.error "Unsolved obligations remaining.")
+
+let named_of_rel_context l =
+ let acc, ctx =
+ List.fold_right
+ (fun decl (subst, ctx) ->
+ let id = match RelDecl.get_name decl with Anonymous -> invalid_arg "named_of_rel_context" | Name id -> id in
+ let d = match decl with
+ | LocalAssum (_,t) -> id, None, substl subst t
+ | LocalDef (_,b,t) -> id, Some (substl subst b), substl subst t in
+ (mkVar id :: subst, d :: ctx))
+ l ([], [])
+ in ctx
+
+let context poly l =
+ let env = Global.env() in
+ let evars = ref (Evd.from_env env) in
+ let _, ((env', fullctx), impls) = interp_context_evars env evars l in
+ let subst = Evarutil.evd_comb0 Evarutil.nf_evars_and_universes evars in
+ let fullctx = Context.Rel.map subst fullctx in
+ let ce t = Pretyping.check_evars env Evd.empty !evars t in
+ let () = List.iter (fun decl -> Context.Rel.Declaration.iter_constr ce decl) fullctx in
+ let ctx =
+ try named_of_rel_context fullctx
+ with e when CErrors.noncritical e ->
+ error "Anonymous variables not allowed in contexts."
+ in
+ let uctx = ref (Evd.universe_context_set !evars) in
+ let fn status (id, b, t) =
+ if Lib.is_modtype () && not (Lib.sections_are_opened ()) then
+ let ctx = Univ.ContextSet.to_context !uctx in
+ (* Declare the universe context once *)
+ let () = uctx := Univ.ContextSet.empty in
+ let decl = (ParameterEntry (None,poly,(t,ctx),None), IsAssumption Logical) in
+ let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id decl in
+ match class_of_constr t with
+ | Some (rels, ((tc,_), args) as _cl) ->
+ add_instance (Typeclasses.new_instance tc Hints.empty_hint_info false (*FIXME*)
+ poly (ConstRef cst));
+ status
+ (* declare_subclasses (ConstRef cst) cl *)
+ | None -> status
+ else
+ let test (x, _) = match x with
+ | ExplByPos (_, Some id') -> Id.equal id id'
+ | _ -> false
+ in
+ let impl = List.exists test impls in
+ let decl = (Discharge, poly, Definitional) in
+ let nstatus =
+ pi3 (Command.declare_assumption false decl (t, !uctx) [] [] impl
+ Vernacexpr.NoInline (Loc.ghost, id))
+ in
+ let () = uctx := Univ.ContextSet.empty in
+ status && nstatus
+ in List.fold_left fn true (List.rev ctx)
diff --git a/vernac/classes.mli b/vernac/classes.mli
new file mode 100644
index 000000000..d2cb788ea
--- /dev/null
+++ b/vernac/classes.mli
@@ -0,0 +1,66 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open Environ
+open Constrexpr
+open Typeclasses
+open Libnames
+
+(** Errors *)
+
+val mismatched_params : env -> constr_expr list -> Context.Rel.t -> 'a
+
+val mismatched_props : env -> constr_expr list -> Context.Rel.t -> 'a
+
+(** Instance declaration *)
+
+val existing_instance : bool -> reference -> Vernacexpr.hint_info_expr option -> unit
+(** globality, reference, optional priority and pattern information *)
+
+val declare_instance_constant :
+ typeclass ->
+ Vernacexpr.hint_info_expr -> (** priority *)
+ bool -> (** globality *)
+ Impargs.manual_explicitation list -> (** implicits *)
+ ?hook:(Globnames.global_reference -> unit) ->
+ Id.t -> (** name *)
+ Id.t Loc.located list option ->
+ bool -> (* polymorphic *)
+ Evd.evar_map -> (* Universes *)
+ Constr.t -> (** body *)
+ Term.types -> (** type *)
+ Names.Id.t
+
+val new_instance :
+ ?abstract:bool -> (** Not abstract by default. *)
+ ?global:bool -> (** Not global by default. *)
+ ?refine:bool -> (** Allow refinement *)
+ Decl_kinds.polymorphic ->
+ local_binder list ->
+ typeclass_constraint ->
+ (bool * constr_expr) option ->
+ ?generalize:bool ->
+ ?tac:unit Proofview.tactic ->
+ ?hook:(Globnames.global_reference -> unit) ->
+ Vernacexpr.hint_info_expr ->
+ Id.t
+
+(** Setting opacity *)
+
+val set_typeclass_transparency : evaluable_global_reference -> bool -> bool -> unit
+
+(** For generation on names based on classes only *)
+
+val id_of_class : typeclass -> Id.t
+
+(** Context command *)
+
+(** returns [false] if, for lack of section, it declares an assumption
+ (unless in a module type). *)
+val context : Decl_kinds.polymorphic -> local_binder list -> bool
diff --git a/vernac/command.ml b/vernac/command.ml
new file mode 100644
index 000000000..049f58aa2
--- /dev/null
+++ b/vernac/command.ml
@@ -0,0 +1,1357 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Pp
+open CErrors
+open Util
+open Flags
+open Term
+open Vars
+open Termops
+open Environ
+open Redexpr
+open Declare
+open Names
+open Libnames
+open Globnames
+open Nameops
+open Constrexpr
+open Constrexpr_ops
+open Topconstr
+open Constrintern
+open Nametab
+open Impargs
+open Reductionops
+open Indtypes
+open Decl_kinds
+open Pretyping
+open Evarutil
+open Evarconv
+open Indschemes
+open Misctypes
+open Vernacexpr
+open Sigma.Notations
+open Context.Rel.Declaration
+open Entries
+
+module RelDecl = Context.Rel.Declaration
+
+let do_universe poly l = Declare.do_universe poly l
+let do_constraint poly l = Declare.do_constraint poly l
+
+let rec under_binders env sigma f n c =
+ if Int.equal n 0 then f env sigma c else
+ match kind_of_term c with
+ | Lambda (x,t,c) ->
+ mkLambda (x,t,under_binders (push_rel (LocalAssum (x,t)) env) sigma f (n-1) c)
+ | LetIn (x,b,t,c) ->
+ mkLetIn (x,b,t,under_binders (push_rel (LocalDef (x,b,t)) env) sigma f (n-1) c)
+ | _ -> assert false
+
+let rec complete_conclusion a cs = function
+ | CProdN (loc,bl,c) -> CProdN (loc,bl,complete_conclusion a cs c)
+ | CLetIn (loc,b,t,c) -> CLetIn (loc,b,t,complete_conclusion a cs c)
+ | CHole (loc, k, _, _) ->
+ let (has_no_args,name,params) = a in
+ if not has_no_args then
+ user_err ~loc
+ (strbrk"Cannot infer the non constant arguments of the conclusion of "
+ ++ pr_id cs ++ str ".");
+ let args = List.map (fun id -> CRef(Ident(loc,id),None)) params in
+ CAppExpl (loc,(None,Ident(loc,name),None),List.rev args)
+ | c -> c
+
+(* Commands of the interface *)
+
+(* 1| Constant definitions *)
+
+let red_constant_entry n ce sigma = function
+ | None -> ce
+ | Some red ->
+ let proof_out = ce.const_entry_body in
+ let env = Global.env () in
+ let (redfun, _) = reduction_of_red_expr env red in
+ let redfun env sigma c =
+ let sigma = Sigma.Unsafe.of_evar_map sigma in
+ let Sigma (c, _, _) = redfun.e_redfun env sigma c in
+ c
+ in
+ { ce with const_entry_body = Future.chain ~greedy:true ~pure:true proof_out
+ (fun ((body,ctx),eff) -> (under_binders env sigma redfun n body,ctx),eff) }
+
+let warn_implicits_in_term =
+ CWarnings.create ~name:"implicits-in-term" ~category:"implicits"
+ (fun () ->
+ strbrk "Implicit arguments declaration relies on type." ++ spc () ++
+ strbrk "The term declares more implicits than the type here.")
+
+let interp_definition pl bl p red_option c ctypopt =
+ let env = Global.env() in
+ let ctx = Evd.make_evar_universe_context env pl in
+ let evdref = ref (Evd.from_ctx ctx) in
+ let impls, ((env_bl, ctx), imps1) = interp_context_evars env evdref bl in
+ let nb_args = List.length ctx in
+ let imps,pl,ce =
+ match ctypopt with
+ None ->
+ let subst = evd_comb0 Evd.nf_univ_variables evdref in
+ let ctx = Context.Rel.map (Vars.subst_univs_constr subst) ctx in
+ let env_bl = push_rel_context ctx env in
+ let c, imps2 = interp_constr_evars_impls ~impls env_bl evdref c in
+ let nf,subst = Evarutil.e_nf_evars_and_universes evdref in
+ let body = nf (it_mkLambda_or_LetIn c ctx) in
+ let vars = Universes.universes_of_constr body in
+ let evd = Evd.restrict_universe_context !evdref vars in
+ let pl, uctx = Evd.universe_context ?names:pl evd in
+ imps1@(Impargs.lift_implicits nb_args imps2), pl,
+ definition_entry ~univs:uctx ~poly:p body
+ | Some ctyp ->
+ let ty, impsty = interp_type_evars_impls ~impls env_bl evdref ctyp in
+ let subst = evd_comb0 Evd.nf_univ_variables evdref in
+ let ctx = Context.Rel.map (Vars.subst_univs_constr subst) ctx in
+ let env_bl = push_rel_context ctx env in
+ let c, imps2 = interp_casted_constr_evars_impls ~impls env_bl evdref c ty in
+ let nf, subst = Evarutil.e_nf_evars_and_universes evdref in
+ let body = nf (it_mkLambda_or_LetIn c ctx) in
+ let typ = nf (it_mkProd_or_LetIn ty ctx) in
+ let beq b1 b2 = if b1 then b2 else not b2 in
+ let impl_eq (x,y,z) (x',y',z') = beq x x' && beq y y' && beq z z' in
+ (* Check that all implicit arguments inferable from the term
+ are inferable from the type *)
+ let chk (key,va) =
+ impl_eq (List.assoc_f Pervasives.(=) key impsty) va (* FIXME *)
+ in
+ if not (try List.for_all chk imps2 with Not_found -> false)
+ then warn_implicits_in_term ();
+ let vars = Univ.LSet.union (Universes.universes_of_constr body)
+ (Universes.universes_of_constr typ) in
+ let ctx = Evd.restrict_universe_context !evdref vars in
+ let pl, uctx = Evd.universe_context ?names:pl ctx in
+ imps1@(Impargs.lift_implicits nb_args impsty), pl,
+ definition_entry ~types:typ ~poly:p
+ ~univs:uctx body
+ in
+ red_constant_entry (Context.Rel.length ctx) ce !evdref red_option, !evdref, pl, imps
+
+let check_definition (ce, evd, _, imps) =
+ check_evars_are_solved (Global.env ()) evd (Evd.empty,evd);
+ ce
+
+let warn_local_declaration =
+ CWarnings.create ~name:"local-declaration" ~category:"scope"
+ (fun (id,kind) ->
+ pr_id id ++ strbrk " is declared as a local " ++ str kind)
+
+let get_locality id ~kind = function
+| Discharge ->
+ (** If a Let is defined outside a section, then we consider it as a local definition *)
+ warn_local_declaration (id,kind);
+ true
+| Local -> true
+| Global -> false
+
+let declare_global_definition ident ce local k pl imps =
+ let local = get_locality ident ~kind:"definition" local in
+ let kn = declare_constant ident ~local (DefinitionEntry ce, IsDefinition k) in
+ let gr = ConstRef kn in
+ let () = maybe_declare_manual_implicits false gr imps in
+ let () = Universes.register_universe_binders gr pl in
+ let () = definition_message ident in
+ gr
+
+let declare_definition_hook = ref ignore
+let set_declare_definition_hook = (:=) declare_definition_hook
+let get_declare_definition_hook () = !declare_definition_hook
+
+let warn_definition_not_visible =
+ CWarnings.create ~name:"definition-not-visible" ~category:"implicits"
+ (fun ident ->
+ strbrk "Section definition " ++
+ pr_id ident ++ strbrk " is not visible from current goals")
+
+let declare_definition ident (local, p, k) ce pl imps hook =
+ let fix_exn = Future.fix_exn_of ce.const_entry_body in
+ let () = !declare_definition_hook ce in
+ let r = match local with
+ | Discharge when Lib.sections_are_opened () ->
+ let c = SectionLocalDef ce in
+ let _ = declare_variable ident (Lib.cwd(), c, IsDefinition k) in
+ let () = definition_message ident in
+ let gr = VarRef ident in
+ let () = maybe_declare_manual_implicits false gr imps in
+ let () = if Pfedit.refining () then
+ warn_definition_not_visible ident
+ in
+ gr
+ | Discharge | Local | Global ->
+ declare_global_definition ident ce local k pl imps in
+ Lemmas.call_hook fix_exn hook local r
+
+let _ = Obligations.declare_definition_ref :=
+ (fun i k c imps hook -> declare_definition i k c [] imps hook)
+
+let do_definition ident k pl bl red_option c ctypopt hook =
+ let (ce, evd, pl', imps as def) =
+ interp_definition pl bl (pi2 k) red_option c ctypopt
+ in
+ if Flags.is_program_mode () then
+ let env = Global.env () in
+ let (c,ctx), sideff = Future.force ce.const_entry_body in
+ assert(Safe_typing.empty_private_constants = sideff);
+ assert(Univ.ContextSet.is_empty ctx);
+ let typ = match ce.const_entry_type with
+ | Some t -> t
+ | None -> Retyping.get_type_of env evd c
+ in
+ Obligations.check_evars env evd;
+ let obls, _, c, cty =
+ Obligations.eterm_obligations env ident evd 0 c typ
+ in
+ let ctx = Evd.evar_universe_context evd in
+ let hook = Lemmas.mk_hook (fun l r _ -> Lemmas.call_hook (fun exn -> exn) hook l r) in
+ ignore(Obligations.add_definition
+ ident ~term:c cty ctx ?pl ~implicits:imps ~kind:k ~hook obls)
+ else let ce = check_definition def in
+ ignore(declare_definition ident k ce pl' imps
+ (Lemmas.mk_hook
+ (fun l r -> Lemmas.call_hook (fun exn -> exn) hook l r;r)))
+
+(* 2| Variable/Hypothesis/Parameter/Axiom declarations *)
+
+let declare_assumption is_coe (local,p,kind) (c,ctx) pl imps impl nl (_,ident) =
+match local with
+| Discharge when Lib.sections_are_opened () ->
+ let decl = (Lib.cwd(), SectionLocalAssum ((c,ctx),p,impl), IsAssumption kind) in
+ let _ = declare_variable ident decl in
+ let () = assumption_message ident in
+ let () =
+ if is_verbose () && Pfedit.refining () then
+ Feedback.msg_info (str"Variable" ++ spc () ++ pr_id ident ++
+ strbrk " is not visible from current goals")
+ in
+ let r = VarRef ident in
+ let () = Typeclasses.declare_instance None true r in
+ let () = if is_coe then Class.try_add_new_coercion r ~local:true false in
+ (r,Univ.Instance.empty,true)
+
+| Global | Local | Discharge ->
+ let local = get_locality ident ~kind:"axiom" local in
+ let inl = match nl with
+ | NoInline -> None
+ | DefaultInline -> Some (Flags.get_inline_level())
+ | InlineAt i -> Some i
+ in
+ let ctx = Univ.ContextSet.to_context ctx in
+ let decl = (ParameterEntry (None,p,(c,ctx),inl), IsAssumption kind) in
+ let kn = declare_constant ident ~local decl in
+ let gr = ConstRef kn in
+ let () = maybe_declare_manual_implicits false gr imps in
+ let () = Universes.register_universe_binders gr pl in
+ let () = assumption_message ident in
+ let () = Typeclasses.declare_instance None false gr in
+ let () = if is_coe then Class.try_add_new_coercion gr local p in
+ let inst =
+ if p (* polymorphic *) then Univ.UContext.instance ctx
+ else Univ.Instance.empty
+ in
+ (gr,inst,Lib.is_modtype_strict ())
+
+let interp_assumption evdref env impls bl c =
+ let c = prod_constr_expr c bl in
+ interp_type_evars_impls env evdref ~impls c
+
+let declare_assumptions idl is_coe k (c,ctx) pl imps impl_is_on nl =
+ let refs, status, _ =
+ List.fold_left (fun (refs,status,ctx) id ->
+ let ref',u',status' =
+ declare_assumption is_coe k (c,ctx) pl imps impl_is_on nl id in
+ (ref',u')::refs, status' && status, Univ.ContextSet.empty)
+ ([],true,ctx) idl
+ in
+ List.rev refs, status
+
+let do_assumptions_unbound_univs (_, poly, _ as kind) nl l =
+ let open Context.Named.Declaration in
+ let env = Global.env () in
+ let evdref = ref (Evd.from_env env) in
+ let l =
+ if poly then
+ (* Separate declarations so that A B : Type puts A and B in different levels. *)
+ List.fold_right (fun (is_coe,(idl,c)) acc ->
+ List.fold_right (fun id acc ->
+ (is_coe, ([id], c)) :: acc) idl acc)
+ l []
+ else l
+ in
+ (* We intepret all declarations in the same evar_map, i.e. as a telescope. *)
+ let _,l = List.fold_map (fun (env,ienv) (is_coe,(idl,c)) ->
+ let t,imps = interp_assumption evdref env ienv [] c in
+ let env =
+ push_named_context (List.map (fun (_,id) -> LocalAssum (id,t)) idl) env in
+ let ienv = List.fold_right (fun (_,id) ienv ->
+ let impls = compute_internalization_data env Variable t imps in
+ Id.Map.add id impls ienv) idl ienv in
+ ((env,ienv),((is_coe,idl),t,imps)))
+ (env,empty_internalization_env) l
+ in
+ let evd = solve_remaining_evars all_and_fail_flags env !evdref (Evd.empty,!evdref) in
+ (* The universe constraints come from the whole telescope. *)
+ let evd = Evd.nf_constraints evd in
+ let ctx = Evd.universe_context_set evd in
+ let l = List.map (on_pi2 (nf_evar evd)) l in
+ pi2 (List.fold_left (fun (subst,status,ctx) ((is_coe,idl),t,imps) ->
+ let t = replace_vars subst t in
+ let (refs,status') = declare_assumptions idl is_coe kind (t,ctx) [] imps false nl in
+ let subst' = List.map2
+ (fun (_,id) (c,u) -> (id,Universes.constr_of_global_univ (c,u)))
+ idl refs
+ in
+ (subst'@subst, status' && status,
+ (* The universe constraints are declared with the first declaration only. *)
+ Univ.ContextSet.empty)) ([],true,ctx) l)
+
+let do_assumptions_bound_univs coe kind nl id pl c =
+ let env = Global.env () in
+ let ctx = Evd.make_evar_universe_context env pl in
+ let evdref = ref (Evd.from_ctx ctx) in
+ let ty, impls = interp_type_evars_impls env evdref c in
+ let nf, subst = Evarutil.e_nf_evars_and_universes evdref in
+ let ty = nf ty in
+ let vars = Universes.universes_of_constr ty in
+ let evd = Evd.restrict_universe_context !evdref vars in
+ let pl, uctx = Evd.universe_context ?names:pl evd in
+ let uctx = Univ.ContextSet.of_context uctx in
+ let (_, _, st) = declare_assumption coe kind (ty, uctx) pl impls false nl id in
+ st
+
+let do_assumptions kind nl l = match l with
+| [coe, ([id, Some pl], c)] ->
+ let () = match kind with
+ | (Discharge, _, _) when Lib.sections_are_opened () ->
+ let loc = fst id in
+ let msg = Pp.str "Section variables cannot be polymorphic." in
+ user_err ~loc msg
+ | _ -> ()
+ in
+ do_assumptions_bound_univs coe kind nl id (Some pl) c
+| _ ->
+ let map (coe, (idl, c)) =
+ let map (id, univs) = match univs with
+ | None -> id
+ | Some _ ->
+ let loc = fst id in
+ let msg =
+ Pp.str "Assumptions with bound universes can only be defined one at a time." in
+ user_err ~loc msg
+ in
+ (coe, (List.map map idl, c))
+ in
+ let l = List.map map l in
+ do_assumptions_unbound_univs kind nl l
+
+(* 3a| Elimination schemes for mutual inductive definitions *)
+
+(* 3b| Mutual inductive definitions *)
+
+let push_types env idl tl =
+ List.fold_left2 (fun env id t -> Environ.push_rel (LocalAssum (Name id,t)) env)
+ env idl tl
+
+type structured_one_inductive_expr = {
+ ind_name : Id.t;
+ ind_univs : lident list option;
+ ind_arity : constr_expr;
+ ind_lc : (Id.t * constr_expr) list
+}
+
+type structured_inductive_expr =
+ local_binder list * structured_one_inductive_expr list
+
+let minductive_message warn = function
+ | [] -> error "No inductive definition."
+ | [x] -> (pr_id x ++ str " is defined" ++
+ if warn then str " as a non-primitive record" else mt())
+ | l -> hov 0 (prlist_with_sep pr_comma pr_id l ++
+ spc () ++ str "are defined")
+
+let check_all_names_different indl =
+ let ind_names = List.map (fun ind -> ind.ind_name) indl in
+ let cstr_names = List.map_append (fun ind -> List.map fst ind.ind_lc) indl in
+ let l = List.duplicates Id.equal ind_names in
+ let () = match l with
+ | [] -> ()
+ | t :: _ -> raise (InductiveError (SameNamesTypes t))
+ in
+ let l = List.duplicates Id.equal cstr_names in
+ let () = match l with
+ | [] -> ()
+ | c :: _ -> raise (InductiveError (SameNamesConstructors (List.hd l)))
+ in
+ let l = List.intersect Id.equal ind_names cstr_names in
+ match l with
+ | [] -> ()
+ | _ -> raise (InductiveError (SameNamesOverlap l))
+
+let mk_mltype_data evdref env assums arity indname =
+ let is_ml_type = is_sort env !evdref arity in
+ (is_ml_type,indname,assums)
+
+let prepare_param = function
+ | LocalAssum (na,t) -> out_name na, LocalAssumEntry t
+ | LocalDef (na,b,_) -> out_name na, LocalDefEntry b
+
+(** Make the arity conclusion flexible to avoid generating an upper bound universe now,
+ only if the universe does not appear anywhere else.
+ This is really a hack to stay compatible with the semantics of template polymorphic
+ inductives which are recognized when a "Type" appears at the end of the conlusion in
+ the source syntax. *)
+
+let rec check_anonymous_type ind =
+ let open Glob_term in
+ match ind with
+ | GSort (_, GType []) -> true
+ | GProd (_, _, _, _, e)
+ | GLetIn (_, _, _, e)
+ | GLambda (_, _, _, _, e)
+ | GApp (_, e, _)
+ | GCast (_, e, _) -> check_anonymous_type e
+ | _ -> false
+
+let make_conclusion_flexible evdref ty poly =
+ if poly && isArity ty then
+ let _, concl = destArity ty in
+ match concl with
+ | Type u ->
+ (match Univ.universe_level u with
+ | Some u ->
+ evdref := Evd.make_flexible_variable !evdref true u
+ | None -> ())
+ | _ -> ()
+ else ()
+
+let is_impredicative env u =
+ u = Prop Null || (is_impredicative_set env && u = Prop Pos)
+
+let interp_ind_arity env evdref ind =
+ let c = intern_gen IsType env ind.ind_arity in
+ let imps = Implicit_quantifiers.implicits_of_glob_constr ~with_products:true c in
+ let t, impls = understand_tcc_evars env evdref ~expected_type:IsType c, imps in
+ let pseudo_poly = check_anonymous_type c in
+ let () = if not (Reduction.is_arity env t) then
+ user_err ~loc:(constr_loc ind.ind_arity) (str "Not an arity")
+ in
+ t, pseudo_poly, impls
+
+let interp_cstrs evdref env impls mldata arity ind =
+ let cnames,ctyps = List.split ind.ind_lc in
+ (* Complete conclusions of constructor types if given in ML-style syntax *)
+ let ctyps' = List.map2 (complete_conclusion mldata) cnames ctyps in
+ (* Interpret the constructor types *)
+ let ctyps'', cimpls = List.split (List.map (interp_type_evars_impls evdref env ~impls) ctyps') in
+ (cnames, ctyps'', cimpls)
+
+let sign_level env evd sign =
+ fst (List.fold_right
+ (fun d (lev,env) ->
+ match d with
+ | LocalDef _ -> lev, push_rel d env
+ | LocalAssum _ ->
+ let s = destSort (Reduction.whd_all env
+ (nf_evar evd (Retyping.get_type_of env evd (RelDecl.get_type d))))
+ in
+ let u = univ_of_sort s in
+ (Univ.sup u lev, push_rel d env))
+ sign (Univ.type0m_univ,env))
+
+let sup_list min = List.fold_left Univ.sup min
+
+let extract_level env evd min tys =
+ let sorts = List.map (fun ty ->
+ let ctx, concl = Reduction.dest_prod_assum env ty in
+ sign_level env evd (LocalAssum (Anonymous, concl) :: ctx)) tys
+ in sup_list min sorts
+
+let is_flexible_sort evd u =
+ match Univ.Universe.level u with
+ | Some l -> Evd.is_flexible_level evd l
+ | None -> false
+
+let inductive_levels env evdref poly arities inds =
+ let destarities = List.map (fun x -> x, Reduction.dest_arity env x) arities in
+ let levels = List.map (fun (x,(ctx,a)) ->
+ if a = Prop Null then None
+ else Some (univ_of_sort a)) destarities
+ in
+ let cstrs_levels, min_levels, sizes =
+ CList.split3
+ (List.map2 (fun (_,tys,_) (arity,(ctx,du)) ->
+ let len = List.length tys in
+ let minlev = Sorts.univ_of_sort du in
+ let minlev =
+ if len > 1 && not (is_impredicative env du) then
+ Univ.sup minlev Univ.type0_univ
+ else minlev
+ in
+ let minlev =
+ (** Indices contribute. *)
+ if Indtypes.is_indices_matter () && List.length ctx > 0 then (
+ let ilev = sign_level env !evdref ctx in
+ Univ.sup ilev minlev)
+ else minlev
+ in
+ let clev = extract_level env !evdref minlev tys in
+ (clev, minlev, len)) inds destarities)
+ in
+ (* Take the transitive closure of the system of constructors *)
+ (* level constraints and remove the recursive dependencies *)
+ let levels' = Universes.solve_constraints_system (Array.of_list levels)
+ (Array.of_list cstrs_levels) (Array.of_list min_levels)
+ in
+ let evd, arities =
+ CList.fold_left3 (fun (evd, arities) cu (arity,(ctx,du)) len ->
+ if is_impredicative env du then
+ (** Any product is allowed here. *)
+ evd, arity :: arities
+ else (** If in a predicative sort, or asked to infer the type,
+ we take the max of:
+ - indices (if in indices-matter mode)
+ - constructors
+ - Type(1) if there is more than 1 constructor
+ *)
+ (** Constructors contribute. *)
+ let evd =
+ if Sorts.is_set du then
+ if not (Evd.check_leq evd cu Univ.type0_univ) then
+ raise (Indtypes.InductiveError Indtypes.LargeNonPropInductiveNotInType)
+ else evd
+ else evd
+ (* Evd.set_leq_sort env evd (Type cu) du *)
+ in
+ let evd =
+ if len >= 2 && Univ.is_type0m_univ cu then
+ (** "Polymorphic" type constraint and more than one constructor,
+ should not land in Prop. Add constraint only if it would
+ land in Prop directly (no informative arguments as well). *)
+ Evd.set_leq_sort env evd (Prop Pos) du
+ else evd
+ in
+ let duu = Sorts.univ_of_sort du in
+ let evd =
+ if not (Univ.is_small_univ duu) && Univ.Universe.equal cu duu then
+ if is_flexible_sort evd duu && not (Evd.check_leq evd Univ.type0_univ duu) then
+ Evd.set_eq_sort env evd (Prop Null) du
+ else evd
+ else Evd.set_eq_sort env evd (Type cu) du
+ in
+ (evd, arity :: arities))
+ (!evdref,[]) (Array.to_list levels') destarities sizes
+ in evdref := evd; List.rev arities
+
+let check_named (loc, na) = match na with
+| Name _ -> ()
+| Anonymous ->
+ let msg = str "Parameters must be named." in
+ user_err ~loc msg
+
+
+let check_param = function
+| LocalRawDef (na, _) -> check_named na
+| LocalRawAssum (nas, Default _, _) -> List.iter check_named nas
+| LocalRawAssum (nas, Generalized _, _) -> ()
+| LocalPattern _ -> assert false
+
+let interp_mutual_inductive (paramsl,indl) notations poly prv finite =
+ check_all_names_different indl;
+ List.iter check_param paramsl;
+ let env0 = Global.env() in
+ let pl = (List.hd indl).ind_univs in
+ let ctx = Evd.make_evar_universe_context env0 pl in
+ let evdref = ref Evd.(from_ctx ctx) in
+ let _, ((env_params, ctx_params), userimpls) =
+ interp_context_evars env0 evdref paramsl
+ in
+ let indnames = List.map (fun ind -> ind.ind_name) indl in
+
+ (* Names of parameters as arguments of the inductive type (defs removed) *)
+ let assums = List.filter is_local_assum ctx_params in
+ let params = List.map (RelDecl.get_name %> out_name) assums in
+
+ (* Interpret the arities *)
+ let arities = List.map (interp_ind_arity env_params evdref) indl in
+
+ let fullarities = List.map (fun (c, _, _) -> it_mkProd_or_LetIn c ctx_params) arities in
+ let env_ar = push_types env0 indnames fullarities in
+ let env_ar_params = push_rel_context ctx_params env_ar in
+
+ (* Compute interpretation metadatas *)
+ let indimpls = List.map (fun (_, _, impls) -> userimpls @
+ lift_implicits (Context.Rel.nhyps ctx_params) impls) arities in
+ let arities = List.map pi1 arities and aritypoly = List.map pi2 arities in
+ let impls = compute_internalization_env env0 (Inductive params) indnames fullarities indimpls in
+ let mldatas = List.map2 (mk_mltype_data evdref env_params params) arities indnames in
+
+ let constructors =
+ Metasyntax.with_syntax_protection (fun () ->
+ (* Temporary declaration of notations and scopes *)
+ List.iter (Metasyntax.set_notation_for_interpretation impls) notations;
+ (* Interpret the constructor types *)
+ List.map3 (interp_cstrs env_ar_params evdref impls) mldatas arities indl)
+ () in
+
+ (* Try further to solve evars, and instantiate them *)
+ let sigma = solve_remaining_evars all_and_fail_flags env_params !evdref (Evd.empty,!evdref) in
+ evdref := sigma;
+ (* Compute renewed arities *)
+ let nf,_ = e_nf_evars_and_universes evdref in
+ let arities = List.map nf arities in
+ let constructors = List.map (fun (idl,cl,impsl) -> (idl,List.map nf cl,impsl)) constructors in
+ let _ = List.iter2 (fun ty poly -> make_conclusion_flexible evdref ty poly) arities aritypoly in
+ let arities = inductive_levels env_ar_params evdref poly arities constructors in
+ let nf',_ = e_nf_evars_and_universes evdref in
+ let nf x = nf' (nf x) in
+ let arities = List.map nf' arities in
+ let constructors = List.map (fun (idl,cl,impsl) -> (idl,List.map nf' cl,impsl)) constructors in
+ let ctx_params = Context.Rel.map nf ctx_params in
+ let evd = !evdref in
+ let pl, uctx = Evd.universe_context ?names:pl evd in
+ List.iter (check_evars env_params Evd.empty evd) arities;
+ Context.Rel.iter (check_evars env0 Evd.empty evd) ctx_params;
+ List.iter (fun (_,ctyps,_) ->
+ List.iter (check_evars env_ar_params Evd.empty evd) ctyps)
+ constructors;
+
+ (* Build the inductive entries *)
+ let entries = List.map4 (fun ind arity template (cnames,ctypes,cimpls) -> {
+ mind_entry_typename = ind.ind_name;
+ mind_entry_arity = arity;
+ mind_entry_template = template;
+ mind_entry_consnames = cnames;
+ mind_entry_lc = ctypes
+ }) indl arities aritypoly constructors in
+ let impls =
+ let len = Context.Rel.nhyps ctx_params in
+ List.map2 (fun indimpls (_,_,cimpls) ->
+ indimpls, List.map (fun impls ->
+ userimpls @ (lift_implicits len impls)) cimpls) indimpls constructors
+ in
+ (* Build the mutual inductive entry *)
+ { mind_entry_params = List.map prepare_param ctx_params;
+ mind_entry_record = None;
+ mind_entry_finite = finite;
+ mind_entry_inds = entries;
+ mind_entry_polymorphic = poly;
+ mind_entry_private = if prv then Some false else None;
+ mind_entry_universes = uctx;
+ },
+ pl, impls
+
+(* Very syntactical equality *)
+let eq_local_binders bl1 bl2 =
+ List.equal local_binder_eq bl1 bl2
+
+let extract_coercions indl =
+ let mkqid (_,((_,id),_)) = qualid_of_ident id in
+ let extract lc = List.filter (fun (iscoe,_) -> iscoe) lc in
+ List.map mkqid (List.flatten(List.map (fun (_,_,_,lc) -> extract lc) indl))
+
+let extract_params indl =
+ let paramsl = List.map (fun (_,params,_,_) -> params) indl in
+ match paramsl with
+ | [] -> anomaly (Pp.str "empty list of inductive types")
+ | params::paramsl ->
+ if not (List.for_all (eq_local_binders params) paramsl) then error
+ "Parameters should be syntactically the same for each inductive type.";
+ params
+
+let extract_inductive indl =
+ List.map (fun (((_,indname),pl),_,ar,lc) -> {
+ ind_name = indname; ind_univs = pl;
+ ind_arity = Option.cata (fun x -> x) (CSort (Loc.ghost,GType [])) ar;
+ ind_lc = List.map (fun (_,((_,id),t)) -> (id,t)) lc
+ }) indl
+
+let extract_mutual_inductive_declaration_components indl =
+ let indl,ntnl = List.split indl in
+ let params = extract_params indl in
+ let coes = extract_coercions indl in
+ let indl = extract_inductive indl in
+ (params,indl), coes, List.flatten ntnl
+
+let is_recursive mie =
+ let rec is_recursive_constructor lift typ =
+ match Term.kind_of_term typ with
+ | Prod (_,arg,rest) ->
+ Termops.dependent (mkRel lift) arg ||
+ is_recursive_constructor (lift+1) rest
+ | LetIn (na,b,t,rest) -> is_recursive_constructor (lift+1) rest
+ | _ -> false
+ in
+ match mie.mind_entry_inds with
+ | [ind] ->
+ let nparams = List.length mie.mind_entry_params in
+ List.exists (fun t -> is_recursive_constructor (nparams+1) t) ind.mind_entry_lc
+ | _ -> false
+
+let declare_mutual_inductive_with_eliminations mie pl impls =
+ (* spiwack: raises an error if the structure is supposed to be non-recursive,
+ but isn't *)
+ begin match mie.mind_entry_finite with
+ | BiFinite when is_recursive mie ->
+ if Option.has_some mie.mind_entry_record then
+ error "Records declared with the keywords Record or Structure cannot be recursive. You can, however, define recursive records using the Inductive or CoInductive command."
+ else
+ error ("Types declared with the keyword Variant cannot be recursive. Recursive types are defined with the Inductive and CoInductive command.")
+ | _ -> ()
+ end;
+ let names = List.map (fun e -> e.mind_entry_typename) mie.mind_entry_inds in
+ let (_, kn), prim = declare_mind mie in
+ let mind = Global.mind_of_delta_kn kn in
+ List.iteri (fun i (indimpls, constrimpls) ->
+ let ind = (mind,i) in
+ let gr = IndRef ind in
+ maybe_declare_manual_implicits false gr indimpls;
+ Universes.register_universe_binders gr pl;
+ List.iteri
+ (fun j impls ->
+ maybe_declare_manual_implicits false
+ (ConstructRef (ind, succ j)) impls)
+ constrimpls)
+ impls;
+ let warn_prim = match mie.mind_entry_record with Some (Some _) -> not prim | _ -> false in
+ if_verbose Feedback.msg_info (minductive_message warn_prim names);
+ if mie.mind_entry_private == None
+ then declare_default_schemes mind;
+ mind
+
+type one_inductive_impls =
+ Impargs.manual_explicitation list (* for inds *)*
+ Impargs.manual_explicitation list list (* for constrs *)
+
+let do_mutual_inductive indl poly prv finite =
+ let indl,coes,ntns = extract_mutual_inductive_declaration_components indl in
+ (* Interpret the types *)
+ let mie,pl,impls = interp_mutual_inductive indl ntns poly prv finite in
+ (* Declare the mutual inductive block with its associated schemes *)
+ ignore (declare_mutual_inductive_with_eliminations mie pl impls);
+ (* Declare the possible notations of inductive types *)
+ List.iter Metasyntax.add_notation_interpretation ntns;
+ (* Declare the coercions *)
+ List.iter (fun qid -> Class.try_add_new_coercion (locate qid) false poly) coes;
+ (* If positivity is assumed declares itself as unsafe. *)
+ if Environ.deactivated_guard (Global.env ()) then Feedback.feedback Feedback.AddedAxiom else ()
+
+(* 3c| Fixpoints and co-fixpoints *)
+
+(* An (unoptimized) function that maps preorders to partial orders...
+
+ Input: a list of associations (x,[y1;...;yn]), all yi distincts
+ and different of x, meaning x<=y1, ..., x<=yn
+
+ Output: a list of associations (x,Inr [y1;...;yn]), collecting all
+ distincts yi greater than x, _or_, (x, Inl y) meaning that
+ x is in the same class as y (in which case, x occurs
+ nowhere else in the association map)
+
+ partial_order : ('a * 'a list) list -> ('a * ('a,'a list) union) list
+*)
+
+let rec partial_order cmp = function
+ | [] -> []
+ | (x,xge)::rest ->
+ let rec browse res xge' = function
+ | [] ->
+ let res = List.map (function
+ | (z, Inr zge) when List.mem_f cmp x zge ->
+ (z, Inr (List.union cmp zge xge'))
+ | r -> r) res in
+ (x,Inr xge')::res
+ | y::xge ->
+ let rec link y =
+ try match List.assoc_f cmp y res with
+ | Inl z -> link z
+ | Inr yge ->
+ if List.mem_f cmp x yge then
+ let res = List.remove_assoc_f cmp y res in
+ let res = List.map (function
+ | (z, Inl t) ->
+ if cmp t y then (z, Inl x) else (z, Inl t)
+ | (z, Inr zge) ->
+ if List.mem_f cmp y zge then
+ (z, Inr (List.add_set cmp x (List.remove cmp y zge)))
+ else
+ (z, Inr zge)) res in
+ browse ((y,Inl x)::res) xge' (List.union cmp xge (List.remove cmp x yge))
+ else
+ browse res (List.add_set cmp y (List.union cmp xge' yge)) xge
+ with Not_found -> browse res (List.add_set cmp y xge') xge
+ in link y
+ in browse (partial_order cmp rest) [] xge
+
+let non_full_mutual_message x xge y yge isfix rest =
+ let reason =
+ if Id.List.mem x yge then
+ pr_id y ++ str " depends on " ++ pr_id x ++ strbrk " but not conversely"
+ else if Id.List.mem y xge then
+ pr_id x ++ str " depends on " ++ pr_id y ++ strbrk " but not conversely"
+ else
+ pr_id y ++ str " and " ++ pr_id x ++ strbrk " are not mutually dependent" in
+ let e = if List.is_empty rest then reason else strbrk "e.g., " ++ reason in
+ let k = if isfix then "fixpoint" else "cofixpoint" in
+ let w =
+ if isfix
+ then strbrk "Well-foundedness check may fail unexpectedly." ++ fnl()
+ else mt () in
+ strbrk "Not a fully mutually defined " ++ str k ++ fnl () ++
+ str "(" ++ e ++ str ")." ++ fnl () ++ w
+
+let warn_non_full_mutual =
+ CWarnings.create ~name:"non-full-mutual" ~category:"fixpoints"
+ (fun (x,xge,y,yge,isfix,rest) ->
+ non_full_mutual_message x xge y yge isfix rest)
+
+let check_mutuality env isfix fixl =
+ let names = List.map fst fixl in
+ let preorder =
+ List.map (fun (id,def) ->
+ (id, List.filter (fun id' -> not (Id.equal id id') && occur_var env id' def) names))
+ fixl in
+ let po = partial_order Id.equal preorder in
+ match List.filter (function (_,Inr _) -> true | _ -> false) po with
+ | (x,Inr xge)::(y,Inr yge)::rest ->
+ warn_non_full_mutual (x,xge,y,yge,isfix,rest)
+ | _ -> ()
+
+type structured_fixpoint_expr = {
+ fix_name : Id.t;
+ fix_univs : lident list option;
+ fix_annot : Id.t Loc.located option;
+ fix_binders : local_binder list;
+ fix_body : constr_expr option;
+ fix_type : constr_expr
+}
+
+let interp_fix_context env evdref isfix fix =
+ let before, after = if isfix then split_at_annot fix.fix_binders fix.fix_annot else [], fix.fix_binders in
+ let impl_env, ((env', ctx), imps) = interp_context_evars env evdref before in
+ let impl_env', ((env'', ctx'), imps') = interp_context_evars ~impl_env ~shift:(List.length before) env' evdref after in
+ let annot = Option.map (fun _ -> List.length (assums_of_rel_context ctx)) fix.fix_annot in
+ ((env'', ctx' @ ctx), (impl_env',imps @ imps'), annot)
+
+let interp_fix_ccl evdref impls (env,_) fix =
+ interp_type_evars_impls ~impls env evdref fix.fix_type
+
+let interp_fix_body env_rec evdref impls (_,ctx) fix ccl =
+ Option.map (fun body ->
+ let env = push_rel_context ctx env_rec in
+ let body = interp_casted_constr_evars env evdref ~impls body ccl in
+ it_mkLambda_or_LetIn body ctx) fix.fix_body
+
+let build_fix_type (_,ctx) ccl = it_mkProd_or_LetIn ccl ctx
+
+let declare_fix ?(opaque = false) (_,poly,_ as kind) pl ctx f ((def,_),eff) t imps =
+ let ce = definition_entry ~opaque ~types:t ~poly ~univs:ctx ~eff def in
+ declare_definition f kind ce pl imps (Lemmas.mk_hook (fun _ r -> r))
+
+let _ = Obligations.declare_fix_ref :=
+ (fun ?opaque k ctx f d t imps -> declare_fix ?opaque k [] ctx f d t imps)
+
+let prepare_recursive_declaration fixnames fixtypes fixdefs =
+ let defs = List.map (subst_vars (List.rev fixnames)) fixdefs in
+ let names = List.map (fun id -> Name id) fixnames in
+ (Array.of_list names, Array.of_list fixtypes, Array.of_list defs)
+
+(* Jump over let-bindings. *)
+
+let compute_possible_guardness_evidences (ids,_,na) =
+ match na with
+ | Some i -> [i]
+ | None ->
+ (* If recursive argument was not given by user, we try all args.
+ An earlier approach was to look only for inductive arguments,
+ but doing it properly involves delta-reduction, and it finally
+ doesn't seem to worth the effort (except for huge mutual
+ fixpoints ?) *)
+ List.interval 0 (List.length ids - 1)
+
+type recursive_preentry =
+ Id.t list * constr option list * types list
+
+(* Wellfounded definition *)
+
+open Coqlib
+
+let contrib_name = "Program"
+let subtac_dir = [contrib_name]
+let fixsub_module = subtac_dir @ ["Wf"]
+let tactics_module = subtac_dir @ ["Tactics"]
+
+let init_reference dir s () = Coqlib.gen_reference "Command" dir s
+let init_constant dir s () = Coqlib.gen_constant "Command" dir s
+let make_ref l s = init_reference l s
+let fix_proto = init_constant tactics_module "fix_proto"
+let fix_sub_ref = make_ref fixsub_module "Fix_sub"
+let measure_on_R_ref = make_ref fixsub_module "MR"
+let well_founded = init_constant ["Init"; "Wf"] "well_founded"
+let mkSubset name typ prop =
+ mkApp (Universes.constr_of_global (delayed_force build_sigma).typ,
+ [| typ; mkLambda (name, typ, prop) |])
+let sigT = Lazy.from_fun build_sigma_type
+
+let make_qref s = Qualid (Loc.ghost, qualid_of_string s)
+let lt_ref = make_qref "Init.Peano.lt"
+
+let rec telescope = function
+ | [] -> assert false
+ | [LocalAssum (n, t)] -> t, [LocalDef (n, mkRel 1, t)], mkRel 1
+ | LocalAssum (n, t) :: tl ->
+ let ty, tys, (k, constr) =
+ List.fold_left
+ (fun (ty, tys, (k, constr)) decl ->
+ let t = RelDecl.get_type decl in
+ let pred = mkLambda (RelDecl.get_name decl, t, ty) in
+ let ty = Universes.constr_of_global (Lazy.force sigT).typ in
+ let intro = Universes.constr_of_global (Lazy.force sigT).intro in
+ let sigty = mkApp (ty, [|t; pred|]) in
+ let intro = mkApp (intro, [|lift k t; lift k pred; mkRel k; constr|]) in
+ (sigty, pred :: tys, (succ k, intro)))
+ (t, [], (2, mkRel 1)) tl
+ in
+ let (last, subst) = List.fold_right2
+ (fun pred decl (prev, subst) ->
+ let t = RelDecl.get_type decl in
+ let p1 = Universes.constr_of_global (Lazy.force sigT).proj1 in
+ let p2 = Universes.constr_of_global (Lazy.force sigT).proj2 in
+ let proj1 = applistc p1 [t; pred; prev] in
+ let proj2 = applistc p2 [t; pred; prev] in
+ (lift 1 proj2, LocalDef (get_name decl, proj1, t) :: subst))
+ (List.rev tys) tl (mkRel 1, [])
+ in ty, (LocalDef (n, last, t) :: subst), constr
+
+ | LocalDef (n, b, t) :: tl -> let ty, subst, term = telescope tl in
+ ty, (LocalDef (n, b, t) :: subst), lift 1 term
+
+let nf_evar_context sigma ctx =
+ List.map (map_constr (Evarutil.nf_evar sigma)) ctx
+
+let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
+ Coqlib.check_required_library ["Coq";"Program";"Wf"];
+ let env = Global.env() in
+ let ctx = Evd.make_evar_universe_context env pl in
+ let evdref = ref (Evd.from_ctx ctx) in
+ let _, ((env', binders_rel), impls) = interp_context_evars env evdref bl in
+ let len = List.length binders_rel in
+ let top_env = push_rel_context binders_rel env in
+ let top_arity = interp_type_evars top_env evdref arityc in
+ let full_arity = it_mkProd_or_LetIn top_arity binders_rel in
+ let argtyp, letbinders, make = telescope binders_rel in
+ let argname = Id.of_string "recarg" in
+ let arg = LocalAssum (Name argname, argtyp) in
+ let binders = letbinders @ [arg] in
+ let binders_env = push_rel_context binders_rel env in
+ let rel, _ = interp_constr_evars_impls env evdref r in
+ let relty = Typing.unsafe_type_of env !evdref rel in
+ let relargty =
+ let error () =
+ user_err ~loc:(constr_loc r)
+ ~hdr:"Command.build_wellfounded"
+ (Printer.pr_constr_env env !evdref rel ++ str " is not an homogeneous binary relation.")
+ in
+ try
+ let ctx, ar = Reductionops.splay_prod_n env !evdref 2 relty in
+ match ctx, kind_of_term ar with
+ | [LocalAssum (_,t); LocalAssum (_,u)], Sort (Prop Null)
+ when Reductionops.is_conv env !evdref t u -> t
+ | _, _ -> error ()
+ with e when CErrors.noncritical e -> error ()
+ in
+ let measure = interp_casted_constr_evars binders_env evdref measure relargty in
+ let wf_rel, wf_rel_fun, measure_fn =
+ let measure_body, measure =
+ it_mkLambda_or_LetIn measure letbinders,
+ it_mkLambda_or_LetIn measure binders
+ in
+ let comb = Universes.constr_of_global (delayed_force measure_on_R_ref) in
+ let wf_rel = mkApp (comb, [| argtyp; relargty; rel; measure |]) in
+ let wf_rel_fun x y =
+ mkApp (rel, [| subst1 x measure_body;
+ subst1 y measure_body |])
+ in wf_rel, wf_rel_fun, measure
+ in
+ let wf_proof = mkApp (delayed_force well_founded, [| argtyp ; wf_rel |]) in
+ let argid' = Id.of_string (Id.to_string argname ^ "'") in
+ let wfarg len = LocalAssum (Name argid',
+ mkSubset (Name argid') argtyp
+ (wf_rel_fun (mkRel 1) (mkRel (len + 1))))
+ in
+ let intern_bl = wfarg 1 :: [arg] in
+ let _intern_env = push_rel_context intern_bl env in
+ let proj = (*FIXME*)Universes.constr_of_global (delayed_force build_sigma).Coqlib.proj1 in
+ let wfargpred = mkLambda (Name argid', argtyp, wf_rel_fun (mkRel 1) (mkRel 3)) in
+ let projection = (* in wfarg :: arg :: before *)
+ mkApp (proj, [| argtyp ; wfargpred ; mkRel 1 |])
+ in
+ let top_arity_let = it_mkLambda_or_LetIn top_arity letbinders in
+ let intern_arity = substl [projection] top_arity_let in
+ (* substitute the projection of wfarg for something,
+ now intern_arity is in wfarg :: arg *)
+ let intern_fun_arity_prod = it_mkProd_or_LetIn intern_arity [wfarg 1] in
+ let intern_fun_binder = LocalAssum (Name (add_suffix recname "'"), intern_fun_arity_prod) in
+ let curry_fun =
+ let wfpred = mkLambda (Name argid', argtyp, wf_rel_fun (mkRel 1) (mkRel (2 * len + 4))) in
+ let intro = (*FIXME*)Universes.constr_of_global (delayed_force build_sigma).Coqlib.intro in
+ let arg = mkApp (intro, [| argtyp; wfpred; lift 1 make; mkRel 1 |]) in
+ let app = mkApp (mkRel (2 * len + 2 (* recproof + orig binders + current binders *)), [| arg |]) in
+ let rcurry = mkApp (rel, [| measure; lift len measure |]) in
+ let lam = LocalAssum (Name (Id.of_string "recproof"), rcurry) in
+ let body = it_mkLambda_or_LetIn app (lam :: binders_rel) in
+ let ty = it_mkProd_or_LetIn (lift 1 top_arity) (lam :: binders_rel) in
+ LocalDef (Name recname, body, ty)
+ in
+ let fun_bl = intern_fun_binder :: [arg] in
+ let lift_lets = Termops.lift_rel_context 1 letbinders in
+ let intern_body =
+ let ctx = LocalAssum (Name recname, get_type curry_fun) :: binders_rel in
+ let (r, l, impls, scopes) =
+ Constrintern.compute_internalization_data env
+ Constrintern.Recursive full_arity impls
+ in
+ let newimpls = Id.Map.singleton recname
+ (r, l, impls @ [(Some (Id.of_string "recproof", Impargs.Manual, (true, false)))],
+ scopes @ [None]) in
+ interp_casted_constr_evars (push_rel_context ctx env) evdref
+ ~impls:newimpls body (lift 1 top_arity)
+ in
+ let intern_body_lam = it_mkLambda_or_LetIn intern_body (curry_fun :: lift_lets @ fun_bl) in
+ let prop = mkLambda (Name argname, argtyp, top_arity_let) in
+ let def =
+ mkApp (Universes.constr_of_global (delayed_force fix_sub_ref),
+ [| argtyp ; wf_rel ;
+ Evarutil.e_new_evar env evdref
+ ~src:(Loc.ghost, Evar_kinds.QuestionMark (Evar_kinds.Define false)) wf_proof;
+ prop |])
+ in
+ let def = Typing.e_solve_evars env evdref def in
+ let _ = evdref := Evarutil.nf_evar_map !evdref in
+ let def = mkApp (def, [|intern_body_lam|]) in
+ let binders_rel = nf_evar_context !evdref binders_rel in
+ let binders = nf_evar_context !evdref binders in
+ let top_arity = Evarutil.nf_evar !evdref top_arity in
+ let hook, recname, typ =
+ if List.length binders_rel > 1 then
+ let name = add_suffix recname "_func" in
+ let hook l gr _ =
+ let body = it_mkLambda_or_LetIn (mkApp (Universes.constr_of_global gr, [|make|])) binders_rel in
+ let ty = it_mkProd_or_LetIn top_arity binders_rel in
+ let pl, univs = Evd.universe_context ?names:pl !evdref in
+ (*FIXME poly? *)
+ let ce = definition_entry ~poly ~types:ty ~univs (Evarutil.nf_evar !evdref body) in
+ (** FIXME: include locality *)
+ let c = Declare.declare_constant recname (DefinitionEntry ce, IsDefinition Definition) in
+ let gr = ConstRef c in
+ if Impargs.is_implicit_args () || not (List.is_empty impls) then
+ Impargs.declare_manual_implicits false gr [impls]
+ in
+ let typ = it_mkProd_or_LetIn top_arity binders in
+ hook, name, typ
+ else
+ let typ = it_mkProd_or_LetIn top_arity binders_rel in
+ let hook l gr _ =
+ if Impargs.is_implicit_args () || not (List.is_empty impls) then
+ Impargs.declare_manual_implicits false gr [impls]
+ in hook, recname, typ
+ in
+ let hook = Lemmas.mk_hook hook in
+ let fullcoqc = Evarutil.nf_evar !evdref def in
+ let fullctyp = Evarutil.nf_evar !evdref typ in
+ Obligations.check_evars env !evdref;
+ let evars, _, evars_def, evars_typ =
+ Obligations.eterm_obligations env recname !evdref 0 fullcoqc fullctyp
+ in
+ let ctx = Evd.evar_universe_context !evdref in
+ ignore(Obligations.add_definition recname ~term:evars_def ?pl
+ evars_typ ctx evars ~hook)
+
+let interp_recursive isfix fixl notations =
+ let open Context.Named.Declaration in
+ let env = Global.env() in
+ let fixnames = List.map (fun fix -> fix.fix_name) fixl in
+
+ (* Interp arities allowing for unresolved types *)
+ let all_universes =
+ List.fold_right (fun sfe acc ->
+ match sfe.fix_univs , acc with
+ | None , acc -> acc
+ | x , None -> x
+ | Some ls , Some us ->
+ if not (CList.for_all2eq (fun x y -> Id.equal (snd x) (snd y)) ls us) then
+ error "(co)-recursive definitions should all have the same universe binders";
+ Some us) fixl None in
+ let ctx = Evd.make_evar_universe_context env all_universes in
+ let evdref = ref (Evd.from_ctx ctx) in
+ let fixctxs, fiximppairs, fixannots =
+ List.split3 (List.map (interp_fix_context env evdref isfix) fixl) in
+ let fixctximpenvs, fixctximps = List.split fiximppairs in
+ let fixccls,fixcclimps = List.split (List.map3 (interp_fix_ccl evdref) fixctximpenvs fixctxs fixl) in
+ let fixtypes = List.map2 build_fix_type fixctxs fixccls in
+ let fixtypes = List.map (nf_evar !evdref) fixtypes in
+ let fiximps = List.map3
+ (fun ctximps cclimps (_,ctx) -> ctximps@(Impargs.lift_implicits (List.length ctx) cclimps))
+ fixctximps fixcclimps fixctxs in
+ let rec_sign =
+ List.fold_left2
+ (fun env' id t ->
+ if Flags.is_program_mode () then
+ let sort = Evarutil.evd_comb1 (Typing.type_of ~refresh:true env) evdref t in
+ let fixprot =
+ try
+ let app = mkApp (delayed_force fix_proto, [|sort; t|]) in
+ Typing.e_solve_evars env evdref app
+ with e when CErrors.noncritical e -> t
+ in
+ LocalAssum (id,fixprot) :: env'
+ else LocalAssum (id,t) :: env')
+ [] fixnames fixtypes
+ in
+ let env_rec = push_named_context rec_sign env in
+
+ (* Get interpretation metadatas *)
+ let impls = compute_internalization_env env Recursive fixnames fixtypes fiximps in
+
+ (* Interp bodies with rollback because temp use of notations/implicit *)
+ let fixdefs =
+ Metasyntax.with_syntax_protection (fun () ->
+ List.iter (Metasyntax.set_notation_for_interpretation impls) notations;
+ List.map4
+ (fun fixctximpenv -> interp_fix_body env_rec evdref (Id.Map.fold Id.Map.add fixctximpenv impls))
+ fixctximpenvs fixctxs fixl fixccls)
+ () in
+
+ (* Instantiate evars and check all are resolved *)
+ let evd = solve_unif_constraints_with_heuristics env_rec !evdref in
+ let evd, nf = nf_evars_and_universes evd in
+ let fixdefs = List.map (Option.map nf) fixdefs in
+ let fixtypes = List.map nf fixtypes in
+ let fixctxnames = List.map (fun (_,ctx) -> List.map RelDecl.get_name ctx) fixctxs in
+
+ (* Build the fix declaration block *)
+ (env,rec_sign,all_universes,evd), (fixnames,fixdefs,fixtypes), List.combine3 fixctxnames fiximps fixannots
+
+let check_recursive isfix env evd (fixnames,fixdefs,_) =
+ check_evars_are_solved env evd (Evd.empty,evd);
+ if List.for_all Option.has_some fixdefs then begin
+ let fixdefs = List.map Option.get fixdefs in
+ check_mutuality env isfix (List.combine fixnames fixdefs)
+ end
+
+let interp_fixpoint l ntns =
+ let (env,_,pl,evd),fix,info = interp_recursive true l ntns in
+ check_recursive true env evd fix;
+ (fix,pl,Evd.evar_universe_context evd,info)
+
+let interp_cofixpoint l ntns =
+ let (env,_,pl,evd),fix,info = interp_recursive false l ntns in
+ check_recursive false env evd fix;
+ (fix,pl,Evd.evar_universe_context evd,info)
+
+let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) indexes ntns =
+ if List.exists Option.is_empty fixdefs then
+ (* Some bodies to define by proof *)
+ let thms =
+ List.map3 (fun id t (len,imps,_) -> ((id,pl),(t,(len,imps))))
+ fixnames fixtypes fiximps in
+ let init_tac =
+ Some (List.map (Option.cata Tacmach.refine_no_check Tacticals.tclIDTAC)
+ fixdefs) in
+ let init_tac =
+ Option.map (List.map Proofview.V82.tactic) init_tac
+ in
+ let evd = Evd.from_ctx ctx in
+ Lemmas.start_proof_with_initialization (Global,poly,DefinitionBody Fixpoint)
+ evd (Some(false,indexes,init_tac)) thms None (Lemmas.mk_hook (fun _ _ -> ()))
+ else begin
+ (* We shortcut the proof process *)
+ let fixdefs = List.map Option.get fixdefs in
+ let fixdecls = prepare_recursive_declaration fixnames fixtypes fixdefs in
+ let env = Global.env() in
+ let indexes = search_guard Loc.ghost env indexes fixdecls in
+ let fiximps = List.map (fun (n,r,p) -> r) fiximps in
+ let vars = Universes.universes_of_constr (mkFix ((indexes,0),fixdecls)) in
+ let fixdecls =
+ List.map_i (fun i _ -> mkFix ((indexes,i),fixdecls)) 0 fixnames in
+ let evd = Evd.from_ctx ctx in
+ let evd = Evd.restrict_universe_context evd vars in
+ let fixdecls = List.map Safe_typing.mk_pure_proof fixdecls in
+ let pl, ctx = Evd.universe_context ?names:pl evd in
+ ignore (List.map4 (declare_fix (local, poly, Fixpoint) pl ctx)
+ fixnames fixdecls fixtypes fiximps);
+ (* Declare the recursive definitions *)
+ fixpoint_message (Some indexes) fixnames;
+ end;
+ (* Declare notations *)
+ List.iter Metasyntax.add_notation_interpretation ntns
+
+let declare_cofixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ntns =
+ if List.exists Option.is_empty fixdefs then
+ (* Some bodies to define by proof *)
+ let thms =
+ List.map3 (fun id t (len,imps,_) -> ((id,pl),(t,(len,imps))))
+ fixnames fixtypes fiximps in
+ let init_tac =
+ Some (List.map (Option.cata Tacmach.refine_no_check Tacticals.tclIDTAC)
+ fixdefs) in
+ let init_tac =
+ Option.map (List.map Proofview.V82.tactic) init_tac
+ in
+ let evd = Evd.from_ctx ctx in
+ Lemmas.start_proof_with_initialization (Global,poly, DefinitionBody CoFixpoint)
+ evd (Some(true,[],init_tac)) thms None (Lemmas.mk_hook (fun _ _ -> ()))
+ else begin
+ (* We shortcut the proof process *)
+ let fixdefs = List.map Option.get fixdefs in
+ let fixdecls = prepare_recursive_declaration fixnames fixtypes fixdefs in
+ let fixdecls = List.map_i (fun i _ -> mkCoFix (i,fixdecls)) 0 fixnames in
+ let vars = Universes.universes_of_constr (List.hd fixdecls) in
+ let fixdecls = List.map Safe_typing.mk_pure_proof fixdecls in
+ let fiximps = List.map (fun (len,imps,idx) -> imps) fiximps in
+ let evd = Evd.from_ctx ctx in
+ let evd = Evd.restrict_universe_context evd vars in
+ let pl, ctx = Evd.universe_context ?names:pl evd in
+ ignore (List.map4 (declare_fix (local, poly, CoFixpoint) pl ctx)
+ fixnames fixdecls fixtypes fiximps);
+ (* Declare the recursive definitions *)
+ cofixpoint_message fixnames
+ end;
+ (* Declare notations *)
+ List.iter Metasyntax.add_notation_interpretation ntns
+
+let extract_decreasing_argument limit = function
+ | (na,CStructRec) -> na
+ | (na,_) when not limit -> na
+ | _ -> error
+ "Only structural decreasing is supported for a non-Program Fixpoint"
+
+let extract_fixpoint_components limit l =
+ let fixl, ntnl = List.split l in
+ let fixl = List.map (fun (((_,id),pl),ann,bl,typ,def) ->
+ let ann = extract_decreasing_argument limit ann in
+ {fix_name = id; fix_annot = ann; fix_univs = pl;
+ fix_binders = bl; fix_body = def; fix_type = typ}) fixl in
+ fixl, List.flatten ntnl
+
+let extract_cofixpoint_components l =
+ let fixl, ntnl = List.split l in
+ List.map (fun (((_,id),pl),bl,typ,def) ->
+ {fix_name = id; fix_annot = None; fix_univs = pl;
+ fix_binders = bl; fix_body = def; fix_type = typ}) fixl,
+ List.flatten ntnl
+
+let out_def = function
+ | Some def -> def
+ | None -> error "Program Fixpoint needs defined bodies."
+
+let collect_evars_of_term evd c ty =
+ let evars = Evar.Set.union (Evd.evars_of_term c) (Evd.evars_of_term ty) in
+ Evar.Set.fold (fun ev acc -> Evd.add acc ev (Evd.find_undefined evd ev))
+ evars (Evd.from_ctx (Evd.evar_universe_context evd))
+
+let do_program_recursive local p fixkind fixl ntns =
+ let isfix = fixkind != Obligations.IsCoFixpoint in
+ let (env, rec_sign, pl, evd), fix, info =
+ interp_recursive isfix fixl ntns
+ in
+ (* Program-specific code *)
+ (* Get the interesting evars, those that were not instanciated *)
+ let evd = Typeclasses.resolve_typeclasses ~filter:Typeclasses.no_goals ~fail:true env evd in
+ (* Solve remaining evars *)
+ let evd = nf_evar_map_undefined evd in
+ let collect_evars id def typ imps =
+ (* Generalize by the recursive prototypes *)
+ let def =
+ nf_evar evd (Termops.it_mkNamedLambda_or_LetIn def rec_sign)
+ and typ =
+ nf_evar evd (Termops.it_mkNamedProd_or_LetIn typ rec_sign)
+ in
+ let evm = collect_evars_of_term evd def typ in
+ let evars, _, def, typ =
+ Obligations.eterm_obligations env id evm
+ (List.length rec_sign) def typ
+ in (id, def, typ, imps, evars)
+ in
+ let (fixnames,fixdefs,fixtypes) = fix in
+ let fiximps = List.map pi2 info in
+ let fixdefs = List.map out_def fixdefs in
+ let defs = List.map4 collect_evars fixnames fixdefs fixtypes fiximps in
+ let () = if isfix then begin
+ let possible_indexes = List.map compute_possible_guardness_evidences info in
+ let fixdecls = Array.of_list (List.map (fun x -> Name x) fixnames),
+ Array.of_list fixtypes,
+ Array.of_list (List.map (subst_vars (List.rev fixnames)) fixdefs)
+ in
+ let indexes =
+ Pretyping.search_guard
+ Loc.ghost (Global.env ()) possible_indexes fixdecls in
+ List.iteri (fun i _ ->
+ Inductive.check_fix env
+ ((indexes,i),fixdecls))
+ fixl
+ end in
+ let ctx = Evd.evar_universe_context evd in
+ let kind = match fixkind with
+ | Obligations.IsFixpoint _ -> (local, p, Fixpoint)
+ | Obligations.IsCoFixpoint -> (local, p, CoFixpoint)
+ in
+ Obligations.add_mutual_definitions defs ~kind ?pl ctx ntns fixkind
+
+let do_program_fixpoint local poly l =
+ let g = List.map (fun ((_,wf,_,_,_),_) -> wf) l in
+ match g, l with
+ | [(n, CWfRec r)], [((((_,id),pl),_,bl,typ,def),ntn)] ->
+ let recarg =
+ match n with
+ | Some n -> mkIdentC (snd n)
+ | None ->
+ user_err ~hdr:"do_program_fixpoint"
+ (str "Recursive argument required for well-founded fixpoints")
+ in build_wellfounded (id, pl, n, bl, typ, out_def def) poly r recarg ntn
+
+ | [(n, CMeasureRec (m, r))], [((((_,id),pl),_,bl,typ,def),ntn)] ->
+ build_wellfounded (id, pl, n, bl, typ, out_def def) poly
+ (Option.default (CRef (lt_ref,None)) r) m ntn
+
+ | _, _ when List.for_all (fun (n, ro) -> ro == CStructRec) g ->
+ let fixl,ntns = extract_fixpoint_components true l in
+ let fixkind = Obligations.IsFixpoint g in
+ do_program_recursive local poly fixkind fixl ntns
+
+ | _, _ ->
+ user_err ~hdr:"do_program_fixpoint"
+ (str "Well-founded fixpoints not allowed in mutually recursive blocks")
+
+let check_safe () =
+ let open Declarations in
+ let flags = Environ.typing_flags (Global.env ()) in
+ flags.check_universes && flags.check_guarded
+
+let do_fixpoint local poly l =
+ if Flags.is_program_mode () then do_program_fixpoint local poly l
+ else
+ let fixl, ntns = extract_fixpoint_components true l in
+ let (_, _, _, info as fix) = interp_fixpoint fixl ntns in
+ let possible_indexes =
+ List.map compute_possible_guardness_evidences info in
+ declare_fixpoint local poly fix possible_indexes ntns;
+ if not (check_safe ()) then Feedback.feedback Feedback.AddedAxiom else ()
+
+let do_cofixpoint local poly l =
+ let fixl,ntns = extract_cofixpoint_components l in
+ if Flags.is_program_mode () then
+ do_program_recursive local poly Obligations.IsCoFixpoint fixl ntns
+ else
+ let cofix = interp_cofixpoint fixl ntns in
+ declare_cofixpoint local poly cofix ntns;
+ if not (check_safe ()) then Feedback.feedback Feedback.AddedAxiom else ()
diff --git a/vernac/command.mli b/vernac/command.mli
new file mode 100644
index 000000000..616afb91f
--- /dev/null
+++ b/vernac/command.mli
@@ -0,0 +1,176 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open Term
+open Entries
+open Libnames
+open Globnames
+open Vernacexpr
+open Constrexpr
+open Decl_kinds
+open Redexpr
+open Pfedit
+
+(** This file is about the interpretation of raw commands into typed
+ ones and top-level declaration of the main Gallina objects *)
+
+val do_universe : polymorphic -> Id.t Loc.located list -> unit
+val do_constraint : polymorphic ->
+ (Misctypes.glob_level * Univ.constraint_type * Misctypes.glob_level) list -> unit
+
+(** {6 Hooks for Pcoq} *)
+
+val set_declare_definition_hook : (Safe_typing.private_constants definition_entry -> unit) -> unit
+val get_declare_definition_hook : unit -> (Safe_typing.private_constants definition_entry -> unit)
+
+(** {6 Definitions/Let} *)
+
+val interp_definition :
+ lident list option -> local_binder list -> polymorphic -> red_expr option -> constr_expr ->
+ constr_expr option -> Safe_typing.private_constants definition_entry * Evd.evar_map *
+ Universes.universe_binders * Impargs.manual_implicits
+
+val declare_definition : Id.t -> definition_kind ->
+ Safe_typing.private_constants definition_entry -> Universes.universe_binders -> Impargs.manual_implicits ->
+ Globnames.global_reference Lemmas.declaration_hook -> Globnames.global_reference
+
+val do_definition : Id.t -> definition_kind -> lident list option ->
+ local_binder list -> red_expr option -> constr_expr ->
+ constr_expr option -> unit Lemmas.declaration_hook -> unit
+
+(** {6 Parameters/Assumptions} *)
+
+(* val interp_assumption : env -> evar_map ref -> *)
+(* local_binder list -> constr_expr -> *)
+(* types Univ.in_universe_context_set * Impargs.manual_implicits *)
+
+(** returns [false] if the assumption is neither local to a section,
+ nor in a module type and meant to be instantiated. *)
+val declare_assumption : coercion_flag -> assumption_kind ->
+ types Univ.in_universe_context_set ->
+ Universes.universe_binders -> Impargs.manual_implicits ->
+ bool (** implicit *) -> Vernacexpr.inline -> variable Loc.located ->
+ global_reference * Univ.Instance.t * bool
+
+val do_assumptions : locality * polymorphic * assumption_object_kind ->
+ Vernacexpr.inline -> (plident list * constr_expr) with_coercion list -> bool
+
+(* val declare_assumptions : variable Loc.located list -> *)
+(* coercion_flag -> assumption_kind -> types Univ.in_universe_context_set -> *)
+(* Impargs.manual_implicits -> bool -> Vernacexpr.inline -> bool *)
+
+(** {6 Inductive and coinductive types} *)
+
+(** Extracting the semantical components out of the raw syntax of mutual
+ inductive declarations *)
+
+type structured_one_inductive_expr = {
+ ind_name : Id.t;
+ ind_univs : lident list option;
+ ind_arity : constr_expr;
+ ind_lc : (Id.t * constr_expr) list
+}
+
+type structured_inductive_expr =
+ local_binder list * structured_one_inductive_expr list
+
+val extract_mutual_inductive_declaration_components :
+ (one_inductive_expr * decl_notation list) list ->
+ structured_inductive_expr * (*coercions:*) qualid list * decl_notation list
+
+(** Typing mutual inductive definitions *)
+
+type one_inductive_impls =
+ Impargs.manual_implicits (** for inds *)*
+ Impargs.manual_implicits list (** for constrs *)
+
+val interp_mutual_inductive :
+ structured_inductive_expr -> decl_notation list -> polymorphic ->
+ private_flag -> Decl_kinds.recursivity_kind ->
+ mutual_inductive_entry * Universes.universe_binders * one_inductive_impls list
+
+(** Registering a mutual inductive definition together with its
+ associated schemes *)
+
+val declare_mutual_inductive_with_eliminations :
+ mutual_inductive_entry -> Universes.universe_binders -> one_inductive_impls list ->
+ mutual_inductive
+
+(** Entry points for the vernacular commands Inductive and CoInductive *)
+
+val do_mutual_inductive :
+ (one_inductive_expr * decl_notation list) list -> polymorphic ->
+ private_flag -> Decl_kinds.recursivity_kind -> unit
+
+(** {6 Fixpoints and cofixpoints} *)
+
+type structured_fixpoint_expr = {
+ fix_name : Id.t;
+ fix_univs : lident list option;
+ fix_annot : Id.t Loc.located option;
+ fix_binders : local_binder list;
+ fix_body : constr_expr option;
+ fix_type : constr_expr
+}
+
+(** Extracting the semantical components out of the raw syntax of
+ (co)fixpoints declarations *)
+
+val extract_fixpoint_components : bool ->
+ (fixpoint_expr * decl_notation list) list ->
+ structured_fixpoint_expr list * decl_notation list
+
+val extract_cofixpoint_components :
+ (cofixpoint_expr * decl_notation list) list ->
+ structured_fixpoint_expr list * decl_notation list
+
+(** Typing global fixpoints and cofixpoint_expr *)
+
+type recursive_preentry =
+ Id.t list * constr option list * types list
+
+val interp_fixpoint :
+ structured_fixpoint_expr list -> decl_notation list ->
+ recursive_preentry * lident list option * Evd.evar_universe_context *
+ (Name.t list * Impargs.manual_implicits * int option) list
+
+val interp_cofixpoint :
+ structured_fixpoint_expr list -> decl_notation list ->
+ recursive_preentry * lident list option * Evd.evar_universe_context *
+ (Name.t list * Impargs.manual_implicits * int option) list
+
+(** Registering fixpoints and cofixpoints in the environment *)
+
+val declare_fixpoint :
+ locality -> polymorphic ->
+ recursive_preentry * lident list option * Evd.evar_universe_context *
+ (Name.t list * Impargs.manual_implicits * int option) list ->
+ lemma_possible_guards -> decl_notation list -> unit
+
+val declare_cofixpoint : locality -> polymorphic ->
+ recursive_preentry * lident list option * Evd.evar_universe_context *
+ (Name.t list * Impargs.manual_implicits * int option) list ->
+ decl_notation list -> unit
+
+(** Entry points for the vernacular commands Fixpoint and CoFixpoint *)
+
+val do_fixpoint :
+ (* When [false], assume guarded. *)
+ locality -> polymorphic -> (fixpoint_expr * decl_notation list) list -> unit
+
+val do_cofixpoint :
+ (* When [false], assume guarded. *)
+ locality -> polymorphic -> (cofixpoint_expr * decl_notation list) list -> unit
+
+(** Utils *)
+
+val check_mutuality : Environ.env -> bool -> (Id.t * types) list -> unit
+
+val declare_fix : ?opaque:bool -> definition_kind -> Universes.universe_binders -> Univ.universe_context -> Id.t ->
+ Safe_typing.private_constants Entries.proof_output -> types -> Impargs.manual_implicits -> global_reference
diff --git a/vernac/discharge.ml b/vernac/discharge.ml
new file mode 100644
index 000000000..e24d5e74f
--- /dev/null
+++ b/vernac/discharge.ml
@@ -0,0 +1,120 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open CErrors
+open Util
+open Term
+open Vars
+open Declarations
+open Cooking
+open Entries
+open Context.Rel.Declaration
+
+(********************************)
+(* Discharging mutual inductive *)
+
+let detype_param =
+ function
+ | LocalAssum (Name id, p) -> id, LocalAssumEntry p
+ | LocalDef (Name id, p,_) -> id, LocalDefEntry p
+ | _ -> anomaly (Pp.str "Unnamed inductive local variable")
+
+(* Replace
+
+ Var(y1)..Var(yq):C1..Cq |- Ij:Bj
+ Var(y1)..Var(yq):C1..Cq; I1..Ip:B1..Bp |- ci : Ti
+
+ by
+
+ |- Ij: (y1..yq:C1..Cq)Bj
+ I1..Ip:(B1 y1..yq)..(Bp y1..yq) |- ci : (y1..yq:C1..Cq)Ti[Ij:=(Ij y1..yq)]
+*)
+
+let abstract_inductive hyps nparams inds =
+ let ntyp = List.length inds in
+ let nhyp = Context.Named.length hyps in
+ let args = Context.Named.to_instance (List.rev hyps) in
+ let args = Array.of_list args in
+ let subs = List.init ntyp (fun k -> lift nhyp (mkApp(mkRel (k+1),args))) in
+ let inds' =
+ List.map
+ (function (tname,arity,template,cnames,lc) ->
+ let lc' = List.map (substl subs) lc in
+ let lc'' = List.map (fun b -> Termops.it_mkNamedProd_wo_LetIn b hyps) lc' in
+ let arity' = Termops.it_mkNamedProd_wo_LetIn arity hyps in
+ (tname,arity',template,cnames,lc''))
+ inds in
+ let nparams' = nparams + Array.length args in
+(* To be sure to be the same as before, should probably be moved to process_inductive *)
+ let params' = let (_,arity,_,_,_) = List.hd inds' in
+ let (params,_) = decompose_prod_n_assum nparams' arity in
+ List.map detype_param params
+ in
+ let ind'' =
+ List.map
+ (fun (a,arity,template,c,lc) ->
+ let _, short_arity = decompose_prod_n_assum nparams' arity in
+ let shortlc =
+ List.map (fun c -> snd (decompose_prod_n_assum nparams' c)) lc in
+ { mind_entry_typename = a;
+ mind_entry_arity = short_arity;
+ mind_entry_template = template;
+ mind_entry_consnames = c;
+ mind_entry_lc = shortlc })
+ inds'
+ in (params',ind'')
+
+let refresh_polymorphic_type_of_inductive (_,mip) =
+ match mip.mind_arity with
+ | RegularArity s -> s.mind_user_arity, false
+ | TemplateArity ar ->
+ let ctx = List.rev mip.mind_arity_ctxt in
+ mkArity (List.rev ctx, Type ar.template_level), true
+
+let process_inductive (sechyps,abs_ctx) modlist mib =
+ let nparams = mib.mind_nparams in
+ let subst, univs =
+ if mib.mind_polymorphic then
+ let inst = Univ.UContext.instance mib.mind_universes in
+ let cstrs = Univ.UContext.constraints mib.mind_universes in
+ inst, Univ.UContext.make (inst, Univ.subst_instance_constraints inst cstrs)
+ else Univ.Instance.empty, mib.mind_universes
+ in
+ let inds =
+ Array.map_to_list
+ (fun mip ->
+ let ty, template = refresh_polymorphic_type_of_inductive (mib,mip) in
+ let arity = expmod_constr modlist ty in
+ let arity = Vars.subst_instance_constr subst arity in
+ let lc = Array.map
+ (fun c -> Vars.subst_instance_constr subst (expmod_constr modlist c))
+ mip.mind_user_lc
+ in
+ (mip.mind_typename,
+ arity, template,
+ Array.to_list mip.mind_consnames,
+ Array.to_list lc))
+ mib.mind_packets in
+ let sechyps' = Context.Named.map (expmod_constr modlist) sechyps in
+ let (params',inds') = abstract_inductive sechyps' nparams inds in
+ let abs_ctx = Univ.instantiate_univ_context abs_ctx in
+ let univs = Univ.UContext.union abs_ctx univs in
+ let record = match mib.mind_record with
+ | Some (Some (id, _, _)) -> Some (Some id)
+ | Some None -> Some None
+ | None -> None
+ in
+ { mind_entry_record = record;
+ mind_entry_finite = mib.mind_finite;
+ mind_entry_params = params';
+ mind_entry_inds = inds';
+ mind_entry_polymorphic = mib.mind_polymorphic;
+ mind_entry_private = mib.mind_private;
+ mind_entry_universes = univs;
+ }
diff --git a/vernac/discharge.mli b/vernac/discharge.mli
new file mode 100644
index 000000000..18d1b6776
--- /dev/null
+++ b/vernac/discharge.mli
@@ -0,0 +1,14 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Declarations
+open Entries
+open Opaqueproof
+
+val process_inductive :
+ Context.Named.t Univ.in_universe_context -> work_list -> mutual_inductive_body -> mutual_inductive_entry
diff --git a/vernac/doc.tex b/vernac/doc.tex
new file mode 100644
index 000000000..f2550fda1
--- /dev/null
+++ b/vernac/doc.tex
@@ -0,0 +1,10 @@
+
+\newpage
+\section*{The Coq toplevel}
+
+\ocwsection \label{toplevel}
+This chapter describes the highest modules of the \Coq\ system.
+They are organized as follows:
+
+\bigskip
+\begin{center}\epsfig{file=toplevel.dep.ps,width=\linewidth}\end{center}
diff --git a/vernac/explainErr.ml b/vernac/explainErr.ml
new file mode 100644
index 000000000..17897460c
--- /dev/null
+++ b/vernac/explainErr.ml
@@ -0,0 +1,129 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Pp
+open CErrors
+open Indtypes
+open Type_errors
+open Pretype_errors
+open Indrec
+
+let guill s = str "\"" ++ str s ++ str "\""
+
+(** Invariant : exceptions embedded in EvaluatedError satisfy
+ Errors.noncritical *)
+
+exception EvaluatedError of std_ppcmds * exn option
+
+(** Registration of generic errors
+ Nota: explain_exn does NOT end with a newline anymore!
+*)
+
+let explain_exn_default = function
+ (* Basic interaction exceptions *)
+ | Stream.Error txt -> hov 0 (str "Syntax error: " ++ str txt ++ str ".")
+ | Compat.Token.Error txt -> hov 0 (str "Syntax error: " ++ str txt ++ str ".")
+ | CLexer.Error.E err -> hov 0 (str (CLexer.Error.to_string err))
+ | Sys_error msg -> hov 0 (str "System error: " ++ guill msg)
+ | Out_of_memory -> hov 0 (str "Out of memory.")
+ | Stack_overflow -> hov 0 (str "Stack overflow.")
+ | Timeout -> hov 0 (str "Timeout!")
+ | Sys.Break -> hov 0 (fnl () ++ str "User interrupt.")
+ (* Exceptions with pre-evaluated error messages *)
+ | EvaluatedError (msg,None) -> msg
+ | EvaluatedError (msg,Some reraise) -> msg ++ CErrors.print reraise
+ (* Otherwise, not handled here *)
+ | _ -> raise CErrors.Unhandled
+
+let _ = CErrors.register_handler explain_exn_default
+
+
+(** Pre-explain a vernac interpretation error *)
+
+let wrap_vernac_error with_header (exn, info) strm =
+ if with_header then
+ let header = Pp.tag (Pp.Tag.inj Ppstyle.error_tag Ppstyle.tag) (str "Error:") in
+ let e = EvaluatedError (hov 0 (header ++ spc () ++ strm), None) in
+ (e, info)
+ else
+ (EvaluatedError (strm, None), info)
+
+let process_vernac_interp_error with_header exn = match fst exn with
+ | Univ.UniverseInconsistency i ->
+ let msg =
+ if !Constrextern.print_universes then
+ str "." ++ spc() ++
+ Univ.explain_universe_inconsistency Universes.pr_with_global_universes i
+ else
+ mt() in
+ wrap_vernac_error with_header exn (str "Universe inconsistency" ++ msg ++ str ".")
+ | TypeError(ctx,te) ->
+ wrap_vernac_error with_header exn (Himsg.explain_type_error ctx Evd.empty te)
+ | PretypeError(ctx,sigma,te) ->
+ wrap_vernac_error with_header exn (Himsg.explain_pretype_error ctx sigma te)
+ | Typeclasses_errors.TypeClassError(env, te) ->
+ wrap_vernac_error with_header exn (Himsg.explain_typeclass_error env te)
+ | InductiveError e ->
+ wrap_vernac_error with_header exn (Himsg.explain_inductive_error e)
+ | Modops.ModuleTypingError e ->
+ wrap_vernac_error with_header exn (Himsg.explain_module_error e)
+ | Modintern.ModuleInternalizationError e ->
+ wrap_vernac_error with_header exn (Himsg.explain_module_internalization_error e)
+ | RecursionSchemeError e ->
+ wrap_vernac_error with_header exn (Himsg.explain_recursion_scheme_error e)
+ | Cases.PatternMatchingError (env,sigma,e) ->
+ wrap_vernac_error with_header exn (Himsg.explain_pattern_matching_error env sigma e)
+ | Tacred.ReductionTacticError e ->
+ wrap_vernac_error with_header exn (Himsg.explain_reduction_tactic_error e)
+ | Logic.RefinerError e ->
+ wrap_vernac_error with_header exn (Himsg.explain_refiner_error e)
+ | Nametab.GlobalizationError q ->
+ wrap_vernac_error with_header exn
+ (str "The reference" ++ spc () ++ Libnames.pr_qualid q ++
+ spc () ++ str "was not found" ++
+ spc () ++ str "in the current" ++ spc () ++ str "environment.")
+ | Refiner.FailError (i,s) ->
+ let s = Lazy.force s in
+ wrap_vernac_error with_header exn
+ (str "Tactic failure" ++
+ (if Pp.is_empty s then s else str ": " ++ s) ++
+ if Int.equal i 0 then str "." else str " (level " ++ int i ++ str").")
+ | AlreadyDeclared msg ->
+ wrap_vernac_error with_header exn (msg ++ str ".")
+ | _ ->
+ exn
+
+let rec strip_wrapping_exceptions = function
+ | Logic_monad.TacticFailure e ->
+ strip_wrapping_exceptions e
+ | exc -> exc
+
+let additional_error_info = ref []
+
+let register_additional_error_info f =
+ additional_error_info := f :: !additional_error_info
+
+let process_vernac_interp_error ?(allow_uncaught=true) ?(with_header=true) (exc, info) =
+ let exc = strip_wrapping_exceptions exc in
+ let e = process_vernac_interp_error with_header (exc, info) in
+ let () =
+ if not allow_uncaught && not (CErrors.handled (fst e)) then
+ let (e, info) = e in
+ let msg = str "Uncaught exception " ++ str (Printexc.to_string e) in
+ let err = CErrors.make_anomaly msg in
+ Util.iraise (err, info)
+ in
+ let e' =
+ try Some (CList.find_map (fun f -> f e) !additional_error_info)
+ with _ -> None
+ in
+ match e' with
+ | None -> e
+ | Some (None, loc) -> (fst e, Loc.add_loc (snd e) loc)
+ | Some (Some msg, loc) ->
+ (EvaluatedError (msg, Some (fst e)), Loc.add_loc (snd e) loc)
diff --git a/vernac/explainErr.mli b/vernac/explainErr.mli
new file mode 100644
index 000000000..a67c887af
--- /dev/null
+++ b/vernac/explainErr.mli
@@ -0,0 +1,21 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** Toplevel Exception *)
+exception EvaluatedError of Pp.std_ppcmds * exn option
+
+(** Pre-explain a vernac interpretation error *)
+
+val process_vernac_interp_error : ?allow_uncaught:bool -> ?with_header:bool -> Util.iexn -> Util.iexn
+
+(** General explain function. Should not be used directly now,
+ see instead function [Errors.print] and variants *)
+
+val explain_exn_default : exn -> Pp.std_ppcmds
+
+val register_additional_error_info : (Util.iexn -> (Pp.std_ppcmds option * Loc.t) option) -> unit
diff --git a/vernac/himsg.ml b/vernac/himsg.ml
new file mode 100644
index 000000000..6cff805fc
--- /dev/null
+++ b/vernac/himsg.ml
@@ -0,0 +1,1268 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Pp
+open Util
+open Names
+open Nameops
+open Namegen
+open Term
+open Termops
+open Indtypes
+open Environ
+open Pretype_errors
+open Type_errors
+open Typeclasses_errors
+open Indrec
+open Cases
+open Logic
+open Printer
+open Evd
+open Context.Rel.Declaration
+
+module RelDecl = Context.Rel.Declaration
+
+(* This simplifies the typing context of Cases clauses *)
+(* hope it does not disturb other typing contexts *)
+let contract env lc =
+ let l = ref [] in
+ let contract_context decl env =
+ match decl with
+ | LocalDef (_,c',_) when isRel c' ->
+ l := (Vars.substl !l c') :: !l;
+ env
+ | _ ->
+ let t = Vars.substl !l (RelDecl.get_type decl) in
+ let decl = decl |> RelDecl.map_name (named_hd env t) |> RelDecl.map_value (Vars.substl !l) |> RelDecl.set_type t in
+ l := (mkRel 1) :: List.map (Vars.lift 1) !l;
+ push_rel decl env
+ in
+ let env = process_rel_context contract_context env in
+ (env, List.map (Vars.substl !l) lc)
+
+let contract2 env a b = match contract env [a;b] with
+ | env, [a;b] -> env,a,b | _ -> assert false
+
+let contract3 env a b c = match contract env [a;b;c] with
+ | env, [a;b;c] -> env,a,b,c | _ -> assert false
+
+let contract4 env a b c d = match contract env [a;b;c;d] with
+ | env, [a;b;c;d] -> (env,a,b,c),d | _ -> assert false
+
+let contract1_vect env a v =
+ match contract env (a :: Array.to_list v) with
+ | env, a::l -> env,a,Array.of_list l
+ | _ -> assert false
+
+let rec contract3' env a b c = function
+ | OccurCheck (evk,d) -> let x,d = contract4 env a b c d in x,OccurCheck(evk,d)
+ | NotClean ((evk,args),env',d) ->
+ let env',d,args = contract1_vect env' d args in
+ contract3 env a b c,NotClean((evk,args),env',d)
+ | ConversionFailed (env',t1,t2) ->
+ let (env',t1,t2) = contract2 env' t1 t2 in
+ contract3 env a b c, ConversionFailed (env',t1,t2)
+ | NotSameArgSize | NotSameHead | NoCanonicalStructure
+ | MetaOccurInBody _ | InstanceNotSameType _ | ProblemBeyondCapabilities
+ | UnifUnivInconsistency _ as x -> contract3 env a b c, x
+ | CannotSolveConstraint ((pb,env',t,u),x) ->
+ let env',t,u = contract2 env' t u in
+ let y,x = contract3' env a b c x in
+ y,CannotSolveConstraint ((pb,env',t,u),x)
+
+(** Ad-hoc reductions *)
+
+let j_nf_betaiotaevar sigma j =
+ { uj_val = Evarutil.nf_evar sigma j.uj_val;
+ uj_type = Reductionops.nf_betaiota sigma j.uj_type }
+
+let jv_nf_betaiotaevar sigma jl =
+ Array.map (j_nf_betaiotaevar sigma) jl
+
+(** Printers *)
+
+let pr_lconstr c = quote (pr_lconstr c)
+let pr_lconstr_env e s c = quote (pr_lconstr_env e s c)
+let pr_ljudge_env e s c = let v,t = pr_ljudge_env e s c in (quote v,quote t)
+
+(** A canonisation procedure for constr such that comparing there
+ externalisation catches more equalities *)
+let canonize_constr c =
+ (* replaces all the names in binders by [dn] ("default name"),
+ ensures that [alpha]-equivalent terms will have the same
+ externalisation. *)
+ let dn = Name.Anonymous in
+ let rec canonize_binders c =
+ match Term.kind_of_term c with
+ | Prod (_,t,b) -> mkProd(dn,t,b)
+ | Lambda (_,t,b) -> mkLambda(dn,t,b)
+ | LetIn (_,u,t,b) -> mkLetIn(dn,u,t,b)
+ | _ -> Term.map_constr canonize_binders c
+ in
+ canonize_binders c
+
+(** Tries to realize when the two terms, albeit different are printed the same. *)
+let display_eq ~flags env sigma t1 t2 =
+ (* terms are canonized, then their externalisation is compared syntactically *)
+ let open Constrextern in
+ let t1 = canonize_constr t1 in
+ let t2 = canonize_constr t2 in
+ let ct1 = Flags.with_options flags (fun () -> extern_constr false env sigma t1) () in
+ let ct2 = Flags.with_options flags (fun () -> extern_constr false env sigma t2) () in
+ Constrexpr_ops.constr_expr_eq ct1 ct2
+
+(** This function adds some explicit printing flags if the two arguments are
+ printed alike. *)
+let rec pr_explicit_aux env sigma t1 t2 = function
+| [] ->
+ (** no specified flags: default. *)
+ (quote (Printer.pr_lconstr_env env sigma t1), quote (Printer.pr_lconstr_env env sigma t2))
+| flags :: rem ->
+ let equal = display_eq ~flags env sigma t1 t2 in
+ if equal then
+ (** The two terms are the same from the user point of view *)
+ pr_explicit_aux env sigma t1 t2 rem
+ else
+ let open Constrextern in
+ let ct1 = Flags.with_options flags (fun () -> extern_constr false env sigma t1) ()
+ in
+ let ct2 = Flags.with_options flags (fun () -> extern_constr false env sigma t2) ()
+ in
+ quote (Ppconstr.pr_lconstr_expr ct1), quote (Ppconstr.pr_lconstr_expr ct2)
+
+let explicit_flags =
+ let open Constrextern in
+ [ []; (** First, try with the current flags *)
+ [print_implicits]; (** Then with implicit *)
+ [print_universes]; (** Then with universes *)
+ [print_universes; print_implicits]; (** With universes AND implicits *)
+ [print_implicits; print_coercions; print_no_symbol]; (** Then more! *)
+ [print_universes; print_implicits; print_coercions; print_no_symbol] (** and more! *) ]
+
+let pr_explicit env sigma t1 t2 = pr_explicit_aux env sigma t1 t2 explicit_flags
+
+let pr_db env i =
+ try
+ match env |> lookup_rel i |> get_name with
+ | Name id -> pr_id id
+ | Anonymous -> str "<>"
+ with Not_found -> str "UNBOUND_REL_" ++ int i
+
+let explain_unbound_rel env sigma n =
+ let pe = pr_ne_context_of (str "In environment") env sigma in
+ str "Unbound reference: " ++ pe ++
+ str "The reference " ++ int n ++ str " is free."
+
+let explain_unbound_var env v =
+ let var = pr_id v in
+ str "No such section variable or assumption: " ++ var ++ str "."
+
+let explain_not_type env sigma j =
+ let j = Evarutil.j_nf_evar sigma j in
+ let pe = pr_ne_context_of (str "In environment") env sigma in
+ let pc,pt = pr_ljudge_env env sigma j in
+ pe ++ str "The term" ++ brk(1,1) ++ pc ++ spc () ++
+ str "has type" ++ spc () ++ pt ++ spc () ++
+ str "which should be Set, Prop or Type."
+
+let explain_bad_assumption env sigma j =
+ let pe = pr_ne_context_of (str "In environment") env sigma in
+ let pc,pt = pr_ljudge_env env sigma j in
+ pe ++ str "Cannot declare a variable or hypothesis over the term" ++
+ brk(1,1) ++ pc ++ spc () ++ str "of type" ++ spc () ++ pt ++ spc () ++
+ str "because this term is not a type."
+
+let explain_reference_variables id c =
+ (* c is intended to be a global reference *)
+ let pc = pr_global (Globnames.global_of_constr c) in
+ pc ++ strbrk " depends on the variable " ++ pr_id id ++
+ strbrk " which is not declared in the context."
+
+let rec pr_disjunction pr = function
+ | [a] -> pr a
+ | [a;b] -> pr a ++ str " or" ++ spc () ++ pr b
+ | a::l -> pr a ++ str "," ++ spc () ++ pr_disjunction pr l
+ | [] -> assert false
+
+let pr_puniverses f env (c,u) =
+ f env c ++
+ (if Flags.is_universe_polymorphism () && not (Univ.Instance.is_empty u) then
+ str"(*" ++ Univ.Instance.pr Universes.pr_with_global_universes u ++ str"*)"
+ else mt())
+
+let explain_elim_arity env sigma ind sorts c pj okinds =
+ let env = make_all_name_different env in
+ let pi = pr_inductive env (fst ind) in
+ let pc = pr_lconstr_env env sigma c in
+ let msg = match okinds with
+ | Some(kp,ki,explanation) ->
+ let pki = pr_sort_family ki in
+ let pkp = pr_sort_family kp in
+ let explanation = match explanation with
+ | NonInformativeToInformative ->
+ "proofs can be eliminated only to build proofs"
+ | StrongEliminationOnNonSmallType ->
+ "strong elimination on non-small inductive types leads to paradoxes"
+ | WrongArity ->
+ "wrong arity" in
+ let ppar = pr_disjunction (fun s -> quote (pr_sort_family s)) sorts in
+ let ppt = pr_lconstr_env env sigma ((strip_prod_assum pj.uj_type)) in
+ hov 0
+ (str "the return type has sort" ++ spc () ++ ppt ++ spc () ++
+ str "while it" ++ spc () ++ str "should be " ++ ppar ++ str ".") ++
+ fnl () ++
+ hov 0
+ (str "Elimination of an inductive object of sort " ++
+ pki ++ brk(1,0) ++
+ str "is not allowed on a predicate in sort " ++ pkp ++ fnl () ++
+ str "because" ++ spc () ++ str explanation ++ str ".")
+ | None ->
+ str "ill-formed elimination predicate."
+ in
+ hov 0 (
+ str "Incorrect elimination of" ++ spc () ++ pc ++ spc () ++
+ str "in the inductive type" ++ spc () ++ quote pi ++ str ":") ++
+ fnl () ++ msg
+
+let explain_case_not_inductive env sigma cj =
+ let cj = Evarutil.j_nf_evar sigma cj in
+ let env = make_all_name_different env in
+ let pc = pr_lconstr_env env sigma cj.uj_val in
+ let pct = pr_lconstr_env env sigma cj.uj_type in
+ match kind_of_term cj.uj_type with
+ | Evar _ ->
+ str "Cannot infer a type for this expression."
+ | _ ->
+ str "The term" ++ brk(1,1) ++ pc ++ spc () ++
+ str "has type" ++ brk(1,1) ++ pct ++ spc () ++
+ str "which is not a (co-)inductive type."
+
+let explain_number_branches env sigma cj expn =
+ let cj = Evarutil.j_nf_evar sigma cj in
+ let env = make_all_name_different env in
+ let pc = pr_lconstr_env env sigma cj.uj_val in
+ let pct = pr_lconstr_env env sigma cj.uj_type in
+ str "Matching on term" ++ brk(1,1) ++ pc ++ spc () ++
+ str "of type" ++ brk(1,1) ++ pct ++ spc () ++
+ str "expects " ++ int expn ++ str " branches."
+
+let explain_ill_formed_branch env sigma c ci actty expty =
+ let simp t = Reduction.nf_betaiota env (Evarutil.nf_evar sigma t) in
+ let c = Evarutil.nf_evar sigma c in
+ let env = make_all_name_different env in
+ let pc = pr_lconstr_env env sigma c in
+ let pa, pe = pr_explicit env sigma (simp actty) (simp expty) in
+ strbrk "In pattern-matching on term" ++ brk(1,1) ++ pc ++
+ spc () ++ strbrk "the branch for constructor" ++ spc () ++
+ quote (pr_puniverses pr_constructor env ci) ++
+ spc () ++ str "has type" ++ brk(1,1) ++ pa ++ spc () ++
+ str "which should be" ++ brk(1,1) ++ pe ++ str "."
+
+let explain_generalization env sigma (name,var) j =
+ let pe = pr_ne_context_of (str "In environment") env sigma in
+ let pv = pr_ltype_env env sigma var in
+ let (pc,pt) = pr_ljudge_env (push_rel_assum (name,var) env) sigma j in
+ pe ++ str "Cannot generalize" ++ brk(1,1) ++ pv ++ spc () ++
+ str "over" ++ brk(1,1) ++ pc ++ str "," ++ spc () ++
+ str "it has type" ++ spc () ++ pt ++
+ spc () ++ str "which should be Set, Prop or Type."
+
+let explain_unification_error env sigma p1 p2 = function
+ | None -> mt()
+ | Some e ->
+ let rec aux p1 p2 = function
+ | OccurCheck (evk,rhs) ->
+ let rhs = Evarutil.nf_evar sigma rhs in
+ [str "cannot define " ++ quote (pr_existential_key sigma evk) ++
+ strbrk " with term " ++ pr_lconstr_env env sigma rhs ++
+ strbrk " that would depend on itself"]
+ | NotClean ((evk,args),env,c) ->
+ let c = Evarutil.nf_evar sigma c in
+ let args = Array.map (Evarutil.nf_evar sigma) args in
+ [str "cannot instantiate " ++ quote (pr_existential_key sigma evk)
+ ++ strbrk " because " ++ pr_lconstr_env env sigma c ++
+ strbrk " is not in its scope" ++
+ (if Array.is_empty args then mt() else
+ strbrk ": available arguments are " ++
+ pr_sequence (pr_lconstr_env env sigma) (List.rev (Array.to_list args)))]
+ | NotSameArgSize | NotSameHead | NoCanonicalStructure ->
+ (* Error speaks from itself *) []
+ | ConversionFailed (env,t1,t2) ->
+ if Term.eq_constr t1 p1 && Term.eq_constr t2 p2 then [] else
+ let env = make_all_name_different env in
+ let t1 = Evarutil.nf_evar sigma t1 in
+ let t2 = Evarutil.nf_evar sigma t2 in
+ if not (Term.eq_constr t1 p1) || not (Term.eq_constr t2 p2) then
+ let t1, t2 = pr_explicit env sigma t1 t2 in
+ [str "cannot unify " ++ t1 ++ strbrk " and " ++ t2]
+ else []
+ | MetaOccurInBody evk ->
+ [str "instance for " ++ quote (pr_existential_key sigma evk) ++
+ strbrk " refers to a metavariable - please report your example" ++
+ strbrk "at " ++ str Coq_config.wwwbugtracker ++ str "."]
+ | InstanceNotSameType (evk,env,t,u) ->
+ let t, u = pr_explicit env sigma t u in
+ [str "unable to find a well-typed instantiation for " ++
+ quote (pr_existential_key sigma evk) ++
+ strbrk ": cannot ensure that " ++
+ t ++ strbrk " is a subtype of " ++ u]
+ | UnifUnivInconsistency p ->
+ if !Constrextern.print_universes then
+ [str "universe inconsistency: " ++
+ Univ.explain_universe_inconsistency Universes.pr_with_global_universes p]
+ else
+ [str "universe inconsistency"]
+ | CannotSolveConstraint ((pb,env,t,u),e) ->
+ let t = Evarutil.nf_evar sigma t in
+ let u = Evarutil.nf_evar sigma u in
+ let env = make_all_name_different env in
+ (strbrk "cannot satisfy constraint " ++ pr_lconstr_env env sigma t ++
+ str " == " ++ pr_lconstr_env env sigma u)
+ :: aux t u e
+ | ProblemBeyondCapabilities ->
+ []
+ in
+ match aux p1 p2 e with
+ | [] -> mt ()
+ | l -> spc () ++ str "(" ++
+ prlist_with_sep pr_semicolon (fun x -> x) l ++ str ")"
+
+let explain_actual_type env sigma j t reason =
+ let env = make_all_name_different env in
+ let j = j_nf_betaiotaevar sigma j in
+ let t = Reductionops.nf_betaiota sigma t in
+ (** Actually print *)
+ let pe = pr_ne_context_of (str "In environment") env sigma in
+ let pc = pr_lconstr_env env sigma (Environ.j_val j) in
+ let (pt, pct) = pr_explicit env sigma t (Environ.j_type j) in
+ let ppreason = explain_unification_error env sigma j.uj_type t reason in
+ pe ++
+ hov 0 (
+ str "The term" ++ brk(1,1) ++ pc ++ spc () ++
+ str "has type" ++ brk(1,1) ++ pct ++ spc () ++
+ str "while it is expected to have type" ++ brk(1,1) ++ pt ++
+ ppreason ++ str ".")
+
+let explain_cant_apply_bad_type env sigma (n,exptyp,actualtyp) rator randl =
+ let randl = jv_nf_betaiotaevar sigma randl in
+ let exptyp = Evarutil.nf_evar sigma exptyp in
+ let actualtyp = Reductionops.nf_betaiota sigma actualtyp in
+ let rator = Evarutil.j_nf_evar sigma rator in
+ let env = make_all_name_different env in
+ let actualtyp, exptyp = pr_explicit env sigma actualtyp exptyp in
+ let nargs = Array.length randl in
+(* let pe = pr_ne_context_of (str "in environment") env sigma in*)
+ let pr,prt = pr_ljudge_env env sigma rator in
+ let term_string1 = str (String.plural nargs "term") in
+ let term_string2 =
+ if nargs>1 then str "The " ++ pr_nth n ++ str " term" else str "This term"
+ in
+ let appl = prvect_with_sep fnl
+ (fun c ->
+ let pc,pct = pr_ljudge_env env sigma c in
+ hov 2 (pc ++ spc () ++ str ": " ++ pct)) randl
+ in
+ str "Illegal application: " ++ (* pe ++ *) fnl () ++
+ str "The term" ++ brk(1,1) ++ pr ++ spc () ++
+ str "of type" ++ brk(1,1) ++ prt ++ spc () ++
+ str "cannot be applied to the " ++ term_string1 ++ fnl () ++
+ str " " ++ v 0 appl ++ fnl () ++ term_string2 ++ str " has type" ++
+ brk(1,1) ++ actualtyp ++ spc () ++
+ str "which should be coercible to" ++ brk(1,1) ++
+ exptyp ++ str "."
+
+let explain_cant_apply_not_functional env sigma rator randl =
+ let randl = Evarutil.jv_nf_evar sigma randl in
+ let rator = Evarutil.j_nf_evar sigma rator in
+ let env = make_all_name_different env in
+ let nargs = Array.length randl in
+(* let pe = pr_ne_context_of (str "in environment") env sigma in*)
+ let pr = pr_lconstr_env env sigma rator.uj_val in
+ let prt = pr_lconstr_env env sigma rator.uj_type in
+ let appl = prvect_with_sep fnl
+ (fun c ->
+ let pc = pr_lconstr_env env sigma c.uj_val in
+ let pct = pr_lconstr_env env sigma c.uj_type in
+ hov 2 (pc ++ spc () ++ str ": " ++ pct)) randl
+ in
+ str "Illegal application (Non-functional construction): " ++
+ (* pe ++ *) fnl () ++
+ str "The expression" ++ brk(1,1) ++ pr ++ spc () ++
+ str "of type" ++ brk(1,1) ++ prt ++ spc () ++
+ str "cannot be applied to the " ++ str (String.plural nargs "term") ++
+ fnl () ++ str " " ++ v 0 appl
+
+let explain_unexpected_type env sigma actual_type expected_type =
+ let actual_type = Evarutil.nf_evar sigma actual_type in
+ let expected_type = Evarutil.nf_evar sigma expected_type in
+ let pract, prexp = pr_explicit env sigma actual_type expected_type in
+ str "Found type" ++ spc () ++ pract ++ spc () ++
+ str "where" ++ spc () ++ prexp ++ str " was expected."
+
+let explain_not_product env sigma c =
+ let c = Evarutil.nf_evar sigma c in
+ let pr = pr_lconstr_env env sigma c in
+ str "The type of this term is a product" ++ spc () ++
+ str "while it is expected to be" ++
+ (if is_Type c then str " a sort" else (brk(1,1) ++ pr)) ++ str "."
+
+(* TODO: use the names *)
+(* (co)fixpoints *)
+let explain_ill_formed_rec_body env sigma err names i fixenv vdefj =
+ let prt_name i =
+ match names.(i) with
+ Name id -> str "Recursive definition of " ++ pr_id id
+ | Anonymous -> str "The " ++ pr_nth i ++ str " definition" in
+
+ let st = match err with
+
+ (* Fixpoint guard errors *)
+ | NotEnoughAbstractionInFixBody ->
+ str "Not enough abstractions in the definition"
+ | RecursionNotOnInductiveType c ->
+ str "Recursive definition on" ++ spc () ++ pr_lconstr_env env sigma c ++
+ spc () ++ str "which should be an inductive type"
+ | RecursionOnIllegalTerm(j,(arg_env, arg),le,lt) ->
+ let arg_env = make_all_name_different arg_env in
+ let called =
+ match names.(j) with
+ Name id -> pr_id id
+ | Anonymous -> str "the " ++ pr_nth i ++ str " definition" in
+ let pr_db x = quote (pr_db env x) in
+ let vars =
+ match (lt,le) with
+ ([],[]) -> assert false
+ | ([],[x]) -> str "a subterm of " ++ pr_db x
+ | ([],_) -> str "a subterm of the following variables: " ++
+ pr_sequence pr_db le
+ | ([x],_) -> pr_db x
+ | _ ->
+ str "one of the following variables: " ++
+ pr_sequence pr_db lt in
+ str "Recursive call to " ++ called ++ spc () ++
+ strbrk "has principal argument equal to" ++ spc () ++
+ pr_lconstr_env arg_env sigma arg ++ strbrk " instead of " ++ vars
+
+ | NotEnoughArgumentsForFixCall j ->
+ let called =
+ match names.(j) with
+ Name id -> pr_id id
+ | Anonymous -> str "the " ++ pr_nth i ++ str " definition" in
+ str "Recursive call to " ++ called ++ str " has not enough arguments"
+
+ (* CoFixpoint guard errors *)
+ | CodomainNotInductiveType c ->
+ str "The codomain is" ++ spc () ++ pr_lconstr_env env sigma c ++ spc () ++
+ str "which should be a coinductive type"
+ | NestedRecursiveOccurrences ->
+ str "Nested recursive occurrences"
+ | UnguardedRecursiveCall c ->
+ str "Unguarded recursive call in" ++ spc () ++ pr_lconstr_env env sigma c
+ | RecCallInTypeOfAbstraction c ->
+ str "Recursive call forbidden in the domain of an abstraction:" ++
+ spc () ++ pr_lconstr_env env sigma c
+ | RecCallInNonRecArgOfConstructor c ->
+ str "Recursive call on a non-recursive argument of constructor" ++
+ spc () ++ pr_lconstr_env env sigma c
+ | RecCallInTypeOfDef c ->
+ str "Recursive call forbidden in the type of a recursive definition" ++
+ spc () ++ pr_lconstr_env env sigma c
+ | RecCallInCaseFun c ->
+ str "Invalid recursive call in a branch of" ++
+ spc () ++ pr_lconstr_env env sigma c
+ | RecCallInCaseArg c ->
+ str "Invalid recursive call in the argument of \"match\" in" ++ spc () ++
+ pr_lconstr_env env sigma c
+ | RecCallInCasePred c ->
+ str "Invalid recursive call in the \"return\" clause of \"match\" in" ++
+ spc () ++ pr_lconstr_env env sigma c
+ | NotGuardedForm c ->
+ str "Sub-expression " ++ pr_lconstr_env env sigma c ++
+ strbrk " not in guarded form (should be a constructor," ++
+ strbrk " an abstraction, a match, a cofix or a recursive call)"
+ | ReturnPredicateNotCoInductive c ->
+ str "The return clause of the following pattern matching should be" ++
+ strbrk " a coinductive type:" ++
+ spc () ++ pr_lconstr_env env sigma c
+ in
+ prt_name i ++ str " is ill-formed." ++ fnl () ++
+ pr_ne_context_of (str "In environment") env sigma ++
+ st ++ str "." ++ fnl () ++
+ (try (* May fail with unresolved globals. *)
+ let fixenv = make_all_name_different fixenv in
+ let pvd = pr_lconstr_env fixenv sigma vdefj.(i).uj_val in
+ str"Recursive definition is:" ++ spc () ++ pvd ++ str "."
+ with e when CErrors.noncritical e -> mt ())
+
+let explain_ill_typed_rec_body env sigma i names vdefj vargs =
+ let vdefj = Evarutil.jv_nf_evar sigma vdefj in
+ let vargs = Array.map (Evarutil.nf_evar sigma) vargs in
+ let env = make_all_name_different env in
+ let pvd = pr_lconstr_env env sigma vdefj.(i).uj_val in
+ let pvdt, pv = pr_explicit env sigma vdefj.(i).uj_type vargs.(i) in
+ str "The " ++
+ (match vdefj with [|_|] -> mt () | _ -> pr_nth (i+1) ++ spc ()) ++
+ str "recursive definition" ++ spc () ++ pvd ++ spc () ++
+ str "has type" ++ spc () ++ pvdt ++ spc () ++
+ str "while it should be" ++ spc () ++ pv ++ str "."
+
+let explain_cant_find_case_type env sigma c =
+ let c = Evarutil.nf_evar sigma c in
+ let env = make_all_name_different env in
+ let pe = pr_lconstr_env env sigma c in
+ str "Cannot infer the return type of pattern-matching on" ++ ws 1 ++
+ pe ++ str "."
+
+let explain_occur_check env sigma ev rhs =
+ let rhs = Evarutil.nf_evar sigma rhs in
+ let env = make_all_name_different env in
+ let pt = pr_lconstr_env env sigma rhs in
+ str "Cannot define " ++ pr_existential_key sigma ev ++ str " with term" ++
+ brk(1,1) ++ pt ++ spc () ++ str "that would depend on itself."
+
+let pr_trailing_ne_context_of env sigma =
+ if List.is_empty (Environ.rel_context env) &&
+ List.is_empty (Environ.named_context env)
+ then str "."
+ else (str " in environment:"++ pr_context_unlimited env sigma)
+
+let rec explain_evar_kind env sigma evk ty = function
+ | Evar_kinds.NamedHole id ->
+ strbrk "the existential variable named " ++ pr_id id
+ | Evar_kinds.QuestionMark _ ->
+ strbrk "this placeholder of type " ++ ty
+ | Evar_kinds.CasesType false ->
+ strbrk "the type of this pattern-matching problem"
+ | Evar_kinds.CasesType true ->
+ strbrk "a subterm of type " ++ ty ++
+ strbrk " in the type of this pattern-matching problem"
+ | Evar_kinds.BinderType (Name id) ->
+ strbrk "the type of " ++ Nameops.pr_id id
+ | Evar_kinds.BinderType Anonymous ->
+ strbrk "the type of this anonymous binder"
+ | Evar_kinds.ImplicitArg (c,(n,ido),b) ->
+ let id = Option.get ido in
+ strbrk "the implicit parameter " ++ pr_id id ++ spc () ++ str "of" ++
+ spc () ++ Nametab.pr_global_env Id.Set.empty c ++
+ strbrk " whose type is " ++ ty
+ | Evar_kinds.InternalHole -> strbrk "an internal placeholder of type " ++ ty
+ | Evar_kinds.TomatchTypeParameter (tyi,n) ->
+ strbrk "the " ++ pr_nth n ++
+ strbrk " argument of the inductive type (" ++ pr_inductive env tyi ++
+ strbrk ") of this term"
+ | Evar_kinds.GoalEvar ->
+ strbrk "an existential variable of type " ++ ty
+ | Evar_kinds.ImpossibleCase ->
+ strbrk "the type of an impossible pattern-matching clause"
+ | Evar_kinds.MatchingVar _ ->
+ assert false
+ | Evar_kinds.VarInstance id ->
+ strbrk "an instance of type " ++ ty ++
+ str " for the variable " ++ pr_id id
+ | Evar_kinds.SubEvar evk' ->
+ let evi = Evd.find sigma evk' in
+ let pc = match evi.evar_body with
+ | Evar_defined c -> pr_lconstr_env env sigma (Evarutil.nf_evar sigma c)
+ | Evar_empty -> assert false in
+ let ty' = Evarutil.nf_evar sigma evi.evar_concl in
+ pr_existential_key sigma evk ++ str " of type " ++ ty ++
+ str " in the partial instance " ++ pc ++
+ str " found for " ++ explain_evar_kind env sigma evk'
+ (pr_lconstr_env env sigma ty') (snd evi.evar_source)
+
+let explain_typeclass_resolution env sigma evi k =
+ match Typeclasses.class_of_constr evi.evar_concl with
+ | Some _ ->
+ let env = Evd.evar_filtered_env evi in
+ fnl () ++ str "Could not find an instance for " ++
+ pr_lconstr_env env sigma evi.evar_concl ++
+ pr_trailing_ne_context_of env sigma
+ | _ -> mt()
+
+let explain_placeholder_kind env sigma c e =
+ match e with
+ | Some (SeveralInstancesFound n) ->
+ strbrk " (several distinct possible type class instances found)"
+ | None ->
+ match Typeclasses.class_of_constr c with
+ | Some _ -> strbrk " (no type class instance found)"
+ | _ -> mt ()
+
+let explain_unsolvable_implicit env sigma evk explain =
+ let evi = Evarutil.nf_evar_info sigma (Evd.find_undefined sigma evk) in
+ let env = Evd.evar_filtered_env evi in
+ let type_of_hole = pr_lconstr_env env sigma evi.evar_concl in
+ let pe = pr_trailing_ne_context_of env sigma in
+ strbrk "Cannot infer " ++
+ explain_evar_kind env sigma evk type_of_hole (snd evi.evar_source) ++
+ explain_placeholder_kind env sigma evi.evar_concl explain ++ pe
+
+let explain_var_not_found env id =
+ str "The variable" ++ spc () ++ pr_id id ++
+ spc () ++ str "was not found" ++
+ spc () ++ str "in the current" ++ spc () ++ str "environment" ++ str "."
+
+let explain_wrong_case_info env (ind,u) ci =
+ let pi = pr_inductive (Global.env()) ind in
+ if eq_ind ci.ci_ind ind then
+ str "Pattern-matching expression on an object of inductive type" ++
+ spc () ++ pi ++ spc () ++ str "has invalid information."
+ else
+ let pc = pr_inductive (Global.env()) ci.ci_ind in
+ str "A term of inductive type" ++ spc () ++ pi ++ spc () ++
+ str "was given to a pattern-matching expression on the inductive type" ++
+ spc () ++ pc ++ str "."
+
+let explain_cannot_unify env sigma m n e =
+ let env = make_all_name_different env in
+ let m = Evarutil.nf_evar sigma m in
+ let n = Evarutil.nf_evar sigma n in
+ let pm, pn = pr_explicit env sigma m n in
+ let ppreason = explain_unification_error env sigma m n e in
+ let pe = pr_ne_context_of (str "In environment") env sigma in
+ pe ++ str "Unable to unify" ++ brk(1,1) ++ pm ++ spc () ++
+ str "with" ++ brk(1,1) ++ pn ++ ppreason ++ str "."
+
+let explain_cannot_unify_local env sigma m n subn =
+ let pm = pr_lconstr_env env sigma m in
+ let pn = pr_lconstr_env env sigma n in
+ let psubn = pr_lconstr_env env sigma subn in
+ str "Unable to unify" ++ brk(1,1) ++ pm ++ spc () ++
+ str "with" ++ brk(1,1) ++ pn ++ spc () ++ str "as" ++ brk(1,1) ++
+ psubn ++ str " contains local variables."
+
+let explain_refiner_cannot_generalize env sigma ty =
+ str "Cannot find a well-typed generalisation of the goal with type: " ++
+ pr_lconstr_env env sigma ty ++ str "."
+
+let explain_no_occurrence_found env sigma c id =
+ str "Found no subterm matching " ++ pr_lconstr_env env sigma c ++
+ str " in " ++
+ (match id with
+ | Some id -> pr_id id
+ | None -> str"the current goal") ++ str "."
+
+let explain_cannot_unify_binding_type env sigma m n =
+ let pm = pr_lconstr_env env sigma m in
+ let pn = pr_lconstr_env env sigma n in
+ str "This binding has type" ++ brk(1,1) ++ pm ++ spc () ++
+ str "which should be unifiable with" ++ brk(1,1) ++ pn ++ str "."
+
+let explain_cannot_find_well_typed_abstraction env sigma p l e =
+ str "Abstracting over the " ++
+ str (String.plural (List.length l) "term") ++ spc () ++
+ hov 0 (pr_enum (pr_lconstr_env env sigma) l) ++ spc () ++
+ str "leads to a term" ++ spc () ++ pr_lconstr_goal_style_env env sigma p ++
+ spc () ++ str "which is ill-typed." ++
+ (match e with None -> mt () | Some e -> fnl () ++ str "Reason is: " ++ e)
+
+let explain_wrong_abstraction_type env sigma na abs expected result =
+ let ppname = match na with Name id -> pr_id id ++ spc () | _ -> mt () in
+ str "Cannot instantiate metavariable " ++ ppname ++ strbrk "of type " ++
+ pr_lconstr_env env sigma expected ++ strbrk " with abstraction " ++
+ pr_lconstr_env env sigma abs ++ strbrk " of incompatible type " ++
+ pr_lconstr_env env sigma result ++ str "."
+
+let explain_abstraction_over_meta _ m n =
+ strbrk "Too complex unification problem: cannot find a solution for both " ++
+ pr_name m ++ spc () ++ str "and " ++ pr_name n ++ str "."
+
+let explain_non_linear_unification env sigma m t =
+ strbrk "Cannot unambiguously instantiate " ++
+ pr_name m ++ str ":" ++
+ strbrk " which would require to abstract twice on " ++
+ pr_lconstr_env env sigma t ++ str "."
+
+let explain_unsatisfied_constraints env sigma cst =
+ strbrk "Unsatisfied constraints: " ++
+ Univ.pr_constraints (Evd.pr_evd_level sigma) cst ++
+ spc () ++ str "(maybe a bugged tactic)."
+
+let explain_type_error env sigma err =
+ let env = make_all_name_different env in
+ match err with
+ | UnboundRel n ->
+ explain_unbound_rel env sigma n
+ | UnboundVar v ->
+ explain_unbound_var env v
+ | NotAType j ->
+ explain_not_type env sigma j
+ | BadAssumption c ->
+ explain_bad_assumption env sigma c
+ | ReferenceVariables (id,c) ->
+ explain_reference_variables id c
+ | ElimArity (ind, aritylst, c, pj, okinds) ->
+ explain_elim_arity env sigma ind aritylst c pj okinds
+ | CaseNotInductive cj ->
+ explain_case_not_inductive env sigma cj
+ | NumberBranches (cj, n) ->
+ explain_number_branches env sigma cj n
+ | IllFormedBranch (c, i, actty, expty) ->
+ explain_ill_formed_branch env sigma c i actty expty
+ | Generalization (nvar, c) ->
+ explain_generalization env sigma nvar c
+ | ActualType (j, pt) ->
+ explain_actual_type env sigma j pt None
+ | CantApplyBadType (t, rator, randl) ->
+ explain_cant_apply_bad_type env sigma t rator randl
+ | CantApplyNonFunctional (rator, randl) ->
+ explain_cant_apply_not_functional env sigma rator randl
+ | IllFormedRecBody (err, lna, i, fixenv, vdefj) ->
+ explain_ill_formed_rec_body env sigma err lna i fixenv vdefj
+ | IllTypedRecBody (i, lna, vdefj, vargs) ->
+ explain_ill_typed_rec_body env sigma i lna vdefj vargs
+ | WrongCaseInfo (ind,ci) ->
+ explain_wrong_case_info env ind ci
+ | UnsatisfiedConstraints cst ->
+ explain_unsatisfied_constraints env sigma cst
+
+let pr_position (cl,pos) =
+ let clpos = match cl with
+ | None -> str " of the goal"
+ | Some (id,Locus.InHyp) -> str " of hypothesis " ++ pr_id id
+ | Some (id,Locus.InHypTypeOnly) -> str " of the type of hypothesis " ++ pr_id id
+ | Some (id,Locus.InHypValueOnly) -> str " of the body of hypothesis " ++ pr_id id in
+ int pos ++ clpos
+
+let explain_cannot_unify_occurrences env sigma nested ((cl2,pos2),t2) ((cl1,pos1),t1) e =
+ if nested then
+ str "Found nested occurrences of the pattern at positions " ++
+ int pos1 ++ strbrk " and " ++ pr_position (cl2,pos2) ++ str "."
+ else
+ let ppreason = match e with None -> mt() | Some (c1,c2,e) -> explain_unification_error env sigma c1 c2 (Some e) in
+ str "Found incompatible occurrences of the pattern" ++ str ":" ++
+ spc () ++ str "Matched term " ++ pr_lconstr_env env sigma t2 ++
+ strbrk " at position " ++ pr_position (cl2,pos2) ++
+ strbrk " is not compatible with matched term " ++
+ pr_lconstr_env env sigma t1 ++ strbrk " at position " ++
+ pr_position (cl1,pos1) ++ ppreason ++ str "."
+
+let pr_constraints printenv env sigma evars cstrs =
+ let (ev, evi) = Evar.Map.choose evars in
+ if Evar.Map.for_all (fun ev' evi' ->
+ eq_named_context_val evi.evar_hyps evi'.evar_hyps) evars
+ then
+ let l = Evar.Map.bindings evars in
+ let env' = reset_with_named_context evi.evar_hyps env in
+ let pe =
+ if printenv then
+ pr_ne_context_of (str "In environment:") env' sigma
+ else mt ()
+ in
+ let evs =
+ prlist
+ (fun (ev, evi) -> fnl () ++ pr_existential_key sigma ev ++
+ str " : " ++ pr_lconstr_env env' sigma evi.evar_concl ++ fnl ()) l
+ in
+ h 0 (pe ++ evs ++ pr_evar_constraints cstrs)
+ else
+ let filter evk _ = Evar.Map.mem evk evars in
+ pr_evar_map_filter ~with_univs:false filter sigma
+
+let explain_unsatisfiable_constraints env sigma constr comp =
+ let (_, constraints) = Evd.extract_all_conv_pbs sigma in
+ let undef = Evd.undefined_map (Evarutil.nf_evar_map_undefined sigma) in
+ (** Only keep evars that are subject to resolution and members of the given
+ component. *)
+ let is_kept evk evi = match comp with
+ | None -> Typeclasses.is_resolvable evi
+ | Some comp -> Typeclasses.is_resolvable evi && Evar.Set.mem evk comp
+ in
+ let undef =
+ let m = Evar.Map.filter is_kept undef in
+ if Evar.Map.is_empty m then undef
+ else m
+ in
+ match constr with
+ | None ->
+ str "Unable to satisfy the following constraints:" ++ fnl () ++
+ pr_constraints true env sigma undef constraints
+ | Some (ev, k) ->
+ let cstr =
+ let remaining = Evar.Map.remove ev undef in
+ if not (Evar.Map.is_empty remaining) then
+ str "With the following constraints:" ++ fnl () ++
+ pr_constraints false env sigma remaining constraints
+ else mt ()
+ in
+ let info = Evar.Map.find ev undef in
+ explain_typeclass_resolution env sigma info k ++ fnl () ++ cstr
+
+let explain_pretype_error env sigma err =
+ let env = Evardefine.env_nf_betaiotaevar sigma env in
+ let env = make_all_name_different env in
+ match err with
+ | CantFindCaseType c -> explain_cant_find_case_type env sigma c
+ | ActualTypeNotCoercible (j,t,e) ->
+ let {uj_val = c; uj_type = actty} = j in
+ let (env, c, actty, expty), e = contract3' env c actty t e in
+ let j = {uj_val = c; uj_type = actty} in
+ explain_actual_type env sigma j expty (Some e)
+ | UnifOccurCheck (ev,rhs) -> explain_occur_check env sigma ev rhs
+ | UnsolvableImplicit (evk,exp) -> explain_unsolvable_implicit env sigma evk exp
+ | VarNotFound id -> explain_var_not_found env id
+ | UnexpectedType (actual,expect) ->
+ let env, actual, expect = contract2 env actual expect in
+ explain_unexpected_type env sigma actual expect
+ | NotProduct c -> explain_not_product env sigma c
+ | CannotUnify (m,n,e) ->
+ let env, m, n = contract2 env m n in
+ explain_cannot_unify env sigma m n e
+ | CannotUnifyLocal (m,n,sn) -> explain_cannot_unify_local env sigma m n sn
+ | CannotGeneralize ty -> explain_refiner_cannot_generalize env sigma ty
+ | NoOccurrenceFound (c, id) -> explain_no_occurrence_found env sigma c id
+ | CannotUnifyBindingType (m,n) -> explain_cannot_unify_binding_type env sigma m n
+ | CannotFindWellTypedAbstraction (p,l,e) ->
+ explain_cannot_find_well_typed_abstraction env sigma p l
+ (Option.map (fun (env',e) -> explain_type_error env' sigma e) e)
+ | WrongAbstractionType (n,a,t,u) ->
+ explain_wrong_abstraction_type env sigma n a t u
+ | AbstractionOverMeta (m,n) -> explain_abstraction_over_meta env m n
+ | NonLinearUnification (m,c) -> explain_non_linear_unification env sigma m c
+ | TypingError t -> explain_type_error env sigma t
+ | CannotUnifyOccurrences (b,c1,c2,e) -> explain_cannot_unify_occurrences env sigma b c1 c2 e
+ | UnsatisfiableConstraints (c,comp) -> explain_unsatisfiable_constraints env sigma c comp
+(* Module errors *)
+
+open Modops
+
+let explain_not_match_error = function
+ | InductiveFieldExpected _ ->
+ strbrk "an inductive definition is expected"
+ | DefinitionFieldExpected ->
+ strbrk "a definition is expected"
+ | ModuleFieldExpected ->
+ strbrk "a module is expected"
+ | ModuleTypeFieldExpected ->
+ strbrk "a module type is expected"
+ | NotConvertibleInductiveField id | NotConvertibleConstructorField id ->
+ str "types given to " ++ pr_id id ++ str " differ"
+ | NotConvertibleBodyField ->
+ str "the body of definitions differs"
+ | NotConvertibleTypeField (env, typ1, typ2) ->
+ str "expected type" ++ spc () ++
+ quote (Printer.safe_pr_lconstr_env env Evd.empty typ2) ++ spc () ++
+ str "but found type" ++ spc () ++
+ quote (Printer.safe_pr_lconstr_env env Evd.empty typ1)
+ | NotSameConstructorNamesField ->
+ str "constructor names differ"
+ | NotSameInductiveNameInBlockField ->
+ str "inductive types names differ"
+ | FiniteInductiveFieldExpected isfinite ->
+ str "type is expected to be " ++
+ str (if isfinite then "coinductive" else "inductive")
+ | InductiveNumbersFieldExpected n ->
+ str "number of inductive types differs"
+ | InductiveParamsNumberField n ->
+ str "inductive type has not the right number of parameters"
+ | RecordFieldExpected isrecord ->
+ str "type is expected " ++ str (if isrecord then "" else "not ") ++
+ str "to be a record"
+ | RecordProjectionsExpected nal ->
+ (if List.length nal >= 2 then str "expected projection names are "
+ else str "expected projection name is ") ++
+ pr_enum (function Name id -> pr_id id | _ -> str "_") nal
+ | NotEqualInductiveAliases ->
+ str "Aliases to inductive types do not match"
+ | NoTypeConstraintExpected ->
+ strbrk "a definition whose type is constrained can only be subtype " ++
+ strbrk "of a definition whose type is itself constrained"
+ | PolymorphicStatusExpected b ->
+ let status b = if b then str"polymorphic" else str"monomorphic" in
+ str "a " ++ status b ++ str" declaration was expected, but a " ++
+ status (not b) ++ str" declaration was found"
+ | IncompatibleInstances ->
+ str"polymorphic universe instances do not match"
+ | IncompatibleUniverses incon ->
+ str"the universe constraints are inconsistent: " ++
+ Univ.explain_universe_inconsistency Universes.pr_with_global_universes incon
+ | IncompatiblePolymorphism (env, t1, t2) ->
+ str "conversion of polymorphic values generates additional constraints: " ++
+ quote (Printer.safe_pr_lconstr_env env Evd.empty t1) ++ spc () ++
+ str "compared to " ++ spc () ++
+ quote (Printer.safe_pr_lconstr_env env Evd.empty t2)
+ | IncompatibleConstraints cst ->
+ str " the expected (polymorphic) constraints do not imply " ++
+ quote (Univ.pr_constraints (Evd.pr_evd_level Evd.empty) cst)
+
+let explain_signature_mismatch l spec why =
+ str "Signature components for label " ++ pr_label l ++
+ str " do not match:" ++ spc () ++ explain_not_match_error why ++ str "."
+
+let explain_label_already_declared l =
+ str "The label " ++ pr_label l ++ str " is already declared."
+
+let explain_application_to_not_path _ =
+ strbrk "A module cannot be applied to another module application or " ++
+ strbrk "with-expression; you must give a name to the intermediate result " ++
+ strbrk "module first."
+
+let explain_not_a_functor () =
+ str "Application of a non-functor."
+
+let explain_is_a_functor () =
+ str "Illegal use of a functor."
+
+let explain_incompatible_module_types mexpr1 mexpr2 =
+ let open Declarations in
+ let rec get_arg = function
+ | NoFunctor _ -> 0
+ | MoreFunctor (_, _, ty) -> succ (get_arg ty)
+ in
+ let len1 = get_arg mexpr1.mod_type in
+ let len2 = get_arg mexpr2.mod_type in
+ if len1 <> len2 then
+ str "Incompatible module types: module expects " ++ int len2 ++
+ str " arguments, found " ++ int len1 ++ str "."
+ else str "Incompatible module types."
+
+let explain_not_equal_module_paths mp1 mp2 =
+ str "Non equal modules."
+
+let explain_no_such_label l =
+ str "No such label " ++ pr_label l ++ str "."
+
+let explain_incompatible_labels l l' =
+ str "Opening and closing labels are not the same: " ++
+ pr_label l ++ str " <> " ++ pr_label l' ++ str "!"
+
+let explain_not_a_module s =
+ quote (str s) ++ str " is not a module."
+
+let explain_not_a_module_type s =
+ quote (str s) ++ str " is not a module type."
+
+let explain_not_a_constant l =
+ quote (pr_label l) ++ str " is not a constant."
+
+let explain_incorrect_label_constraint l =
+ str "Incorrect constraint for label " ++
+ quote (pr_label l) ++ str "."
+
+let explain_generative_module_expected l =
+ str "The module " ++ pr_label l ++ str " is not generative." ++
+ strbrk " Only components of generative modules can be changed" ++
+ strbrk " using the \"with\" construct."
+
+let explain_label_missing l s =
+ str "The field " ++ pr_label l ++ str " is missing in "
+ ++ str s ++ str "."
+
+let explain_include_restricted_functor mp =
+ let q = Nametab.shortest_qualid_of_module mp in
+ str "Cannot include the functor " ++ Libnames.pr_qualid q ++
+ strbrk " since it has a restricted signature. " ++
+ strbrk "You may name first an instance of this functor, and include it."
+
+let explain_module_error = function
+ | SignatureMismatch (l,spec,err) -> explain_signature_mismatch l spec err
+ | LabelAlreadyDeclared l -> explain_label_already_declared l
+ | ApplicationToNotPath mexpr -> explain_application_to_not_path mexpr
+ | NotAFunctor -> explain_not_a_functor ()
+ | IsAFunctor -> explain_is_a_functor ()
+ | IncompatibleModuleTypes (m1,m2) -> explain_incompatible_module_types m1 m2
+ | NotEqualModulePaths (mp1,mp2) -> explain_not_equal_module_paths mp1 mp2
+ | NoSuchLabel l -> explain_no_such_label l
+ | IncompatibleLabels (l1,l2) -> explain_incompatible_labels l1 l2
+ | NotAModule s -> explain_not_a_module s
+ | NotAModuleType s -> explain_not_a_module_type s
+ | NotAConstant l -> explain_not_a_constant l
+ | IncorrectWithConstraint l -> explain_incorrect_label_constraint l
+ | GenerativeModuleExpected l -> explain_generative_module_expected l
+ | LabelMissing (l,s) -> explain_label_missing l s
+ | IncludeRestrictedFunctor mp -> explain_include_restricted_functor mp
+
+(* Module internalization errors *)
+
+(*
+let explain_declaration_not_path _ =
+ str "Declaration is not a path."
+
+*)
+
+let explain_not_module_nor_modtype s =
+ quote (str s) ++ str " is not a module or module type."
+
+let explain_incorrect_with_in_module () =
+ str "The syntax \"with\" is not allowed for modules."
+
+let explain_incorrect_module_application () =
+ str "Illegal application to a module type."
+
+open Modintern
+
+let explain_module_internalization_error = function
+ | NotAModuleNorModtype s -> explain_not_module_nor_modtype s
+ | IncorrectWithInModule -> explain_incorrect_with_in_module ()
+ | IncorrectModuleApplication -> explain_incorrect_module_application ()
+
+(* Typeclass errors *)
+
+let explain_not_a_class env c =
+ pr_constr_env env Evd.empty c ++ str" is not a declared type class."
+
+let explain_unbound_method env cid id =
+ str "Unbound method name " ++ Nameops.pr_id (snd id) ++ spc () ++
+ str"of class" ++ spc () ++ pr_global cid ++ str "."
+
+let pr_constr_exprs exprs =
+ hv 0 (List.fold_right
+ (fun d pps -> ws 2 ++ Ppconstr.pr_constr_expr d ++ pps)
+ exprs (mt ()))
+
+let explain_mismatched_contexts env c i j =
+ str"Mismatched contexts while declaring instance: " ++ brk (1,1) ++
+ hov 1 (str"Expected:" ++ brk (1, 1) ++ pr_rel_context env Evd.empty j) ++
+ fnl () ++ brk (1,1) ++
+ hov 1 (str"Found:" ++ brk (1, 1) ++ pr_constr_exprs i)
+
+let explain_typeclass_error env = function
+ | NotAClass c -> explain_not_a_class env c
+ | UnboundMethod (cid, id) -> explain_unbound_method env cid id
+ | MismatchedContextInstance (c,i,j) -> explain_mismatched_contexts env c i j
+
+(* Refiner errors *)
+
+let explain_refiner_bad_type arg ty conclty =
+ str "Refiner was given an argument" ++ brk(1,1) ++
+ pr_lconstr arg ++ spc () ++
+ str "of type" ++ brk(1,1) ++ pr_lconstr ty ++ spc () ++
+ str "instead of" ++ brk(1,1) ++ pr_lconstr conclty ++ str "."
+
+let explain_refiner_unresolved_bindings l =
+ str "Unable to find an instance for the " ++
+ str (String.plural (List.length l) "variable") ++ spc () ++
+ prlist_with_sep pr_comma pr_name l ++ str"."
+
+let explain_refiner_cannot_apply t harg =
+ str "In refiner, a term of type" ++ brk(1,1) ++
+ pr_lconstr t ++ spc () ++ str "could not be applied to" ++ brk(1,1) ++
+ pr_lconstr harg ++ str "."
+
+let explain_refiner_not_well_typed c =
+ str "The term " ++ pr_lconstr c ++ str " is not well-typed."
+
+let explain_intro_needs_product () =
+ str "Introduction tactics needs products."
+
+let explain_does_not_occur_in c hyp =
+ str "The term" ++ spc () ++ pr_lconstr c ++ spc () ++
+ str "does not occur in" ++ spc () ++ pr_id hyp ++ str "."
+
+let explain_non_linear_proof c =
+ str "Cannot refine with term" ++ brk(1,1) ++ pr_lconstr c ++
+ spc () ++ str "because a metavariable has several occurrences."
+
+let explain_meta_in_type c =
+ str "In refiner, a meta appears in the type " ++ brk(1,1) ++ pr_lconstr c ++
+ str " of another meta"
+
+let explain_no_such_hyp id =
+ str "No such hypothesis: " ++ pr_id id
+
+let explain_refiner_error = function
+ | BadType (arg,ty,conclty) -> explain_refiner_bad_type arg ty conclty
+ | UnresolvedBindings t -> explain_refiner_unresolved_bindings t
+ | CannotApply (t,harg) -> explain_refiner_cannot_apply t harg
+ | NotWellTyped c -> explain_refiner_not_well_typed c
+ | IntroNeedsProduct -> explain_intro_needs_product ()
+ | DoesNotOccurIn (c,hyp) -> explain_does_not_occur_in c hyp
+ | NonLinearProof c -> explain_non_linear_proof c
+ | MetaInType c -> explain_meta_in_type c
+ | NoSuchHyp id -> explain_no_such_hyp id
+
+(* Inductive errors *)
+
+let error_non_strictly_positive env c v =
+ let pc = pr_lconstr_env env Evd.empty c in
+ let pv = pr_lconstr_env env Evd.empty v in
+ str "Non strictly positive occurrence of " ++ pv ++ str " in" ++
+ brk(1,1) ++ pc ++ str "."
+
+let error_ill_formed_inductive env c v =
+ let pc = pr_lconstr_env env Evd.empty c in
+ let pv = pr_lconstr_env env Evd.empty v in
+ str "Not enough arguments applied to the " ++ pv ++
+ str " in" ++ brk(1,1) ++ pc ++ str "."
+
+let error_ill_formed_constructor env id c v nparams nargs =
+ let pv = pr_lconstr_env env Evd.empty v in
+ let atomic = Int.equal (nb_prod c) 0 in
+ str "The type of constructor" ++ brk(1,1) ++ pr_id id ++ brk(1,1) ++
+ str "is not valid;" ++ brk(1,1) ++
+ strbrk (if atomic then "it must be " else "its conclusion must be ") ++
+ pv ++
+ (* warning: because of implicit arguments it is difficult to say which
+ parameters must be explicitly given *)
+ (if not (Int.equal nparams 0) then
+ strbrk " applied to its " ++ str (String.plural nparams "parameter")
+ else
+ mt()) ++
+ (if not (Int.equal nargs 0) then
+ str (if not (Int.equal nparams 0) then " and" else " applied") ++
+ strbrk " to some " ++ str (String.plural nargs "argument")
+ else
+ mt()) ++ str "."
+
+let pr_ltype_using_barendregt_convention_env env c =
+ (* Use goal_concl_style as an approximation of Barendregt's convention (?) *)
+ quote (pr_goal_concl_style_env env Evd.empty c)
+
+let error_bad_ind_parameters env c n v1 v2 =
+ let pc = pr_ltype_using_barendregt_convention_env env c in
+ let pv1 = pr_lconstr_env env Evd.empty v1 in
+ let pv2 = pr_lconstr_env env Evd.empty v2 in
+ str "Last occurrence of " ++ pv2 ++ str " must have " ++ pv1 ++
+ str " as " ++ pr_nth n ++ str " argument in" ++ brk(1,1) ++ pc ++ str "."
+
+let error_same_names_types id =
+ str "The name" ++ spc () ++ pr_id id ++ spc () ++
+ str "is used more than once."
+
+let error_same_names_constructors id =
+ str "The constructor name" ++ spc () ++ pr_id id ++ spc () ++
+ str "is used more than once."
+
+let error_same_names_overlap idl =
+ strbrk "The following names are used both as type names and constructor " ++
+ str "names:" ++ spc () ++
+ prlist_with_sep pr_comma pr_id idl ++ str "."
+
+let error_not_an_arity env c =
+ str "The type" ++ spc () ++ pr_lconstr_env env Evd.empty c ++ spc () ++
+ str "is not an arity."
+
+let error_bad_entry () =
+ str "Bad inductive definition."
+
+let error_large_non_prop_inductive_not_in_type () =
+ str "Large non-propositional inductive types must be in Type."
+
+(* Recursion schemes errors *)
+
+let error_not_allowed_case_analysis isrec kind i =
+ str (if isrec then "Induction" else "Case analysis") ++
+ strbrk " on sort " ++ pr_sort Evd.empty kind ++
+ strbrk " is not allowed for inductive definition " ++
+ pr_inductive (Global.env()) (fst i) ++ str "."
+
+let error_not_allowed_dependent_analysis isrec i =
+ str "Dependent " ++ str (if isrec then "Induction" else "Case analysis") ++
+ strbrk " is not allowed for inductive definition " ++
+ pr_inductive (Global.env()) i ++ str "."
+
+let error_not_mutual_in_scheme ind ind' =
+ if eq_ind ind ind' then
+ str "The inductive type " ++ pr_inductive (Global.env()) ind ++
+ str " occurs twice."
+ else
+ str "The inductive types " ++ pr_inductive (Global.env()) ind ++ spc () ++
+ str "and" ++ spc () ++ pr_inductive (Global.env()) ind' ++ spc () ++
+ str "are not mutually defined."
+
+(* Inductive constructions errors *)
+
+let explain_inductive_error = function
+ | NonPos (env,c,v) -> error_non_strictly_positive env c v
+ | NotEnoughArgs (env,c,v) -> error_ill_formed_inductive env c v
+ | NotConstructor (env,id,c,v,n,m) ->
+ error_ill_formed_constructor env id c v n m
+ | NonPar (env,c,n,v1,v2) -> error_bad_ind_parameters env c n v1 v2
+ | SameNamesTypes id -> error_same_names_types id
+ | SameNamesConstructors id -> error_same_names_constructors id
+ | SameNamesOverlap idl -> error_same_names_overlap idl
+ | NotAnArity (env, c) -> error_not_an_arity env c
+ | BadEntry -> error_bad_entry ()
+ | LargeNonPropInductiveNotInType ->
+ error_large_non_prop_inductive_not_in_type ()
+
+(* Recursion schemes errors *)
+
+let explain_recursion_scheme_error = function
+ | NotAllowedCaseAnalysis (isrec,k,i) ->
+ error_not_allowed_case_analysis isrec k i
+ | NotMutualInScheme (ind,ind')-> error_not_mutual_in_scheme ind ind'
+ | NotAllowedDependentAnalysis (isrec, i) ->
+ error_not_allowed_dependent_analysis isrec i
+
+(* Pattern-matching errors *)
+
+let explain_bad_pattern env sigma cstr ty =
+ let env = make_all_name_different env in
+ let pt = pr_lconstr_env env sigma ty in
+ let pc = pr_constructor env cstr in
+ str "Found the constructor " ++ pc ++ brk(1,1) ++
+ str "while matching a term of type " ++ pt ++ brk(1,1) ++
+ str "which is not an inductive type."
+
+let explain_bad_constructor env cstr ind =
+ let pi = pr_inductive env ind in
+(* let pc = pr_constructor env cstr in*)
+ let pt = pr_inductive env (inductive_of_constructor cstr) in
+ str "Found a constructor of inductive type " ++ pt ++ brk(1,1) ++
+ str "while a constructor of " ++ pi ++ brk(1,1) ++
+ str "is expected."
+
+let decline_string n s =
+ if Int.equal n 0 then str "no " ++ str s ++ str "s"
+ else if Int.equal n 1 then str "1 " ++ str s
+ else (int n ++ str " " ++ str s ++ str "s")
+
+let explain_wrong_numarg_constructor env cstr n =
+ str "The constructor " ++ pr_constructor env cstr ++
+ str " (in type " ++ pr_inductive env (inductive_of_constructor cstr) ++
+ str ") expects " ++ decline_string n "argument" ++ str "."
+
+let explain_wrong_numarg_inductive env ind n =
+ str "The inductive type " ++ pr_inductive env ind ++
+ str " expects " ++ decline_string n "argument" ++ str "."
+
+let explain_unused_clause env pats =
+(* Without localisation
+ let s = if List.length pats > 1 then "s" else "" in
+ (str ("Unused clause with pattern"^s) ++ spc () ++
+ hov 0 (pr_sequence pr_cases_pattern pats) ++ str ")")
+*)
+ str "This clause is redundant."
+
+let explain_non_exhaustive env pats =
+ str "Non exhaustive pattern-matching: no clause found for " ++
+ str (String.plural (List.length pats) "pattern") ++
+ spc () ++ hov 0 (prlist_with_sep pr_comma pr_cases_pattern pats)
+
+let explain_cannot_infer_predicate env sigma typs =
+ let env = make_all_name_different env in
+ let pr_branch (cstr,typ) =
+ let cstr,_ = decompose_app cstr in
+ str "For " ++ pr_lconstr_env env sigma cstr ++ str ": " ++ pr_lconstr_env env sigma typ
+ in
+ str "Unable to unify the types found in the branches:" ++
+ spc () ++ hov 0 (prlist_with_sep fnl pr_branch (Array.to_list typs))
+
+let explain_pattern_matching_error env sigma = function
+ | BadPattern (c,t) ->
+ explain_bad_pattern env sigma c t
+ | BadConstructor (c,ind) ->
+ explain_bad_constructor env c ind
+ | WrongNumargConstructor (c,n) ->
+ explain_wrong_numarg_constructor env c n
+ | WrongNumargInductive (c,n) ->
+ explain_wrong_numarg_inductive env c n
+ | UnusedClause tms ->
+ explain_unused_clause env tms
+ | NonExhaustive tms ->
+ explain_non_exhaustive env tms
+ | CannotInferPredicate typs ->
+ explain_cannot_infer_predicate env sigma typs
+
+let explain_reduction_tactic_error = function
+ | Tacred.InvalidAbstraction (env,sigma,c,(env',e)) ->
+ str "The abstracted term" ++ spc () ++
+ quote (pr_goal_concl_style_env env sigma c) ++
+ spc () ++ str "is not well typed." ++ fnl () ++
+ explain_type_error env' Evd.empty e
diff --git a/vernac/himsg.mli b/vernac/himsg.mli
new file mode 100644
index 000000000..ced54fd27
--- /dev/null
+++ b/vernac/himsg.mli
@@ -0,0 +1,42 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Pp
+open Indtypes
+open Environ
+open Type_errors
+open Pretype_errors
+open Typeclasses_errors
+open Indrec
+open Cases
+open Logic
+
+(** This module provides functions to explain the type errors. *)
+
+val explain_type_error : env -> Evd.evar_map -> type_error -> std_ppcmds
+
+val explain_pretype_error : env -> Evd.evar_map -> pretype_error -> std_ppcmds
+
+val explain_inductive_error : inductive_error -> std_ppcmds
+
+val explain_typeclass_error : env -> typeclass_error -> Pp.std_ppcmds
+
+val explain_recursion_scheme_error : recursion_scheme_error -> std_ppcmds
+
+val explain_refiner_error : refiner_error -> std_ppcmds
+
+val explain_pattern_matching_error :
+ env -> Evd.evar_map -> pattern_matching_error -> std_ppcmds
+
+val explain_reduction_tactic_error :
+ Tacred.reduction_tactic_error -> std_ppcmds
+
+val explain_module_error : Modops.module_typing_error -> std_ppcmds
+
+val explain_module_internalization_error :
+ Modintern.module_internalization_error -> std_ppcmds
diff --git a/vernac/ind_tables.ml b/vernac/ind_tables.ml
new file mode 100644
index 000000000..85d0b6194
--- /dev/null
+++ b/vernac/ind_tables.ml
@@ -0,0 +1,203 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* File created by Vincent Siles, Oct 2007, extended into a generic
+ support for generation of inductive schemes by Hugo Herbelin, Nov 2009 *)
+
+(* This file provides support for registering inductive scheme builders,
+ declaring schemes and generating schemes on demand *)
+
+open Names
+open Mod_subst
+open Libobject
+open Nameops
+open Declarations
+open Term
+open CErrors
+open Util
+open Declare
+open Entries
+open Decl_kinds
+open Pp
+
+(**********************************************************************)
+(* Registering schemes in the environment *)
+
+type mutual_scheme_object_function =
+ internal_flag -> mutual_inductive -> constr array Evd.in_evar_universe_context * Safe_typing.private_constants
+type individual_scheme_object_function =
+ internal_flag -> inductive -> constr Evd.in_evar_universe_context * Safe_typing.private_constants
+
+type 'a scheme_kind = string
+
+let scheme_map = Summary.ref Indmap.empty ~name:"Schemes"
+
+let pr_scheme_kind = Pp.str
+
+let cache_one_scheme kind (ind,const) =
+ let map = try Indmap.find ind !scheme_map with Not_found -> String.Map.empty in
+ scheme_map := Indmap.add ind (String.Map.add kind const map) !scheme_map
+
+let cache_scheme (_,(kind,l)) =
+ Array.iter (cache_one_scheme kind) l
+
+let subst_one_scheme subst (ind,const) =
+ (* Remark: const is a def: the result of substitution is a constant *)
+ (subst_ind subst ind,subst_constant subst const)
+
+let subst_scheme (subst,(kind,l)) =
+ (kind,Array.map (subst_one_scheme subst) l)
+
+let discharge_scheme (_,(kind,l)) =
+ Some (kind,Array.map (fun (ind,const) ->
+ (Lib.discharge_inductive ind,Lib.discharge_con const)) l)
+
+let inScheme : string * (inductive * constant) array -> obj =
+ declare_object {(default_object "SCHEME") with
+ cache_function = cache_scheme;
+ load_function = (fun _ -> cache_scheme);
+ subst_function = subst_scheme;
+ classify_function = (fun obj -> Substitute obj);
+ discharge_function = discharge_scheme}
+
+(**********************************************************************)
+(* The table of scheme building functions *)
+
+type individual
+type mutual
+
+type scheme_object_function =
+ | MutualSchemeFunction of mutual_scheme_object_function
+ | IndividualSchemeFunction of individual_scheme_object_function
+
+let scheme_object_table =
+ (Hashtbl.create 17 : (string, string * scheme_object_function) Hashtbl.t)
+
+let declare_scheme_object s aux f =
+ let () =
+ if not (Id.is_valid ("ind" ^ s)) then
+ error ("Illegal induction scheme suffix: " ^ s)
+ in
+ let key = if String.is_empty aux then s else aux in
+ try
+ let _ = Hashtbl.find scheme_object_table key in
+(* let aux_msg = if aux="" then "" else " (with key "^aux^")" in*)
+ user_err ~hdr:"IndTables.declare_scheme_object"
+ (str "Scheme object " ++ str key ++ str " already declared.")
+ with Not_found ->
+ Hashtbl.add scheme_object_table key (s,f);
+ key
+
+let declare_mutual_scheme_object s ?(aux="") f =
+ declare_scheme_object s aux (MutualSchemeFunction f)
+
+let declare_individual_scheme_object s ?(aux="") f =
+ declare_scheme_object s aux (IndividualSchemeFunction f)
+
+(**********************************************************************)
+(* Defining/retrieving schemes *)
+
+let declare_scheme kind indcl =
+ Lib.add_anonymous_leaf (inScheme (kind,indcl))
+
+let () = Declare.set_declare_scheme declare_scheme
+
+let is_visible_name id =
+ try ignore (Nametab.locate (Libnames.qualid_of_ident id)); true
+ with Not_found -> false
+
+let compute_name internal id =
+ match internal with
+ | UserAutomaticRequest | UserIndividualRequest -> id
+ | InternalTacticRequest ->
+ Namegen.next_ident_away_from (add_prefix "internal_" id) is_visible_name
+
+let define internal id c p univs =
+ let fd = declare_constant ~internal in
+ let id = compute_name internal id in
+ let ctx = Evd.normalize_evar_universe_context univs in
+ let c = Vars.subst_univs_fn_constr
+ (Universes.make_opt_subst (Evd.evar_universe_context_subst ctx)) c in
+ let entry = {
+ const_entry_body =
+ Future.from_val ((c,Univ.ContextSet.empty),
+ Safe_typing.empty_private_constants);
+ const_entry_secctx = None;
+ const_entry_type = None;
+ const_entry_polymorphic = p;
+ const_entry_universes = Evd.evar_context_universe_context ctx;
+ const_entry_opaque = false;
+ const_entry_inline_code = false;
+ const_entry_feedback = None;
+ } in
+ let kn = fd id (DefinitionEntry entry, Decl_kinds.IsDefinition Scheme) in
+ let () = match internal with
+ | InternalTacticRequest -> ()
+ | _-> definition_message id
+ in
+ kn
+
+let define_individual_scheme_base kind suff f mode idopt (mind,i as ind) =
+ let (c, ctx), eff = f mode ind in
+ let mib = Global.lookup_mind mind in
+ let id = match idopt with
+ | Some id -> id
+ | None -> add_suffix mib.mind_packets.(i).mind_typename suff in
+ let const = define mode id c mib.mind_polymorphic ctx in
+ declare_scheme kind [|ind,const|];
+ const, Safe_typing.add_private
+ (Safe_typing.private_con_of_scheme kind (Global.safe_env()) [ind,const]) eff
+
+let define_individual_scheme kind mode names (mind,i as ind) =
+ match Hashtbl.find scheme_object_table kind with
+ | _,MutualSchemeFunction f -> assert false
+ | s,IndividualSchemeFunction f ->
+ define_individual_scheme_base kind s f mode names ind
+
+let define_mutual_scheme_base kind suff f mode names mind =
+ let (cl, ctx), eff = f mode mind in
+ let mib = Global.lookup_mind mind in
+ let ids = Array.init (Array.length mib.mind_packets) (fun i ->
+ try Int.List.assoc i names
+ with Not_found -> add_suffix mib.mind_packets.(i).mind_typename suff) in
+ let consts = Array.map2 (fun id cl ->
+ define mode id cl mib.mind_polymorphic ctx) ids cl in
+ let schemes = Array.mapi (fun i cst -> ((mind,i),cst)) consts in
+ declare_scheme kind schemes;
+ consts,
+ Safe_typing.add_private
+ (Safe_typing.private_con_of_scheme
+ kind (Global.safe_env()) (Array.to_list schemes))
+ eff
+
+let define_mutual_scheme kind mode names mind =
+ match Hashtbl.find scheme_object_table kind with
+ | _,IndividualSchemeFunction _ -> assert false
+ | s,MutualSchemeFunction f ->
+ define_mutual_scheme_base kind s f mode names mind
+
+let find_scheme_on_env_too kind ind =
+ let s = String.Map.find kind (Indmap.find ind !scheme_map) in
+ s, Safe_typing.add_private
+ (Safe_typing.private_con_of_scheme
+ kind (Global.safe_env()) [ind, s])
+ Safe_typing.empty_private_constants
+
+let find_scheme ?(mode=InternalTacticRequest) kind (mind,i as ind) =
+ try find_scheme_on_env_too kind ind
+ with Not_found ->
+ match Hashtbl.find scheme_object_table kind with
+ | s,IndividualSchemeFunction f ->
+ define_individual_scheme_base kind s f mode None ind
+ | s,MutualSchemeFunction f ->
+ let ca, eff = define_mutual_scheme_base kind s f mode [] mind in
+ ca.(i), eff
+
+let check_scheme kind ind =
+ try let _ = find_scheme_on_env_too kind ind in true
+ with Not_found -> false
diff --git a/vernac/ind_tables.mli b/vernac/ind_tables.mli
new file mode 100644
index 000000000..20f30d6d1
--- /dev/null
+++ b/vernac/ind_tables.mli
@@ -0,0 +1,51 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Term
+open Names
+open Declare
+
+(** This module provides support for registering inductive scheme builders,
+ declaring schemes and generating schemes on demand *)
+
+(** A scheme is either a "mutual scheme_kind" or an "individual scheme_kind" *)
+
+type mutual
+type individual
+type 'a scheme_kind
+
+type mutual_scheme_object_function =
+ internal_flag -> mutual_inductive -> constr array Evd.in_evar_universe_context * Safe_typing.private_constants
+type individual_scheme_object_function =
+ internal_flag -> inductive -> constr Evd.in_evar_universe_context * Safe_typing.private_constants
+
+(** Main functions to register a scheme builder *)
+
+val declare_mutual_scheme_object : string -> ?aux:string ->
+ mutual_scheme_object_function -> mutual scheme_kind
+
+val declare_individual_scheme_object : string -> ?aux:string ->
+ individual_scheme_object_function ->
+ individual scheme_kind
+
+(** Force generation of a (mutually) scheme with possibly user-level names *)
+
+val define_individual_scheme : individual scheme_kind ->
+ internal_flag (** internal *) ->
+ Id.t option -> inductive -> constant * Safe_typing.private_constants
+
+val define_mutual_scheme : mutual scheme_kind -> internal_flag (** internal *) ->
+ (int * Id.t) list -> mutual_inductive -> constant array * Safe_typing.private_constants
+
+(** Main function to retrieve a scheme in the cache or to generate it *)
+val find_scheme : ?mode:internal_flag -> 'a scheme_kind -> inductive -> constant * Safe_typing.private_constants
+
+val check_scheme : 'a scheme_kind -> inductive -> bool
+
+
+val pr_scheme_kind : 'a scheme_kind -> Pp.std_ppcmds
diff --git a/vernac/indschemes.ml b/vernac/indschemes.ml
new file mode 100644
index 000000000..f7e3f0d95
--- /dev/null
+++ b/vernac/indschemes.ml
@@ -0,0 +1,530 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* Created by Hugo Herbelin from contents related to inductive schemes
+ initially developed by Christine Paulin (induction schemes), Vincent
+ Siles (decidable equality and boolean equality) and Matthieu Sozeau
+ (combined scheme) in file command.ml, Sep 2009 *)
+
+(* This file provides entry points for manually or automatically
+ declaring new schemes *)
+
+open Pp
+open CErrors
+open Util
+open Names
+open Declarations
+open Entries
+open Term
+open Inductive
+open Decl_kinds
+open Indrec
+open Declare
+open Libnames
+open Globnames
+open Goptions
+open Nameops
+open Termops
+open Pretyping
+open Nametab
+open Smartlocate
+open Vernacexpr
+open Ind_tables
+open Auto_ind_decl
+open Eqschemes
+open Elimschemes
+open Context.Rel.Declaration
+
+(* Flags governing automatic synthesis of schemes *)
+
+let elim_flag = ref true
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "automatic declaration of induction schemes";
+ optkey = ["Elimination";"Schemes"];
+ optread = (fun () -> !elim_flag) ;
+ optwrite = (fun b -> elim_flag := b) }
+
+let bifinite_elim_flag = ref false
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "automatic declaration of induction schemes for non-recursive types";
+ optkey = ["Nonrecursive";"Elimination";"Schemes"];
+ optread = (fun () -> !bifinite_elim_flag) ;
+ optwrite = (fun b -> bifinite_elim_flag := b) }
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = true; (* compatibility 2014-09-03*)
+ optname = "automatic declaration of induction schemes for non-recursive types";
+ optkey = ["Record";"Elimination";"Schemes"];
+ optread = (fun () -> !bifinite_elim_flag) ;
+ optwrite = (fun b -> bifinite_elim_flag := b) }
+
+let case_flag = ref false
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "automatic declaration of case analysis schemes";
+ optkey = ["Case";"Analysis";"Schemes"];
+ optread = (fun () -> !case_flag) ;
+ optwrite = (fun b -> case_flag := b) }
+
+let eq_flag = ref false
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "automatic declaration of boolean equality";
+ optkey = ["Boolean";"Equality";"Schemes"];
+ optread = (fun () -> !eq_flag) ;
+ optwrite = (fun b -> eq_flag := b) }
+let _ = (* compatibility *)
+ declare_bool_option
+ { optsync = true;
+ optdepr = true;
+ optname = "automatic declaration of boolean equality";
+ optkey = ["Equality";"Scheme"];
+ optread = (fun () -> !eq_flag) ;
+ optwrite = (fun b -> eq_flag := b) }
+
+let is_eq_flag () = !eq_flag && Flags.version_strictly_greater Flags.V8_2
+
+let eq_dec_flag = ref false
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "automatic declaration of decidable equality";
+ optkey = ["Decidable";"Equality";"Schemes"];
+ optread = (fun () -> !eq_dec_flag) ;
+ optwrite = (fun b -> eq_dec_flag := b) }
+
+let rewriting_flag = ref false
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname ="automatic declaration of rewriting schemes for equality types";
+ optkey = ["Rewriting";"Schemes"];
+ optread = (fun () -> !rewriting_flag) ;
+ optwrite = (fun b -> rewriting_flag := b) }
+
+(* Util *)
+
+let define id internal ctx c t =
+ let f = declare_constant ~internal in
+ let kn = f id
+ (DefinitionEntry
+ { const_entry_body = c;
+ const_entry_secctx = None;
+ const_entry_type = t;
+ const_entry_polymorphic = Flags.is_universe_polymorphism ();
+ const_entry_universes = snd (Evd.universe_context ctx);
+ const_entry_opaque = false;
+ const_entry_inline_code = false;
+ const_entry_feedback = None;
+ },
+ Decl_kinds.IsDefinition Scheme) in
+ definition_message id;
+ kn
+
+(* Boolean equality *)
+
+let declare_beq_scheme_gen internal names kn =
+ ignore (define_mutual_scheme beq_scheme_kind internal names kn)
+
+let alarm what internal msg =
+ let debug = false in
+ match internal with
+ | UserAutomaticRequest
+ | InternalTacticRequest ->
+ (if debug then
+ Feedback.msg_debug
+ (hov 0 msg ++ fnl () ++ what ++ str " not defined.")); None
+ | _ -> Some msg
+
+let try_declare_scheme what f internal names kn =
+ try f internal names kn
+ with e ->
+ let e = CErrors.push e in
+ let msg = match fst e with
+ | ParameterWithoutEquality cst ->
+ alarm what internal
+ (str "Boolean equality not found for parameter " ++ Printer.pr_global cst ++
+ str".")
+ | InductiveWithProduct ->
+ alarm what internal
+ (str "Unable to decide equality of functional arguments.")
+ | InductiveWithSort ->
+ alarm what internal
+ (str "Unable to decide equality of type arguments.")
+ | NonSingletonProp ind ->
+ alarm what internal
+ (str "Cannot extract computational content from proposition " ++
+ quote (Printer.pr_inductive (Global.env()) ind) ++ str ".")
+ | EqNotFound (ind',ind) ->
+ alarm what internal
+ (str "Boolean equality on " ++
+ quote (Printer.pr_inductive (Global.env()) ind') ++
+ strbrk " is missing.")
+ | UndefinedCst s ->
+ alarm what internal
+ (strbrk "Required constant " ++ str s ++ str " undefined.")
+ | AlreadyDeclared msg ->
+ alarm what internal (msg ++ str ".")
+ | DecidabilityMutualNotSupported ->
+ alarm what internal
+ (str "Decidability lemma for mutual inductive types not supported.")
+ | EqUnknown s ->
+ alarm what internal
+ (str "Found unsupported " ++ str s ++ str " while building Boolean equality.")
+ | NoDecidabilityCoInductive ->
+ alarm what internal
+ (str "Scheme Equality is only for inductive types.")
+ | e when CErrors.noncritical e ->
+ alarm what internal
+ (str "Unexpected error during scheme creation: " ++ CErrors.print e)
+ | _ -> iraise e
+ in
+ match msg with
+ | None -> ()
+ | Some msg -> iraise (UserError (None, msg), snd e)
+
+let beq_scheme_msg mind =
+ let mib = Global.lookup_mind mind in
+ (* TODO: mutual inductive case *)
+ str "Boolean equality on " ++
+ pr_enum (fun ind -> quote (Printer.pr_inductive (Global.env()) ind))
+ (List.init (Array.length mib.mind_packets) (fun i -> (mind,i)))
+
+let declare_beq_scheme_with l kn =
+ try_declare_scheme (beq_scheme_msg kn) declare_beq_scheme_gen UserIndividualRequest l kn
+
+let try_declare_beq_scheme kn =
+ (* TODO: handle Fix, eventually handle
+ proof-irrelevance; improve decidability by depending on decidability
+ for the parameters rather than on the bl and lb properties *)
+ try_declare_scheme (beq_scheme_msg kn) declare_beq_scheme_gen UserAutomaticRequest [] kn
+
+let declare_beq_scheme = declare_beq_scheme_with []
+
+(* Case analysis schemes *)
+let declare_one_case_analysis_scheme ind =
+ let (mib,mip) = Global.lookup_inductive ind in
+ let kind = inductive_sort_family mip in
+ let dep =
+ if kind == InProp then case_scheme_kind_from_prop
+ else if not (Inductiveops.has_dependent_elim mib) then
+ case_scheme_kind_from_type
+ else case_dep_scheme_kind_from_type in
+ let kelim = elim_sorts (mib,mip) in
+ (* in case the inductive has a type elimination, generates only one
+ induction scheme, the other ones share the same code with the
+ apropriate type *)
+ if Sorts.List.mem InType kelim then
+ ignore (define_individual_scheme dep UserAutomaticRequest None ind)
+
+(* Induction/recursion schemes *)
+
+let kinds_from_prop =
+ [InType,rect_scheme_kind_from_prop;
+ InProp,ind_scheme_kind_from_prop;
+ InSet,rec_scheme_kind_from_prop]
+
+let kinds_from_type =
+ [InType,rect_dep_scheme_kind_from_type;
+ InProp,ind_dep_scheme_kind_from_type;
+ InSet,rec_dep_scheme_kind_from_type]
+
+let nondep_kinds_from_type =
+ [InType,rect_scheme_kind_from_type;
+ InProp,ind_scheme_kind_from_type;
+ InSet,rec_scheme_kind_from_type]
+
+let declare_one_induction_scheme ind =
+ let (mib,mip) = Global.lookup_inductive ind in
+ let kind = inductive_sort_family mip in
+ let from_prop = kind == InProp in
+ let depelim = Inductiveops.has_dependent_elim mib in
+ let kelim = elim_sorts (mib,mip) in
+ let elims =
+ List.map_filter (fun (sort,kind) ->
+ if Sorts.List.mem sort kelim then Some kind else None)
+ (if from_prop then kinds_from_prop
+ else if depelim then kinds_from_type
+ else nondep_kinds_from_type) in
+ List.iter (fun kind -> ignore (define_individual_scheme kind UserAutomaticRequest None ind))
+ elims
+
+let declare_induction_schemes kn =
+ let mib = Global.lookup_mind kn in
+ if mib.mind_finite <> Decl_kinds.CoFinite then begin
+ for i = 0 to Array.length mib.mind_packets - 1 do
+ declare_one_induction_scheme (kn,i);
+ done;
+ end
+
+(* Decidable equality *)
+
+let declare_eq_decidability_gen internal names kn =
+ let mib = Global.lookup_mind kn in
+ if mib.mind_finite <> Decl_kinds.CoFinite then
+ ignore (define_mutual_scheme eq_dec_scheme_kind internal names kn)
+
+let eq_dec_scheme_msg ind = (* TODO: mutual inductive case *)
+ str "Decidable equality on " ++ quote (Printer.pr_inductive (Global.env()) ind)
+
+let declare_eq_decidability_scheme_with l kn =
+ try_declare_scheme (eq_dec_scheme_msg (kn,0))
+ declare_eq_decidability_gen UserIndividualRequest l kn
+
+let try_declare_eq_decidability kn =
+ try_declare_scheme (eq_dec_scheme_msg (kn,0))
+ declare_eq_decidability_gen UserAutomaticRequest [] kn
+
+let declare_eq_decidability = declare_eq_decidability_scheme_with []
+
+let ignore_error f x =
+ try ignore (f x) with e when CErrors.noncritical e -> ()
+
+let declare_rewriting_schemes ind =
+ if Hipattern.is_inductive_equality ind then begin
+ ignore (define_individual_scheme rew_r2l_scheme_kind UserAutomaticRequest None ind);
+ ignore (define_individual_scheme rew_r2l_dep_scheme_kind UserAutomaticRequest None ind);
+ ignore (define_individual_scheme rew_r2l_forward_dep_scheme_kind
+ UserAutomaticRequest None ind);
+ (* These ones expect the equality to be symmetric; the first one also *)
+ (* needs eq *)
+ ignore_error (define_individual_scheme rew_l2r_scheme_kind UserAutomaticRequest None) ind;
+ ignore_error
+ (define_individual_scheme rew_l2r_dep_scheme_kind UserAutomaticRequest None) ind;
+ ignore_error
+ (define_individual_scheme rew_l2r_forward_dep_scheme_kind UserAutomaticRequest None) ind
+ end
+
+let warn_cannot_build_congruence =
+ CWarnings.create ~name:"cannot-build-congruence" ~category:"schemes"
+ (fun () ->
+ strbrk "Cannot build congruence scheme because eq is not found")
+
+let declare_congr_scheme ind =
+ if Hipattern.is_equality_type (mkInd ind) then begin
+ if
+ try Coqlib.check_required_library Coqlib.logic_module_name; true
+ with e when CErrors.noncritical e -> false
+ then
+ ignore (define_individual_scheme congr_scheme_kind UserAutomaticRequest None ind)
+ else
+ warn_cannot_build_congruence ()
+ end
+
+let declare_sym_scheme ind =
+ if Hipattern.is_inductive_equality ind then
+ (* Expect the equality to be symmetric *)
+ ignore_error (define_individual_scheme sym_scheme_kind UserAutomaticRequest None) ind
+
+(* Scheme command *)
+
+let smart_global_inductive y = smart_global_inductive y
+let rec split_scheme l =
+ let env = Global.env() in
+ match l with
+ | [] -> [],[]
+ | (Some id,t)::q -> let l1,l2 = split_scheme q in
+ ( match t with
+ | InductionScheme (x,y,z) -> ((id,x,smart_global_inductive y,z)::l1),l2
+ | CaseScheme (x,y,z) -> ((id,x,smart_global_inductive y,z)::l1),l2
+ | EqualityScheme x -> l1,((Some id,smart_global_inductive x)::l2)
+ )
+(*
+ if no name has been provided, we build one from the types of the ind
+requested
+*)
+ | (None,t)::q ->
+ let l1,l2 = split_scheme q in
+ let names inds recs isdep y z =
+ let ind = smart_global_inductive y in
+ let sort_of_ind = inductive_sort_family (snd (lookup_mind_specif env ind)) in
+ let z' = interp_elimination_sort z in
+ let suffix = (
+ match sort_of_ind with
+ | InProp ->
+ if isdep then (match z' with
+ | InProp -> inds ^ "_dep"
+ | InSet -> recs ^ "_dep"
+ | InType -> recs ^ "t_dep")
+ else ( match z' with
+ | InProp -> inds
+ | InSet -> recs
+ | InType -> recs ^ "t" )
+ | _ ->
+ if isdep then (match z' with
+ | InProp -> inds
+ | InSet -> recs
+ | InType -> recs ^ "t" )
+ else (match z' with
+ | InProp -> inds ^ "_nodep"
+ | InSet -> recs ^ "_nodep"
+ | InType -> recs ^ "t_nodep")
+ ) in
+ let newid = add_suffix (basename_of_global (IndRef ind)) suffix in
+ let newref = (Loc.ghost,newid) in
+ ((newref,isdep,ind,z)::l1),l2
+ in
+ match t with
+ | CaseScheme (x,y,z) -> names "_case" "_case" x y z
+ | InductionScheme (x,y,z) -> names "_ind" "_rec" x y z
+ | EqualityScheme x -> l1,((None,smart_global_inductive x)::l2)
+
+
+let do_mutual_induction_scheme lnamedepindsort =
+ let lrecnames = List.map (fun ((_,f),_,_,_) -> f) lnamedepindsort
+ and env0 = Global.env() in
+ let sigma, lrecspec, _ =
+ List.fold_right
+ (fun (_,dep,ind,sort) (evd, l, inst) ->
+ let evd, indu, inst =
+ match inst with
+ | None ->
+ let _, ctx = Global.type_of_global_in_context env0 (IndRef ind) in
+ let ctxs = Univ.ContextSet.of_context ctx in
+ let evd = Evd.from_ctx (Evd.evar_universe_context_of ctxs) in
+ let u = Univ.UContext.instance ctx in
+ evd, (ind,u), Some u
+ | Some ui -> evd, (ind, ui), inst
+ in
+ (evd, (indu,dep,interp_elimination_sort sort) :: l, inst))
+ lnamedepindsort (Evd.from_env env0,[],None)
+ in
+ let sigma, listdecl = Indrec.build_mutual_induction_scheme env0 sigma lrecspec in
+ let declare decl fi lrecref =
+ let decltype = Retyping.get_type_of env0 sigma decl in
+ let proof_output = Future.from_val ((decl,Univ.ContextSet.empty),Safe_typing.empty_private_constants) in
+ let cst = define fi UserIndividualRequest sigma proof_output (Some decltype) in
+ ConstRef cst :: lrecref
+ in
+ let _ = List.fold_right2 declare listdecl lrecnames [] in
+ fixpoint_message None lrecnames
+
+let get_common_underlying_mutual_inductive = function
+ | [] -> assert false
+ | (id,(mind,i as ind))::l as all ->
+ match List.filter (fun (_,(mind',_)) -> not (eq_mind mind mind')) l with
+ | (_,ind')::_ ->
+ raise (RecursionSchemeError (NotMutualInScheme (ind,ind')))
+ | [] ->
+ if not (List.distinct_f Int.compare (List.map snd (List.map snd all)))
+ then error "A type occurs twice";
+ mind,
+ List.map_filter
+ (function (Some id,(_,i)) -> Some (i,snd id) | (None,_) -> None) all
+
+let do_scheme l =
+ let ischeme,escheme = split_scheme l in
+(* we want 1 kind of scheme at a time so we check if the user
+tried to declare different schemes at once *)
+ if not (List.is_empty ischeme) && not (List.is_empty escheme)
+ then
+ error "Do not declare equality and induction scheme at the same time."
+ else (
+ if not (List.is_empty ischeme) then do_mutual_induction_scheme ischeme
+ else
+ let mind,l = get_common_underlying_mutual_inductive escheme in
+ declare_beq_scheme_with l mind;
+ declare_eq_decidability_scheme_with l mind
+ )
+
+(**********************************************************************)
+(* Combined scheme *)
+(* Matthieu Sozeau, Dec 2006 *)
+
+let list_split_rev_at index l =
+ let rec aux i acc = function
+ hd :: tl when Int.equal i index -> acc, tl
+ | hd :: tl -> aux (succ i) (hd :: acc) tl
+ | [] -> failwith "List.split_when: Invalid argument"
+ in aux 0 [] l
+
+let fold_left' f = function
+ [] -> invalid_arg "fold_left'"
+ | hd :: tl -> List.fold_left f hd tl
+
+let build_combined_scheme env schemes =
+ let defs = List.map (fun cst -> (* FIXME *)
+ let evd, c = Evd.fresh_constant_instance env (Evd.from_env env) cst in
+ (c, Typeops.type_of_constant_in env c)) schemes in
+(* let nschemes = List.length schemes in *)
+ let find_inductive ty =
+ let (ctx, arity) = decompose_prod ty in
+ let (_, last) = List.hd ctx in
+ match kind_of_term last with
+ | App (ind, args) ->
+ let ind = destInd ind in
+ let (_,spec) = Inductive.lookup_mind_specif env (fst ind) in
+ ctx, ind, spec.mind_nrealargs
+ | _ -> ctx, destInd last, 0
+ in
+ let (c, t) = List.hd defs in
+ let ctx, ind, nargs = find_inductive t in
+ (* Number of clauses, including the predicates quantification *)
+ let prods = nb_prod t - (nargs + 1) in
+ let coqand = Coqlib.build_coq_and () and coqconj = Coqlib.build_coq_conj () in
+ let relargs = rel_vect 0 prods in
+ let concls = List.rev_map
+ (fun (cst, t) -> (* FIXME *)
+ mkApp(mkConstU cst, relargs),
+ snd (decompose_prod_n prods t)) defs in
+ let concl_bod, concl_typ =
+ fold_left'
+ (fun (accb, acct) (cst, x) ->
+ mkApp (coqconj, [| x; acct; cst; accb |]),
+ mkApp (coqand, [| x; acct |])) concls
+ in
+ let ctx, _ =
+ list_split_rev_at prods
+ (List.rev_map (fun (x, y) -> LocalAssum (x, y)) ctx) in
+ let typ = it_mkProd_wo_LetIn concl_typ ctx in
+ let body = it_mkLambda_or_LetIn concl_bod ctx in
+ (body, typ)
+
+let do_combined_scheme name schemes =
+ let csts =
+ List.map (fun x ->
+ let refe = Ident x in
+ let qualid = qualid_of_reference refe in
+ try Nametab.locate_constant (snd qualid)
+ with Not_found -> error ((string_of_qualid (snd qualid))^" is not declared."))
+ schemes
+ in
+ let body,typ = build_combined_scheme (Global.env ()) csts in
+ let proof_output = Future.from_val ((body,Univ.ContextSet.empty),Safe_typing.empty_private_constants) in
+ ignore (define (snd name) UserIndividualRequest Evd.empty proof_output (Some typ));
+ fixpoint_message None [snd name]
+
+(**********************************************************************)
+
+let map_inductive_block f kn n = for i=0 to n-1 do f (kn,i) done
+
+let declare_default_schemes kn =
+ let mib = Global.lookup_mind kn in
+ let n = Array.length mib.mind_packets in
+ if !elim_flag && (mib.mind_finite <> BiFinite || !bifinite_elim_flag)
+ && mib.mind_typing_flags.check_guarded then
+ declare_induction_schemes kn;
+ if !case_flag then map_inductive_block declare_one_case_analysis_scheme kn n;
+ if is_eq_flag() then try_declare_beq_scheme kn;
+ if !eq_dec_flag then try_declare_eq_decidability kn;
+ if !rewriting_flag then map_inductive_block declare_congr_scheme kn n;
+ if !rewriting_flag then map_inductive_block declare_sym_scheme kn n;
+ if !rewriting_flag then map_inductive_block declare_rewriting_schemes kn n
diff --git a/vernac/indschemes.mli b/vernac/indschemes.mli
new file mode 100644
index 000000000..e5d79fd51
--- /dev/null
+++ b/vernac/indschemes.mli
@@ -0,0 +1,49 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Loc
+open Names
+open Term
+open Environ
+open Vernacexpr
+open Misctypes
+
+(** See also Auto_ind_decl, Indrec, Eqscheme, Ind_tables, ... *)
+
+(** Build and register the boolean equalities associated to an inductive type *)
+
+val declare_beq_scheme : mutual_inductive -> unit
+
+val declare_eq_decidability : mutual_inductive -> unit
+
+(** Build and register a congruence scheme for an equality-like inductive type *)
+
+val declare_congr_scheme : inductive -> unit
+
+(** Build and register rewriting schemes for an equality-like inductive type *)
+
+val declare_rewriting_schemes : inductive -> unit
+
+(** Mutual Minimality/Induction scheme *)
+
+val do_mutual_induction_scheme :
+ (Id.t located * bool * inductive * glob_sort) list -> unit
+
+(** Main calls to interpret the Scheme command *)
+
+val do_scheme : (Id.t located option * scheme) list -> unit
+
+(** Combine a list of schemes into a conjunction of them *)
+
+val build_combined_scheme : env -> constant list -> constr * types
+
+val do_combined_scheme : Id.t located -> Id.t located list -> unit
+
+(** Hook called at each inductive type definition *)
+
+val declare_default_schemes : mutual_inductive -> unit
diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml
new file mode 100644
index 000000000..55f33be39
--- /dev/null
+++ b/vernac/lemmas.ml
@@ -0,0 +1,557 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* Created by Hugo Herbelin from contents related to lemma proofs in
+ file command.ml, Aug 2009 *)
+
+open CErrors
+open Util
+open Flags
+open Pp
+open Names
+open Term
+open Declarations
+open Declareops
+open Entries
+open Environ
+open Nameops
+open Globnames
+open Decls
+open Decl_kinds
+open Declare
+open Pretyping
+open Termops
+open Namegen
+open Reductionops
+open Constrexpr
+open Constrintern
+open Impargs
+open Context.Rel.Declaration
+
+module RelDecl = Context.Rel.Declaration
+module NamedDecl = Context.Named.Declaration
+
+type 'a declaration_hook = Decl_kinds.locality -> Globnames.global_reference -> 'a
+let mk_hook hook = hook
+let call_hook fix_exn hook l c =
+ try hook l c
+ with e when CErrors.noncritical e ->
+ let e = CErrors.push e in
+ iraise (fix_exn e)
+
+(* Support for mutually proved theorems *)
+
+let retrieve_first_recthm = function
+ | VarRef id ->
+ (NamedDecl.get_value (Global.lookup_named id),variable_opacity id)
+ | ConstRef cst ->
+ let cb = Global.lookup_constant cst in
+ (Global.body_of_constant_body cb, is_opaque cb)
+ | _ -> assert false
+
+let adjust_guardness_conditions const = function
+ | [] -> const (* Not a recursive statement *)
+ | possible_indexes ->
+ (* Try all combinations... not optimal *)
+ let env = Global.env() in
+ { const with const_entry_body =
+ Future.chain ~greedy:true ~pure:true const.const_entry_body
+ (fun ((body, ctx), eff) ->
+ match kind_of_term body with
+ | Fix ((nv,0),(_,_,fixdefs as fixdecls)) ->
+(* let possible_indexes =
+ List.map2 (fun i c -> match i with Some i -> i | None ->
+ List.interval 0 (List.length ((lam_assum c))))
+ lemma_guard (Array.to_list fixdefs) in
+*)
+ let add c cb e =
+ let exists c e =
+ try ignore(Environ.lookup_constant c e); true
+ with Not_found -> false in
+ if exists c e then e else Environ.add_constant c cb e in
+ let env = List.fold_left (fun env { eff } ->
+ match eff with
+ | SEsubproof (c, cb,_) -> add c cb env
+ | SEscheme (l,_) ->
+ List.fold_left (fun e (_,c,cb,_) -> add c cb e) env l)
+ env (Safe_typing.side_effects_of_private_constants eff) in
+ let indexes =
+ search_guard Loc.ghost env
+ possible_indexes fixdecls in
+ (mkFix ((indexes,0),fixdecls), ctx), eff
+ | _ -> (body, ctx), eff) }
+
+let find_mutually_recursive_statements thms =
+ let n = List.length thms in
+ let inds = List.map (fun (id,(t,impls,annot)) ->
+ let (hyps,ccl) = decompose_prod_assum t in
+ let x = (id,(t,impls)) in
+ match annot with
+ (* Explicit fixpoint decreasing argument is given *)
+ | Some (Some (_,id),CStructRec) ->
+ let i,b,typ = lookup_rel_id id hyps in
+ (match kind_of_term t with
+ | Ind ((kn,_ as ind), u) when
+ let mind = Global.lookup_mind kn in
+ mind.mind_finite == Decl_kinds.Finite && Option.is_empty b ->
+ [ind,x,i],[]
+ | _ ->
+ error "Decreasing argument is not an inductive assumption.")
+ (* Unsupported cases *)
+ | Some (_,(CWfRec _|CMeasureRec _)) ->
+ error "Only structural decreasing is supported for mutual statements."
+ (* Cofixpoint or fixpoint w/o explicit decreasing argument *)
+ | None | Some (None, CStructRec) ->
+ let whnf_hyp_hds = map_rel_context_in_env
+ (fun env c -> fst (whd_all_stack env Evd.empty c))
+ (Global.env()) hyps in
+ let ind_hyps =
+ List.flatten (List.map_i (fun i decl ->
+ let t = RelDecl.get_type decl in
+ match kind_of_term t with
+ | Ind ((kn,_ as ind),u) when
+ let mind = Global.lookup_mind kn in
+ mind.mind_finite <> Decl_kinds.CoFinite && is_local_assum decl ->
+ [ind,x,i]
+ | _ ->
+ []) 0 (List.rev whnf_hyp_hds)) in
+ let ind_ccl =
+ let cclenv = push_rel_context hyps (Global.env()) in
+ let whnf_ccl,_ = whd_all_stack cclenv Evd.empty ccl in
+ match kind_of_term whnf_ccl with
+ | Ind ((kn,_ as ind),u) when
+ let mind = Global.lookup_mind kn in
+ Int.equal mind.mind_ntypes n && mind.mind_finite == Decl_kinds.CoFinite ->
+ [ind,x,0]
+ | _ ->
+ [] in
+ ind_hyps,ind_ccl) thms in
+ let inds_hyps,ind_ccls = List.split inds in
+ let of_same_mutind ((kn,_),_,_) = function ((kn',_),_,_) -> eq_mind kn kn' in
+ (* Check if all conclusions are coinductive in the same type *)
+ (* (degenerated cartesian product since there is at most one coind ccl) *)
+ let same_indccl =
+ List.cartesians_filter (fun hyp oks ->
+ if List.for_all (of_same_mutind hyp) oks
+ then Some (hyp::oks) else None) [] ind_ccls in
+ let ordered_same_indccl =
+ List.filter (List.for_all_i (fun i ((kn,j),_,_) -> Int.equal i j) 0) same_indccl in
+ (* Check if some hypotheses are inductive in the same type *)
+ let common_same_indhyp =
+ List.cartesians_filter (fun hyp oks ->
+ if List.for_all (of_same_mutind hyp) oks
+ then Some (hyp::oks) else None) [] inds_hyps in
+ let ordered_inds,finite,guard =
+ match ordered_same_indccl, common_same_indhyp with
+ | indccl::rest, _ ->
+ assert (List.is_empty rest);
+ (* One occ. of common coind ccls and no common inductive hyps *)
+ if not (List.is_empty common_same_indhyp) then
+ if_verbose Feedback.msg_info (str "Assuming mutual coinductive statements.");
+ flush_all ();
+ indccl, true, []
+ | [], _::_ ->
+ let () = match same_indccl with
+ | ind :: _ ->
+ if List.distinct_f ind_ord (List.map pi1 ind)
+ then
+ if_verbose Feedback.msg_info
+ (strbrk
+ ("Coinductive statements do not follow the order of "^
+ "definition, assuming the proof to be by induction."));
+ flush_all ()
+ | _ -> ()
+ in
+ let possible_guards = List.map (List.map pi3) inds_hyps in
+ (* assume the largest indices as possible *)
+ List.last common_same_indhyp, false, possible_guards
+ | _, [] ->
+ error
+ ("Cannot find common (mutual) inductive premises or coinductive" ^
+ " conclusions in the statements.")
+ in
+ (finite,guard,None), ordered_inds
+
+let look_for_possibly_mutual_statements = function
+ | [id,(t,impls,None)] ->
+ (* One non recursively proved theorem *)
+ None,[id,(t,impls)],None
+ | _::_ as thms ->
+ (* More than one statement and/or an explicit decreasing mark: *)
+ (* we look for a common inductive hyp or a common coinductive conclusion *)
+ let recguard,ordered_inds = find_mutually_recursive_statements thms in
+ let thms = List.map pi2 ordered_inds in
+ Some recguard,thms, Some (List.map (fun (_,_,i) -> succ i) ordered_inds)
+ | [] -> anomaly (Pp.str "Empty list of theorems.")
+
+(* Saving a goal *)
+
+let save ?export_seff id const cstrs pl do_guard (locality,poly,kind) hook =
+ let fix_exn = Future.fix_exn_of const.Entries.const_entry_body in
+ try
+ let const = adjust_guardness_conditions const do_guard in
+ let k = Kindops.logical_kind_of_goal_kind kind in
+ let l,r = match locality with
+ | Discharge when Lib.sections_are_opened () ->
+ let c = SectionLocalDef const in
+ let _ = declare_variable id (Lib.cwd(), c, k) in
+ (Local, VarRef id)
+ | Local | Global | Discharge ->
+ let local = match locality with
+ | Local | Discharge -> true
+ | Global -> false
+ in
+ let kn =
+ declare_constant ?export_seff id ~local (DefinitionEntry const, k) in
+ (locality, ConstRef kn) in
+ definition_message id;
+ Option.iter (Universes.register_universe_binders r) pl;
+ call_hook (fun exn -> exn) hook l r
+ with e when CErrors.noncritical e ->
+ let e = CErrors.push e in
+ iraise (fix_exn e)
+
+let default_thm_id = Id.of_string "Unnamed_thm"
+
+let compute_proof_name locality = function
+ | Some ((loc,id),pl) ->
+ (* We check existence here: it's a bit late at Qed time *)
+ if Nametab.exists_cci (Lib.make_path id) || is_section_variable id ||
+ locality == Global && Nametab.exists_cci (Lib.make_path_except_section id)
+ then
+ user_err ~loc (pr_id id ++ str " already exists.");
+ id, pl
+ | None ->
+ next_global_ident_away default_thm_id (Pfedit.get_all_proof_names ()), None
+
+let save_remaining_recthms (locality,p,kind) norm ctx body opaq i ((id,pl),(t_i,(_,imps))) =
+ let t_i = norm t_i in
+ match body with
+ | None ->
+ (match locality with
+ | Discharge ->
+ let impl = false in (* copy values from Vernacentries *)
+ let k = IsAssumption Conjectural in
+ let c = SectionLocalAssum ((t_i,ctx),p,impl) in
+ let _ = declare_variable id (Lib.cwd(),c,k) in
+ (Discharge, VarRef id,imps)
+ | Local | Global ->
+ let k = IsAssumption Conjectural in
+ let local = match locality with
+ | Local -> true
+ | Global -> false
+ | Discharge -> assert false
+ in
+ let ctx = Univ.ContextSet.to_context ctx in
+ let decl = (ParameterEntry (None,p,(t_i,ctx),None), k) in
+ let kn = declare_constant id ~local decl in
+ (locality,ConstRef kn,imps))
+ | Some body ->
+ let body = norm body in
+ let k = Kindops.logical_kind_of_goal_kind kind in
+ let rec body_i t = match kind_of_term t with
+ | Fix ((nv,0),decls) -> mkFix ((nv,i),decls)
+ | CoFix (0,decls) -> mkCoFix (i,decls)
+ | LetIn(na,t1,ty,t2) -> mkLetIn (na,t1,ty, body_i t2)
+ | Lambda(na,ty,t) -> mkLambda(na,ty,body_i t)
+ | App (t, args) -> mkApp (body_i t, args)
+ | _ -> anomaly Pp.(str "Not a proof by induction: " ++ Printer.pr_constr body) in
+ let body_i = body_i body in
+ match locality with
+ | Discharge ->
+ let const = definition_entry ~types:t_i ~opaque:opaq ~poly:p
+ ~univs:(Univ.ContextSet.to_context ctx) body_i in
+ let c = SectionLocalDef const in
+ let _ = declare_variable id (Lib.cwd(), c, k) in
+ (Discharge,VarRef id,imps)
+ | Local | Global ->
+ let ctx = Univ.ContextSet.to_context ctx in
+ let local = match locality with
+ | Local -> true
+ | Global -> false
+ | Discharge -> assert false
+ in
+ let const =
+ Declare.definition_entry ~types:t_i ~poly:p ~univs:ctx ~opaque:opaq body_i
+ in
+ let kn = declare_constant id ~local (DefinitionEntry const, k) in
+ (locality,ConstRef kn,imps)
+
+let save_hook = ref ignore
+let set_save_hook f = save_hook := f
+
+let save_named ?export_seff proof =
+ let id,const,(cstrs,pl),do_guard,persistence,hook = proof in
+ save ?export_seff id const cstrs pl do_guard persistence hook
+
+let check_anonymity id save_ident =
+ if not (String.equal (atompart_of_id id) (Id.to_string (default_thm_id))) then
+ error "This command can only be used for unnamed theorem."
+
+let save_anonymous ?export_seff proof save_ident =
+ let id,const,(cstrs,pl),do_guard,persistence,hook = proof in
+ check_anonymity id save_ident;
+ save ?export_seff save_ident const cstrs pl do_guard persistence hook
+
+let save_anonymous_with_strength ?export_seff proof kind save_ident =
+ let id,const,(cstrs,pl),do_guard,_,hook = proof in
+ check_anonymity id save_ident;
+ (* we consider that non opaque behaves as local for discharge *)
+ save ?export_seff save_ident const cstrs pl do_guard
+ (Global, const.const_entry_polymorphic, Proof kind) hook
+
+(* Admitted *)
+
+let warn_let_as_axiom =
+ CWarnings.create ~name:"let-as-axiom" ~category:"vernacular"
+ (fun id -> strbrk "Let definition" ++ spc () ++ pr_id id ++
+ spc () ++ strbrk "declared as an axiom.")
+
+let admit (id,k,e) pl hook () =
+ let kn = declare_constant id (ParameterEntry e, IsAssumption Conjectural) in
+ let () = match k with
+ | Global, _, _ -> ()
+ | Local, _, _ | Discharge, _, _ -> warn_let_as_axiom id
+ in
+ let () = assumption_message id in
+ Option.iter (Universes.register_universe_binders (ConstRef kn)) pl;
+ call_hook (fun exn -> exn) hook Global (ConstRef kn)
+
+(* Starting a goal *)
+
+let start_hook = ref ignore
+let set_start_hook = (:=) start_hook
+
+
+let get_proof proof do_guard hook opacity =
+ let (id,(const,univs,persistence)) =
+ Pfedit.cook_this_proof proof
+ in
+ id,{const with const_entry_opaque = opacity},univs,do_guard,persistence,hook
+
+let check_exist =
+ List.iter (fun (loc,id) ->
+ if not (Nametab.exists_cci (Lib.make_path id)) then
+ user_err ~loc (pr_id id ++ str " does not exist.")
+ )
+
+let universe_proof_terminator compute_guard hook =
+ let open Proof_global in
+ make_terminator begin function
+ | Admitted (id,k,pe,(ctx,pl)) ->
+ admit (id,k,pe) pl (hook (Some ctx)) ();
+ Feedback.feedback Feedback.AddedAxiom
+ | Proved (opaque,idopt,proof) ->
+ let is_opaque, export_seff, exports = match opaque with
+ | Vernacexpr.Transparent -> false, true, []
+ | Vernacexpr.Opaque None -> true, false, []
+ | Vernacexpr.Opaque (Some l) -> true, true, l in
+ let proof = get_proof proof compute_guard
+ (hook (Some (fst proof.Proof_global.universes))) is_opaque in
+ begin match idopt with
+ | None -> save_named ~export_seff proof
+ | Some ((_,id),None) -> save_anonymous ~export_seff proof id
+ | Some ((_,id),Some kind) ->
+ save_anonymous_with_strength ~export_seff proof kind id
+ end;
+ check_exist exports
+ end
+
+let standard_proof_terminator compute_guard hook =
+ universe_proof_terminator compute_guard (fun _ -> hook)
+
+let start_proof id ?pl kind sigma ?terminator ?sign c ?init_tac ?(compute_guard=[]) hook =
+ let terminator = match terminator with
+ | None -> standard_proof_terminator compute_guard hook
+ | Some terminator -> terminator compute_guard hook
+ in
+ let sign =
+ match sign with
+ | Some sign -> sign
+ | None -> initialize_named_context_for_proof ()
+ in
+ !start_hook c;
+ Pfedit.start_proof id ?pl kind sigma sign c ?init_tac terminator
+
+let start_proof_univs id ?pl kind sigma ?terminator ?sign c ?init_tac ?(compute_guard=[]) hook =
+ let terminator = match terminator with
+ | None -> universe_proof_terminator compute_guard hook
+ | Some terminator -> terminator compute_guard hook
+ in
+ let sign =
+ match sign with
+ | Some sign -> sign
+ | None -> initialize_named_context_for_proof ()
+ in
+ !start_hook c;
+ Pfedit.start_proof id ?pl kind sigma sign c ?init_tac terminator
+
+let rec_tac_initializer finite guard thms snl =
+ if finite then
+ match List.map (fun ((id,_),(t,_)) -> (id,t)) thms with
+ | (id,_)::l -> Tactics.mutual_cofix id l 0
+ | _ -> assert false
+ else
+ (* nl is dummy: it will be recomputed at Qed-time *)
+ let nl = match snl with
+ | None -> List.map succ (List.map List.last guard)
+ | Some nl -> nl
+ in match List.map2 (fun ((id,_),(t,_)) n -> (id,n,t)) thms nl with
+ | (id,n,_)::l -> Tactics.mutual_fix id n l 0
+ | _ -> assert false
+
+let start_proof_with_initialization kind ctx recguard thms snl hook =
+ let intro_tac (_, (_, (ids, _))) =
+ Tacticals.New.tclMAP (function
+ | Name id -> Tactics.intro_mustbe_force id
+ | Anonymous -> Tactics.intro) (List.rev ids) in
+ let init_tac,guard = match recguard with
+ | Some (finite,guard,init_tac) ->
+ let rec_tac = rec_tac_initializer finite guard thms snl in
+ Some (match init_tac with
+ | None ->
+ if Flags.is_auto_intros () then
+ Tacticals.New.tclTHENS rec_tac (List.map intro_tac thms)
+ else
+ rec_tac
+ | Some tacl ->
+ Tacticals.New.tclTHENS rec_tac
+ (if Flags.is_auto_intros () then
+ List.map2 (fun tac thm -> Tacticals.New.tclTHEN tac (intro_tac thm)) tacl thms
+ else
+ tacl)),guard
+ | None ->
+ let () = match thms with [_] -> () | _ -> assert false in
+ (if Flags.is_auto_intros () then Some (intro_tac (List.hd thms)) else None), [] in
+ match thms with
+ | [] -> anomaly (Pp.str "No proof to start")
+ | ((id,pl),(t,(_,imps)))::other_thms ->
+ let hook ctx strength ref =
+ let ctx = match ctx with
+ | None -> Evd.empty_evar_universe_context
+ | Some ctx -> ctx
+ in
+ let other_thms_data =
+ if List.is_empty other_thms then [] else
+ (* there are several theorems defined mutually *)
+ let body,opaq = retrieve_first_recthm ref in
+ let subst = Evd.evar_universe_context_subst ctx in
+ let norm c = Universes.subst_opt_univs_constr subst c in
+ let ctx = UState.context_set (*FIXME*) ctx in
+ let body = Option.map norm body in
+ List.map_i (save_remaining_recthms kind norm ctx body opaq) 1 other_thms in
+ let thms_data = (strength,ref,imps)::other_thms_data in
+ List.iter (fun (strength,ref,imps) ->
+ maybe_declare_manual_implicits false ref imps;
+ call_hook (fun exn -> exn) hook strength ref) thms_data in
+ start_proof_univs id ?pl kind ctx t ?init_tac (fun ctx -> mk_hook (hook ctx)) ~compute_guard:guard
+
+let start_proof_com ?inference_hook kind thms hook =
+ let env0 = Global.env () in
+ let levels = Option.map snd (fst (List.hd thms)) in
+ let evdref = ref (match levels with
+ | None -> Evd.from_env env0
+ | Some l -> Evd.from_ctx (Evd.make_evar_universe_context env0 l))
+ in
+ let thms = List.map (fun (sopt,(bl,t,guard)) ->
+ let impls, ((env, ctx), imps) = interp_context_evars env0 evdref bl in
+ let t', imps' = interp_type_evars_impls ~impls env evdref t in
+ let flags = all_and_fail_flags in
+ let flags = { flags with use_hook = inference_hook } in
+ evdref := solve_remaining_evars flags env !evdref (Evd.empty,!evdref);
+ let ids = List.map RelDecl.get_name ctx in
+ (compute_proof_name (pi1 kind) sopt,
+ (nf_evar !evdref (it_mkProd_or_LetIn t' ctx),
+ (ids, imps @ lift_implicits (List.length ids) imps'),
+ guard)))
+ thms in
+ let recguard,thms,snl = look_for_possibly_mutual_statements thms in
+ let evd, nf = Evarutil.nf_evars_and_universes !evdref in
+ let thms = List.map (fun (n, (t, info)) -> (n, (nf t, info))) thms in
+ let () =
+ match levels with
+ | None -> ()
+ | Some l -> ignore (Evd.universe_context evd ?names:l)
+ in
+ let evd =
+ if pi2 kind then evd
+ else (* We fix the variables to ensure they won't be lowered to Set *)
+ Evd.fix_undefined_variables evd
+ in
+ start_proof_with_initialization kind evd recguard thms snl hook
+
+
+(* Saving a proof *)
+
+let keep_admitted_vars = ref true
+
+let _ =
+ let open Goptions in
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "keep section variables in admitted proofs";
+ optkey = ["Keep"; "Admitted"; "Variables"];
+ optread = (fun () -> !keep_admitted_vars);
+ optwrite = (fun b -> keep_admitted_vars := b) }
+
+let save_proof ?proof = function
+ | Vernacexpr.Admitted ->
+ let pe =
+ let open Proof_global in
+ match proof with
+ | Some ({ id; entries; persistence = k; universes }, _) ->
+ if List.length entries <> 1 then
+ error "Admitted does not support multiple statements";
+ let { const_entry_secctx; const_entry_type } = List.hd entries in
+ if const_entry_type = None then
+ error "Admitted requires an explicit statement";
+ let typ = Option.get const_entry_type in
+ let ctx = Evd.evar_context_universe_context (fst universes) in
+ let sec_vars = if !keep_admitted_vars then const_entry_secctx else None in
+ Admitted(id, k, (sec_vars, pi2 k, (typ, ctx), None), universes)
+ | None ->
+ let pftree = Pfedit.get_pftreestate () in
+ let id, k, typ = Pfedit.current_proof_statement () in
+ let universes = Proof.initial_euctx pftree in
+ (* This will warn if the proof is complete *)
+ let pproofs, _univs =
+ Proof_global.return_proof ~allow_partial:true () in
+ let sec_vars =
+ if not !keep_admitted_vars then None
+ else match Pfedit.get_used_variables(), pproofs with
+ | Some _ as x, _ -> x
+ | None, (pproof, _) :: _ ->
+ let env = Global.env () in
+ let ids_typ = Environ.global_vars_set env typ in
+ let ids_def = Environ.global_vars_set env pproof in
+ Some (Environ.keep_hyps env (Idset.union ids_typ ids_def))
+ | _ -> None in
+ let names = Pfedit.get_universe_binders () in
+ let evd = Evd.from_ctx universes in
+ let binders, ctx = Evd.universe_context ?names evd in
+ Admitted(id,k,(sec_vars, pi2 k, (typ, ctx), None),
+ (universes, Some binders))
+ in
+ Proof_global.apply_terminator (Proof_global.get_terminator ()) pe
+ | Vernacexpr.Proved (is_opaque,idopt) ->
+ let (proof_obj,terminator) =
+ match proof with
+ | None ->
+ Proof_global.close_proof ~keep_body_ucst_separate:false (fun x -> x)
+ | Some proof -> proof
+ in
+ (* if the proof is given explicitly, nothing has to be deleted *)
+ if Option.is_empty proof then Pfedit.delete_current_proof ();
+ Proof_global.(apply_terminator terminator (Proved (is_opaque,idopt,proof_obj)))
+
+(* Miscellaneous *)
+
+let get_current_context () =
+ Pfedit.get_current_context ()
+
diff --git a/vernac/lemmas.mli b/vernac/lemmas.mli
new file mode 100644
index 000000000..39c089be9
--- /dev/null
+++ b/vernac/lemmas.mli
@@ -0,0 +1,69 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open Term
+open Decl_kinds
+open Pfedit
+
+type 'a declaration_hook
+val mk_hook :
+ (Decl_kinds.locality -> Globnames.global_reference -> 'a) -> 'a declaration_hook
+
+val call_hook :
+ Future.fix_exn -> 'a declaration_hook -> Decl_kinds.locality -> Globnames.global_reference -> 'a
+
+(** A hook start_proof calls on the type of the definition being started *)
+val set_start_hook : (types -> unit) -> unit
+
+val start_proof : Id.t -> ?pl:universe_binders -> goal_kind -> Evd.evar_map ->
+ ?terminator:(lemma_possible_guards -> unit declaration_hook -> Proof_global.proof_terminator) ->
+ ?sign:Environ.named_context_val -> types ->
+ ?init_tac:unit Proofview.tactic -> ?compute_guard:lemma_possible_guards ->
+ unit declaration_hook -> unit
+
+val start_proof_univs : Id.t -> ?pl:universe_binders -> goal_kind -> Evd.evar_map ->
+ ?terminator:(lemma_possible_guards -> (Evd.evar_universe_context option -> unit declaration_hook) -> Proof_global.proof_terminator) ->
+ ?sign:Environ.named_context_val -> types ->
+ ?init_tac:unit Proofview.tactic -> ?compute_guard:lemma_possible_guards ->
+ (Evd.evar_universe_context option -> unit declaration_hook) -> unit
+
+val start_proof_com :
+ ?inference_hook:Pretyping.inference_hook ->
+ goal_kind -> Vernacexpr.proof_expr list ->
+ unit declaration_hook -> unit
+
+val start_proof_with_initialization :
+ goal_kind -> Evd.evar_map ->
+ (bool * lemma_possible_guards * unit Proofview.tactic list option) option ->
+ ((Id.t * universe_binders option) *
+ (types * (Name.t list * Impargs.manual_explicitation list))) list
+ -> int list option -> unit declaration_hook -> unit
+
+val universe_proof_terminator :
+ Proof_global.lemma_possible_guards ->
+ (Evd.evar_universe_context option -> unit declaration_hook) ->
+ Proof_global.proof_terminator
+
+val standard_proof_terminator :
+ Proof_global.lemma_possible_guards -> unit declaration_hook ->
+ Proof_global.proof_terminator
+
+(** {6 ... } *)
+
+(** A hook the next three functions pass to cook_proof *)
+val set_save_hook : (Proof.proof -> unit) -> unit
+
+val save_proof : ?proof:Proof_global.closed_proof -> Vernacexpr.proof_end -> unit
+
+
+(** [get_current_context ()] returns the evar context and env of the
+ current open proof if any, otherwise returns the empty evar context
+ and the current global env *)
+
+val get_current_context : unit -> Evd.evar_map * Environ.env
diff --git a/vernac/locality.ml b/vernac/locality.ml
new file mode 100644
index 000000000..03640676e
--- /dev/null
+++ b/vernac/locality.ml
@@ -0,0 +1,107 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Pp
+
+(** * Managing locality *)
+
+let local_of_bool = function
+ | true -> Decl_kinds.Local
+ | false -> Decl_kinds.Global
+
+let check_locality locality_flag =
+ match locality_flag with
+ | Some b ->
+ let s = if b then "Local" else "Global" in
+ CErrors.user_err ~hdr:"Locality.check_locality"
+ (str "This command does not support the \"" ++ str s ++ str "\" prefix.")
+ | None -> ()
+
+(** Extracting the locality flag *)
+
+(* Commands which supported an inlined Local flag *)
+
+let warn_deprecated_local_syntax =
+ CWarnings.create ~name:"deprecated-local-syntax" ~category:"deprecated"
+ (fun () ->
+ Pp.strbrk "Deprecated syntax: use \"Local\" as a prefix.")
+
+let enforce_locality_full locality_flag local =
+ let local =
+ match locality_flag with
+ | Some false when local ->
+ CErrors.error "Cannot be simultaneously Local and Global."
+ | Some true when local ->
+ CErrors.error "Use only prefix \"Local\"."
+ | None ->
+ if local then begin
+ warn_deprecated_local_syntax ();
+ Some true
+ end else
+ None
+ | Some b -> Some b in
+ local
+
+(** Positioning locality for commands supporting discharging and export
+ outside of modules *)
+
+(* For commands whose default is to discharge and export:
+ Global is the default and is neutral;
+ Local in a section deactivates discharge,
+ Local not in a section deactivates export *)
+let make_non_locality = function Some false -> false | _ -> true
+
+let make_locality = function Some true -> true | _ -> false
+
+let enforce_locality locality_flag local =
+ make_locality (enforce_locality_full locality_flag local)
+
+let enforce_locality_exp locality_flag local =
+ match locality_flag, local with
+ | None, Some local -> local
+ | Some b, None -> local_of_bool b
+ | None, None -> Decl_kinds.Global
+ | Some _, Some _ -> CErrors.error "Local non allowed in this case"
+
+(* For commands whose default is to not discharge but to export:
+ Global in sections forces discharge, Global not in section is the default;
+ Local in sections is the default, Local not in section forces non-export *)
+
+let make_section_locality =
+ function Some b -> b | None -> Lib.sections_are_opened ()
+
+let enforce_section_locality locality_flag local =
+ make_section_locality (enforce_locality_full locality_flag local)
+
+(** Positioning locality for commands supporting export but not discharge *)
+
+(* For commands whose default is to export (if not in section):
+ Global in sections is forbidden, Global not in section is neutral;
+ Local in sections is the default, Local not in section forces non-export *)
+
+let make_module_locality = function
+ | Some false ->
+ if Lib.sections_are_opened () then
+ CErrors.error
+ "This command does not support the Global option in sections.";
+ false
+ | Some true -> true
+ | None -> false
+
+let enforce_module_locality locality_flag local =
+ make_module_locality (enforce_locality_full locality_flag local)
+
+module LocalityFixme = struct
+ let locality = ref None
+ let set l = locality := l
+ let consume () =
+ let l = !locality in
+ locality := None;
+ l
+ let assert_consumed () = check_locality !locality
+end
diff --git a/vernac/locality.mli b/vernac/locality.mli
new file mode 100644
index 000000000..2ec392eef
--- /dev/null
+++ b/vernac/locality.mli
@@ -0,0 +1,51 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** * Managing locality *)
+
+(** Commands which supported an inlined Local flag *)
+
+val enforce_locality_full : bool option -> bool -> bool option
+
+(** * Positioning locality for commands supporting discharging and export
+ outside of modules *)
+
+(** For commands whose default is to discharge and export:
+ Global is the default and is neutral;
+ Local in a section deactivates discharge,
+ Local not in a section deactivates export *)
+
+val make_locality : bool option -> bool
+val make_non_locality : bool option -> bool
+val enforce_locality : bool option -> bool -> bool
+val enforce_locality_exp :
+ bool option -> Decl_kinds.locality option -> Decl_kinds.locality
+
+(** For commands whose default is to not discharge but to export:
+ Global in sections forces discharge, Global not in section is the default;
+ Local in sections is the default, Local not in section forces non-export *)
+
+val make_section_locality : bool option -> bool
+val enforce_section_locality : bool option -> bool -> bool
+
+(** * Positioning locality for commands supporting export but not discharge *)
+
+(** For commands whose default is to export (if not in section):
+ Global in sections is forbidden, Global not in section is neutral;
+ Local in sections is the default, Local not in section forces non-export *)
+
+val make_module_locality : bool option -> bool
+val enforce_module_locality : bool option -> bool -> bool
+
+(* This is the old imperative interface that is still used for
+ * VernacExtend vernaculars. Time permitting this could be trashed too *)
+module LocalityFixme : sig
+ val set : bool option -> unit
+ val consume : unit -> bool option
+ val assert_consumed : unit -> unit
+end
diff --git a/vernac/metasyntax.ml b/vernac/metasyntax.ml
new file mode 100644
index 000000000..0aaf6afd7
--- /dev/null
+++ b/vernac/metasyntax.ml
@@ -0,0 +1,1469 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Pp
+open Flags
+open CErrors
+open Util
+open Names
+open Constrexpr
+open Constrexpr_ops
+open Notation_term
+open Notation_ops
+open Ppextend
+open Extend
+open Libobject
+open Constrintern
+open Vernacexpr
+open Libnames
+open Tok
+open Notation
+open Nameops
+
+(**********************************************************************)
+(* Tokens *)
+
+let cache_token (_,s) = CLexer.add_keyword s
+
+let inToken : string -> obj =
+ declare_object {(default_object "TOKEN") with
+ open_function = (fun i o -> if Int.equal i 1 then cache_token o);
+ cache_function = cache_token;
+ subst_function = Libobject.ident_subst_function;
+ classify_function = (fun o -> Substitute o)}
+
+let add_token_obj s = Lib.add_anonymous_leaf (inToken s)
+
+(**********************************************************************)
+(* Printing grammar entries *)
+
+let entry_buf = Buffer.create 64
+
+type any_entry = AnyEntry : 'a Pcoq.Gram.entry -> any_entry
+
+let grammars : any_entry list String.Map.t ref = ref String.Map.empty
+
+let register_grammar name grams =
+ grammars := String.Map.add name grams !grammars
+
+let pr_entry e =
+ let () = Buffer.clear entry_buf in
+ let ft = Format.formatter_of_buffer entry_buf in
+ let () = Pcoq.Gram.entry_print ft e in
+ str (Buffer.contents entry_buf)
+
+let pr_registered_grammar name =
+ let gram = try Some (String.Map.find name !grammars) with Not_found -> None in
+ match gram with
+ | None -> error "Unknown or unprintable grammar entry."
+ | Some entries ->
+ let pr_one (AnyEntry e) =
+ str "Entry " ++ str (Pcoq.Gram.Entry.name e) ++ str " is" ++ fnl () ++
+ pr_entry e
+ in
+ prlist pr_one entries
+
+let pr_grammar = function
+ | "constr" | "operconstr" | "binder_constr" ->
+ str "Entry constr is" ++ fnl () ++
+ pr_entry Pcoq.Constr.constr ++
+ str "and lconstr is" ++ fnl () ++
+ pr_entry Pcoq.Constr.lconstr ++
+ str "where binder_constr is" ++ fnl () ++
+ pr_entry Pcoq.Constr.binder_constr ++
+ str "and operconstr is" ++ fnl () ++
+ pr_entry Pcoq.Constr.operconstr
+ | "pattern" ->
+ pr_entry Pcoq.Constr.pattern
+ | "vernac" ->
+ str "Entry vernac is" ++ fnl () ++
+ pr_entry Pcoq.Vernac_.vernac ++
+ str "Entry command is" ++ fnl () ++
+ pr_entry Pcoq.Vernac_.command ++
+ str "Entry syntax is" ++ fnl () ++
+ pr_entry Pcoq.Vernac_.syntax ++
+ str "Entry gallina is" ++ fnl () ++
+ pr_entry Pcoq.Vernac_.gallina ++
+ str "Entry gallina_ext is" ++ fnl () ++
+ pr_entry Pcoq.Vernac_.gallina_ext
+ | name -> pr_registered_grammar name
+
+(**********************************************************************)
+(* Parse a format (every terminal starting with a letter or a single
+ quote (except a single quote alone) must be quoted) *)
+
+let parse_format ((loc, str) : lstring) =
+ let str = " "^str in
+ let l = String.length str in
+ let push_token a = function
+ | cur::l -> (a::cur)::l
+ | [] -> [[a]] in
+ let push_white n l =
+ if Int.equal n 0 then l else push_token (UnpTerminal (String.make n ' ')) l in
+ let close_box i b = function
+ | a::(_::_ as l) -> push_token (UnpBox (b,a)) l
+ | _ -> error "Non terminated box in format." in
+ let close_quotation i =
+ if i < String.length str && str.[i] == '\'' && (Int.equal (i+1) l || str.[i+1] == ' ')
+ then i+1
+ else error "Incorrectly terminated quoted expression." in
+ let rec spaces n i =
+ if i < String.length str && str.[i] == ' ' then spaces (n+1) (i+1)
+ else n in
+ let rec nonspaces quoted n i =
+ if i < String.length str && str.[i] != ' ' then
+ if str.[i] == '\'' && quoted &&
+ (i+1 >= String.length str || str.[i+1] == ' ')
+ then if Int.equal n 0 then error "Empty quoted token." else n
+ else nonspaces quoted (n+1) (i+1)
+ else
+ if quoted then error "Spaces are not allowed in (quoted) symbols."
+ else n in
+ let rec parse_non_format i =
+ let n = nonspaces false 0 i in
+ push_token (UnpTerminal (String.sub str i n)) (parse_token (i+n))
+ and parse_quoted n i =
+ if i < String.length str then match str.[i] with
+ (* Parse " // " *)
+ | '/' when i <= String.length str && str.[i+1] == '/' ->
+ (* We forget the useless n spaces... *)
+ push_token (UnpCut PpFnl)
+ (parse_token (close_quotation (i+2)))
+ (* Parse " .. / .. " *)
+ | '/' when i <= String.length str ->
+ let p = spaces 0 (i+1) in
+ push_token (UnpCut (PpBrk (n,p)))
+ (parse_token (close_quotation (i+p+1)))
+ | c ->
+ (* The spaces are real spaces *)
+ push_white n (match c with
+ | '[' ->
+ if i <= String.length str then match str.[i+1] with
+ (* Parse " [h .. ", *)
+ | 'h' when i+1 <= String.length str && str.[i+2] == 'v' ->
+ (parse_box (fun n -> PpHVB n) (i+3))
+ (* Parse " [v .. ", *)
+ | 'v' ->
+ parse_box (fun n -> PpVB n) (i+2)
+ (* Parse " [ .. ", *)
+ | ' ' | '\'' ->
+ parse_box (fun n -> PpHOVB n) (i+1)
+ | _ -> error "\"v\", \"hv\", \" \" expected after \"[\" in format."
+ else error "\"v\", \"hv\" or \" \" expected after \"[\" in format."
+ (* Parse "]" *)
+ | ']' ->
+ ([] :: parse_token (close_quotation (i+1)))
+ (* Parse a non formatting token *)
+ | c ->
+ let n = nonspaces true 0 i in
+ push_token (UnpTerminal (String.sub str (i-1) (n+2)))
+ (parse_token (close_quotation (i+n))))
+ else
+ if Int.equal n 0 then []
+ else error "Ending spaces non part of a format annotation."
+ and parse_box box i =
+ let n = spaces 0 i in
+ close_box i (box n) (parse_token (close_quotation (i+n)))
+ and parse_token i =
+ let n = spaces 0 i in
+ let i = i+n in
+ if i < l then match str.[i] with
+ (* Parse a ' *)
+ | '\'' when i+1 >= String.length str || str.[i+1] == ' ' ->
+ push_white (n-1) (push_token (UnpTerminal "'") (parse_token (i+1)))
+ (* Parse the beginning of a quoted expression *)
+ | '\'' ->
+ parse_quoted (n-1) (i+1)
+ (* Otherwise *)
+ | _ ->
+ push_white (n-1) (parse_non_format i)
+ else push_white n [[]]
+ in
+ try
+ if not (String.is_empty str) then match parse_token 0 with
+ | [l] -> l
+ | _ -> error "Box closed without being opened in format."
+ else
+ error "Empty format."
+ with reraise ->
+ let (e, info) = CErrors.push reraise in
+ let info = Loc.add_loc info loc in
+ iraise (e, info)
+
+(***********************)
+(* Analyzing notations *)
+
+type symbol_token = WhiteSpace of int | String of string
+
+let split_notation_string str =
+ let push_token beg i l =
+ if Int.equal beg i then l else
+ let s = String.sub str beg (i - beg) in
+ String s :: l
+ in
+ let push_whitespace beg i l =
+ if Int.equal beg i then l else WhiteSpace (i-beg) :: l
+ in
+ let rec loop beg i =
+ if i < String.length str then
+ if str.[i] == ' ' then
+ push_token beg i (loop_on_whitespace (i+1) (i+1))
+ else
+ loop beg (i+1)
+ else
+ push_token beg i []
+ and loop_on_whitespace beg i =
+ if i < String.length str then
+ if str.[i] != ' ' then
+ push_whitespace beg i (loop i (i+1))
+ else
+ loop_on_whitespace beg (i+1)
+ else
+ push_whitespace beg i []
+ in
+ loop 0 0
+
+(* Interpret notations with a recursive component *)
+
+let out_nt = function NonTerminal x -> x | _ -> assert false
+
+let msg_expected_form_of_recursive_notation =
+ "In the notation, the special symbol \"..\" must occur in\na configuration of the form \"x symbs .. symbs y\"."
+
+let rec find_pattern nt xl = function
+ | Break n as x :: l, Break n' :: l' when Int.equal n n' ->
+ find_pattern nt (x::xl) (l,l')
+ | Terminal s as x :: l, Terminal s' :: l' when String.equal s s' ->
+ find_pattern nt (x::xl) (l,l')
+ | [], NonTerminal x' :: l' ->
+ (out_nt nt,x',List.rev xl),l'
+ | _, Break s :: _ | Break s :: _, _ ->
+ error ("A break occurs on one side of \"..\" but not on the other side.")
+ | _, Terminal s :: _ | Terminal s :: _, _ ->
+ user_err ~hdr:"Metasyntax.find_pattern"
+ (str "The token \"" ++ str s ++ str "\" occurs on one side of \"..\" but not on the other side.")
+ | _, [] ->
+ error msg_expected_form_of_recursive_notation
+ | ((SProdList _ | NonTerminal _) :: _), _ | _, (SProdList _ :: _) ->
+ anomaly (Pp.str "Only Terminal or Break expected on left, non-SProdList on right")
+
+let rec interp_list_parser hd = function
+ | [] -> [], List.rev hd
+ | NonTerminal id :: tl when Id.equal id ldots_var ->
+ if List.is_empty hd then error msg_expected_form_of_recursive_notation;
+ let hd = List.rev hd in
+ let ((x,y,sl),tl') = find_pattern (List.hd hd) [] (List.tl hd,tl) in
+ let xyl,tl'' = interp_list_parser [] tl' in
+ (* We remember each pair of variable denoting a recursive part to *)
+ (* remove the second copy of it afterwards *)
+ (x,y)::xyl, SProdList (x,sl) :: tl''
+ | (Terminal _ | Break _) as s :: tl ->
+ if List.is_empty hd then
+ let yl,tl' = interp_list_parser [] tl in
+ yl, s :: tl'
+ else
+ interp_list_parser (s::hd) tl
+ | NonTerminal _ as x :: tl ->
+ let xyl,tl' = interp_list_parser [x] tl in
+ xyl, List.rev_append hd tl'
+ | SProdList _ :: _ -> anomaly (Pp.str "Unexpected SProdList in interp_list_parser")
+
+
+(* Find non-terminal tokens of notation *)
+
+(* To protect alphabetic tokens and quotes from being seen as variables *)
+let quote_notation_token x =
+ let n = String.length x in
+ let norm = CLexer.is_ident x in
+ if (n > 0 && norm) || (n > 2 && x.[0] == '\'') then "'"^x^"'"
+ else x
+
+let rec raw_analyze_notation_tokens = function
+ | [] -> []
+ | String ".." :: sl -> NonTerminal ldots_var :: raw_analyze_notation_tokens sl
+ | String "_" :: _ -> error "_ must be quoted."
+ | String x :: sl when CLexer.is_ident x ->
+ NonTerminal (Names.Id.of_string x) :: raw_analyze_notation_tokens sl
+ | String s :: sl ->
+ Terminal (String.drop_simple_quotes s) :: raw_analyze_notation_tokens sl
+ | WhiteSpace n :: sl ->
+ Break n :: raw_analyze_notation_tokens sl
+
+let is_numeral symbs =
+ match List.filter (function Break _ -> false | _ -> true) symbs with
+ | ([Terminal "-"; Terminal x] | [Terminal x]) ->
+ (try let _ = Bigint.of_string x in true with Failure _ -> false)
+ | _ ->
+ false
+
+let rec get_notation_vars = function
+ | [] -> []
+ | NonTerminal id :: sl ->
+ let vars = get_notation_vars sl in
+ if Id.equal id ldots_var then vars else
+ if Id.List.mem id vars then
+ user_err ~hdr:"Metasyntax.get_notation_vars"
+ (str "Variable " ++ pr_id id ++ str " occurs more than once.")
+ else
+ id::vars
+ | (Terminal _ | Break _) :: sl -> get_notation_vars sl
+ | SProdList _ :: _ -> assert false
+
+let analyze_notation_tokens l =
+ let l = raw_analyze_notation_tokens l in
+ let vars = get_notation_vars l in
+ let recvars,l = interp_list_parser [] l in
+ recvars, List.subtract Id.equal vars (List.map snd recvars), l
+
+let error_not_same_scope x y =
+ user_err ~hdr:"Metasyntax.error_not_name_scope"
+ (str "Variables " ++ pr_id x ++ str " and " ++ pr_id y ++ str " must be in the same scope.")
+
+(**********************************************************************)
+(* Build pretty-printing rules *)
+
+let prec_assoc = function
+ | RightA -> (L,E)
+ | LeftA -> (E,L)
+ | NonA -> (L,L)
+
+let precedence_of_entry_type from = function
+ | ETConstr (NumLevel n,BorderProd (_,None)) -> n, Prec n
+ | ETConstr (NumLevel n,BorderProd (b,Some a)) ->
+ n, let (lp,rp) = prec_assoc a in if b == Left then lp else rp
+ | ETConstr (NumLevel n,InternalProd) -> n, Prec n
+ | ETConstr (NextLevel,_) -> from, L
+ | _ -> 0, E (* ?? *)
+
+(* Some breaking examples *)
+(* "x = y" : "x /1 = y" (breaks before any symbol) *)
+(* "x =S y" : "x /1 =S /1 y" (protect from confusion; each side for symmetry)*)
+(* "+ {" : "+ {" may breaks reversibility without space but oth. not elegant *)
+(* "x y" : "x spc y" *)
+(* "{ x } + { y }" : "{ x } / + { y }" *)
+(* "< x , y > { z , t }" : "< x , / y > / { z , / t }" *)
+
+let starts_with_left_bracket s =
+ let l = String.length s in not (Int.equal l 0) &&
+ (s.[0] == '{' || s.[0] == '[' || s.[0] == '(')
+
+let ends_with_right_bracket s =
+ let l = String.length s in not (Int.equal l 0) &&
+ (s.[l-1] == '}' || s.[l-1] == ']' || s.[l-1] == ')')
+
+let is_left_bracket s =
+ starts_with_left_bracket s && not (ends_with_right_bracket s)
+
+let is_right_bracket s =
+ not (starts_with_left_bracket s) && ends_with_right_bracket s
+
+let is_comma s =
+ let l = String.length s in not (Int.equal l 0) &&
+ (s.[0] == ',' || s.[0] == ';')
+
+let is_operator s =
+ let l = String.length s in not (Int.equal l 0) &&
+ (s.[0] == '+' || s.[0] == '*' || s.[0] == '=' ||
+ s.[0] == '-' || s.[0] == '/' || s.[0] == '<' || s.[0] == '>' ||
+ s.[0] == '@' || s.[0] == '\\' || s.[0] == '&' || s.[0] == '~' || s.[0] == '$')
+
+let is_non_terminal = function
+ | NonTerminal _ | SProdList _ -> true
+ | _ -> false
+
+let is_next_non_terminal = function
+| [] -> false
+| pr :: _ -> is_non_terminal pr
+
+let is_next_terminal = function Terminal _ :: _ -> true | _ -> false
+
+let is_next_break = function Break _ :: _ -> true | _ -> false
+
+let add_break n l = UnpCut (PpBrk(n,0)) :: l
+
+let add_break_if_none n = function
+ | ((UnpCut (PpBrk _) :: _) | []) as l -> l
+ | l -> UnpCut (PpBrk(n,0)) :: l
+
+let check_open_binder isopen sl m =
+ let pr_token = function
+ | Terminal s -> str s
+ | Break n -> str "␣"
+ | _ -> assert false
+ in
+ if isopen && not (List.is_empty sl) then
+ user_err (str "as " ++ pr_id m ++
+ str " is a non-closed binder, no such \"" ++
+ prlist_with_sep spc pr_token sl
+ ++ strbrk "\" is allowed to occur.")
+
+(* Heuristics for building default printing rules *)
+
+let index_id id l = List.index Id.equal id l
+
+let make_hunks etyps symbols from =
+ let vars,typs = List.split etyps in
+ let rec make = function
+ | NonTerminal m :: prods ->
+ let i = index_id m vars in
+ let _,prec = precedence_of_entry_type from (List.nth typs (i-1)) in
+ let u = UnpMetaVar (i,prec) in
+ if is_next_non_terminal prods then
+ u :: add_break_if_none 1 (make prods)
+ else
+ u :: make_with_space prods
+ | Terminal s :: prods when List.exists is_non_terminal prods ->
+ if (is_comma s || is_operator s) then
+ (* Always a breakable space after comma or separator *)
+ UnpTerminal s :: add_break_if_none 1 (make prods)
+ else if is_right_bracket s && is_next_terminal prods then
+ (* Always no space after right bracked, but possibly a break *)
+ UnpTerminal s :: add_break_if_none 0 (make prods)
+ else if is_left_bracket s && is_next_non_terminal prods then
+ UnpTerminal s :: make prods
+ else if not (is_next_break prods) then
+ (* Add rigid space, no break, unless user asked for something *)
+ UnpTerminal (s^" ") :: make prods
+ else
+ (* Rely on user spaces *)
+ UnpTerminal s :: make prods
+
+ | Terminal s :: prods ->
+ (* Separate but do not cut a trailing sequence of terminal *)
+ (match prods with
+ | Terminal _ :: _ -> UnpTerminal (s^" ") :: make prods
+ | _ -> UnpTerminal s :: make prods)
+
+ | Break n :: prods ->
+ add_break n (make prods)
+
+ | SProdList (m,sl) :: prods ->
+ let i = index_id m vars in
+ let typ = List.nth typs (i-1) in
+ let _,prec = precedence_of_entry_type from typ in
+ let sl' =
+ (* If no separator: add a break *)
+ if List.is_empty sl then add_break 1 []
+ (* We add NonTerminal for simulation but remove it afterwards *)
+ else snd (List.sep_last (make (sl@[NonTerminal m]))) in
+ let hunk = match typ with
+ | ETConstr _ -> UnpListMetaVar (i,prec,sl')
+ | ETBinder isopen ->
+ check_open_binder isopen sl m;
+ UnpBinderListMetaVar (i,isopen,sl')
+ | _ -> assert false in
+ hunk :: make_with_space prods
+
+ | [] -> []
+
+ and make_with_space prods =
+ match prods with
+ | Terminal s' :: prods'->
+ if is_operator s' then
+ (* A rigid space before operator and a breakable after *)
+ UnpTerminal (" "^s') :: add_break_if_none 1 (make prods')
+ else if is_comma s' then
+ (* No space whatsoever before comma *)
+ make prods
+ else if is_right_bracket s' then
+ make prods
+ else
+ (* A breakable space between any other two terminals *)
+ add_break_if_none 1 (make prods)
+ | (NonTerminal _ | SProdList _) :: _ ->
+ (* A breakable space before a non-terminal *)
+ add_break_if_none 1 (make prods)
+ | Break _ :: _ ->
+ (* Rely on user wish *)
+ make prods
+ | [] -> []
+
+ in make symbols
+
+(* Build default printing rules from explicit format *)
+
+let error_format () = error "The format does not match the notation."
+
+let rec split_format_at_ldots hd = function
+ | UnpTerminal s :: fmt when String.equal s (Id.to_string ldots_var) -> List.rev hd, fmt
+ | u :: fmt ->
+ check_no_ldots_in_box u;
+ split_format_at_ldots (u::hd) fmt
+ | [] -> raise Exit
+
+and check_no_ldots_in_box = function
+ | UnpBox (_,fmt) ->
+ (try
+ let _ = split_format_at_ldots [] fmt in
+ error ("The special symbol \"..\" must occur at the same formatting depth than the variables of which it is the ellipse.")
+ with Exit -> ())
+ | _ -> ()
+
+let skip_var_in_recursive_format = function
+ | UnpTerminal _ :: sl (* skip first var *) ->
+ (* To do, though not so important: check that the names match
+ the names in the notation *)
+ sl
+ | _ -> error_format ()
+
+let read_recursive_format sl fmt =
+ let get_head fmt =
+ let sl = skip_var_in_recursive_format fmt in
+ try split_format_at_ldots [] sl with Exit -> error_format () in
+ let rec get_tail = function
+ | a :: sepfmt, b :: fmt when Pervasives.(=) a b -> get_tail (sepfmt, fmt) (* FIXME *)
+ | [], tail -> skip_var_in_recursive_format tail
+ | _ -> error "The format is not the same on the right and left hand side of the special token \"..\"." in
+ let slfmt, fmt = get_head fmt in
+ slfmt, get_tail (slfmt, fmt)
+
+let hunks_of_format (from,(vars,typs)) symfmt =
+ let rec aux = function
+ | symbs, (UnpTerminal s' as u) :: fmt
+ when String.equal s' (String.make (String.length s') ' ') ->
+ let symbs, l = aux (symbs,fmt) in symbs, u :: l
+ | Terminal s :: symbs, (UnpTerminal s') :: fmt
+ when String.equal s (String.drop_simple_quotes s') ->
+ let symbs, l = aux (symbs,fmt) in symbs, UnpTerminal s :: l
+ | NonTerminal s :: symbs, UnpTerminal s' :: fmt when Id.equal s (Id.of_string s') ->
+ let i = index_id s vars in
+ let _,prec = precedence_of_entry_type from (List.nth typs (i-1)) in
+ let symbs, l = aux (symbs,fmt) in symbs, UnpMetaVar (i,prec) :: l
+ | symbs, UnpBox (a,b) :: fmt ->
+ let symbs', b' = aux (symbs,b) in
+ let symbs', l = aux (symbs',fmt) in
+ symbs', UnpBox (a,b') :: l
+ | symbs, (UnpCut _ as u) :: fmt ->
+ let symbs, l = aux (symbs,fmt) in symbs, u :: l
+ | SProdList (m,sl) :: symbs, fmt ->
+ let i = index_id m vars in
+ let typ = List.nth typs (i-1) in
+ let _,prec = precedence_of_entry_type from typ in
+ let slfmt,fmt = read_recursive_format sl fmt in
+ let sl, slfmt = aux (sl,slfmt) in
+ if not (List.is_empty sl) then error_format ();
+ let symbs, l = aux (symbs,fmt) in
+ let hunk = match typ with
+ | ETConstr _ -> UnpListMetaVar (i,prec,slfmt)
+ | ETBinder isopen ->
+ check_open_binder isopen sl m;
+ UnpBinderListMetaVar (i,isopen,slfmt)
+ | _ -> assert false in
+ symbs, hunk :: l
+ | symbs, [] -> symbs, []
+ | _, _ -> error_format ()
+ in
+ match aux symfmt with
+ | [], l -> l
+ | _ -> error_format ()
+
+(**********************************************************************)
+(* Build parsing rules *)
+
+let assoc_of_type n (_,typ) = precedence_of_entry_type n typ
+
+let is_not_small_constr = function
+ ETConstr _ -> true
+ | ETOther("constr","binder_constr") -> true
+ | _ -> false
+
+let rec define_keywords_aux = function
+ | GramConstrNonTerminal(e,Some _) as n1 :: GramConstrTerminal(IDENT k) :: l
+ when is_not_small_constr e ->
+ Flags.if_verbose Feedback.msg_info (str "Identifier '" ++ str k ++ str "' now a keyword");
+ CLexer.add_keyword k;
+ n1 :: GramConstrTerminal(KEYWORD k) :: define_keywords_aux l
+ | n :: l -> n :: define_keywords_aux l
+ | [] -> []
+
+ (* Ensure that IDENT articulation terminal symbols are keywords *)
+let define_keywords = function
+ | GramConstrTerminal(IDENT k)::l ->
+ Flags.if_verbose Feedback.msg_info (str "Identifier '" ++ str k ++ str "' now a keyword");
+ CLexer.add_keyword k;
+ GramConstrTerminal(KEYWORD k) :: define_keywords_aux l
+ | l -> define_keywords_aux l
+
+let distribute a ll = List.map (fun l -> a @ l) ll
+
+ (* Expand LIST1(t,sep) into the combination of t and t;sep;LIST1(t,sep)
+ as many times as expected in [n] argument *)
+let rec expand_list_rule typ tkl x n i hds ll =
+ if Int.equal i n then
+ let hds =
+ GramConstrListMark (n,true) :: hds
+ @ [GramConstrNonTerminal (ETConstrList (typ,tkl), Some x)] in
+ distribute hds ll
+ else
+ let camlp4_message_name = Some (add_suffix x ("_"^string_of_int n)) in
+ let main = GramConstrNonTerminal (ETConstr typ, camlp4_message_name) in
+ let tks = List.map (fun x -> GramConstrTerminal x) tkl in
+ distribute (GramConstrListMark (i+1,false) :: hds @ [main]) ll @
+ expand_list_rule typ tkl x n (i+1) (main :: tks @ hds) ll
+
+let make_production etyps symbols =
+ let prod =
+ List.fold_right
+ (fun t ll -> match t with
+ | NonTerminal m ->
+ let typ = List.assoc m etyps in
+ distribute [GramConstrNonTerminal (typ, Some m)] ll
+ | Terminal s ->
+ distribute [GramConstrTerminal (CLexer.terminal s)] ll
+ | Break _ ->
+ ll
+ | SProdList (x,sl) ->
+ let tkl = List.flatten
+ (List.map (function Terminal s -> [CLexer.terminal s]
+ | Break _ -> []
+ | _ -> anomaly (Pp.str "Found a non terminal token in recursive notation separator")) sl) in
+ match List.assoc x etyps with
+ | ETConstr typ -> expand_list_rule typ tkl x 1 0 [] ll
+ | ETBinder o ->
+ distribute
+ [GramConstrNonTerminal (ETBinderList (o,tkl), Some x)] ll
+ | _ ->
+ error "Components of recursive patterns in notation must be terms or binders.")
+ symbols [[]] in
+ List.map define_keywords prod
+
+let rec find_symbols c_current c_next c_last = function
+ | [] -> []
+ | NonTerminal id :: sl ->
+ let prec = if not (List.is_empty sl) then c_current else c_last in
+ (id, prec) :: (find_symbols c_next c_next c_last sl)
+ | Terminal s :: sl -> find_symbols c_next c_next c_last sl
+ | Break n :: sl -> find_symbols c_current c_next c_last sl
+ | SProdList (x,_) :: sl' ->
+ (x,c_next)::(find_symbols c_next c_next c_last sl')
+
+let border = function
+ | (_,ETConstr(_,BorderProd (_,a))) :: _ -> a
+ | _ -> None
+
+let recompute_assoc typs =
+ match border typs, border (List.rev typs) with
+ | Some LeftA, Some RightA -> assert false
+ | Some LeftA, _ -> Some LeftA
+ | _, Some RightA -> Some RightA
+ | _ -> None
+
+(**************************************************************************)
+(* Registration of syntax extensions (parsing/printing, no interpretation)*)
+
+let pr_arg_level from = function
+ | (n,L) when Int.equal n from -> str "at next level"
+ | (n,E) -> str "at level " ++ int n
+ | (n,L) -> str "at level below " ++ int n
+ | (n,Prec m) when Int.equal m n -> str "at level " ++ int n
+ | (n,_) -> str "Unknown level"
+
+let pr_level ntn (from,args) =
+ str "at level " ++ int from ++ spc () ++ str "with arguments" ++ spc() ++
+ prlist_with_sep pr_comma (pr_arg_level from) args
+
+let error_incompatible_level ntn oldprec prec =
+ user_err
+ (str "Notation " ++ str ntn ++ str " is already defined" ++ spc() ++
+ pr_level ntn oldprec ++
+ spc() ++ str "while it is now required to be" ++ spc() ++
+ pr_level ntn prec ++ str ".")
+
+type syntax_extension = {
+ synext_level : Notation.level;
+ synext_notation : notation;
+ synext_notgram : notation_grammar;
+ synext_unparsing : unparsing list;
+ synext_extra : (string * string) list;
+ synext_compat : Flags.compat_version option;
+}
+
+let is_active_compat = function
+| None -> true
+| Some v -> 0 <= Flags.version_compare v !Flags.compat_version
+
+type syntax_extension_obj = locality_flag * syntax_extension list
+
+let cache_one_syntax_extension se =
+ let ntn = se.synext_notation in
+ let prec = se.synext_level in
+ let onlyprint = se.synext_notgram.notgram_onlyprinting in
+ try
+ let oldprec = Notation.level_of_notation ntn in
+ if not (Notation.level_eq prec oldprec) then error_incompatible_level ntn oldprec prec
+ with Not_found ->
+ if is_active_compat se.synext_compat then begin
+ (* Reserve the notation level *)
+ Notation.declare_notation_level ntn prec;
+ (* Declare the parsing rule *)
+ if not onlyprint then Egramcoq.extend_constr_grammar prec se.synext_notgram;
+ (* Declare the notation rule *)
+ Notation.declare_notation_rule ntn
+ ~extra:se.synext_extra (se.synext_unparsing, fst prec) se.synext_notgram
+ end
+
+let cache_syntax_extension (_, (_, sy)) =
+ List.iter cache_one_syntax_extension sy
+
+let subst_parsing_rule subst x = x
+
+let subst_printing_rule subst x = x
+
+let subst_syntax_extension (subst, (local, sy)) =
+ let map sy = { sy with
+ synext_notgram = subst_parsing_rule subst sy.synext_notgram;
+ synext_unparsing = subst_printing_rule subst sy.synext_unparsing;
+ } in
+ (local, List.map map sy)
+
+let classify_syntax_definition (local, _ as o) =
+ if local then Dispose else Substitute o
+
+let inSyntaxExtension : syntax_extension_obj -> obj =
+ declare_object {(default_object "SYNTAX-EXTENSION") with
+ open_function = (fun i o -> if Int.equal i 1 then cache_syntax_extension o);
+ cache_function = cache_syntax_extension;
+ subst_function = subst_syntax_extension;
+ classify_function = classify_syntax_definition}
+
+(**************************************************************************)
+(* Precedences *)
+
+(* Interpreting user-provided modifiers *)
+
+(* XXX: We could move this to the parser itself *)
+module NotationMods = struct
+
+type notation_modifier = {
+ assoc : gram_assoc option;
+ level : int option;
+ etyps : (Id.t * simple_constr_prod_entry_key) list;
+
+ (* common to syn_data below *)
+ only_parsing : bool;
+ only_printing : bool;
+ compat : compat_version option;
+ format : string Loc.located option;
+ extra : (string * string) list;
+}
+
+let default = {
+ assoc = None;
+ level = None;
+ etyps = [];
+ only_parsing = false;
+ only_printing = false;
+ compat = None;
+ format = None;
+ extra = [];
+}
+
+end
+
+let interp_modifiers modl = let open NotationMods in
+ let rec interp acc = function
+ | [] -> acc
+ | SetEntryType (s,typ) :: l ->
+ let id = Id.of_string s in
+ if Id.List.mem_assoc id acc.etyps then
+ user_err ~hdr:"Metasyntax.interp_modifiers"
+ (str s ++ str " is already assigned to an entry or constr level.");
+ interp { acc with etyps = (id,typ) :: acc.etyps; } l
+ | SetItemLevel ([],n) :: l ->
+ interp acc l
+ | SetItemLevel (s::idl,n) :: l ->
+ let id = Id.of_string s in
+ if Id.List.mem_assoc id acc.etyps then
+ user_err ~hdr:"Metasyntax.interp_modifiers"
+ (str s ++ str " is already assigned to an entry or constr level.");
+ let typ = ETConstr (n,()) in
+ interp { acc with etyps = (id,typ)::acc.etyps; } (SetItemLevel (idl,n)::l)
+ | SetLevel n :: l ->
+
+ interp { acc with level = Some n; } l
+ | SetAssoc a :: l ->
+ if not (Option.is_empty acc.assoc) then error "An associativity is given more than once.";
+ interp { acc with assoc = Some a; } l
+ | SetOnlyParsing :: l ->
+ interp { acc with only_parsing = true; } l
+ | SetOnlyPrinting :: l ->
+ interp { acc with only_printing = true; } l
+ | SetCompatVersion v :: l ->
+ interp { acc with compat = Some v; } l
+ | SetFormat ("text",s) :: l ->
+ if not (Option.is_empty acc.format) then error "A format is given more than once.";
+ interp { acc with format = Some s; } l
+ | SetFormat (k,(_,s)) :: l ->
+ interp { acc with extra = (k,s)::acc.extra; } l
+ in interp default modl
+
+let check_infix_modifiers modifiers =
+ let t = (interp_modifiers modifiers).NotationMods.etyps in
+ if not (List.is_empty t) then
+ error "Explicit entry level or type unexpected in infix notation."
+
+let check_useless_entry_types recvars mainvars etyps =
+ let vars = let (l1,l2) = List.split recvars in l1@l2@mainvars in
+ match List.filter (fun (x,etyp) -> not (List.mem x vars)) etyps with
+ | (x,_)::_ -> user_err ~hdr:"Metasyntax.check_useless_entry_types"
+ (pr_id x ++ str " is unbound in the notation.")
+ | _ -> ()
+
+let not_a_syntax_modifier = function
+| SetOnlyParsing -> true
+| SetOnlyPrinting -> true
+| SetCompatVersion _ -> true
+| _ -> false
+
+let no_syntax_modifiers mods = List.for_all not_a_syntax_modifier mods
+
+let is_only_parsing mods =
+ let test = function SetOnlyParsing -> true | _ -> false in
+ List.exists test mods
+
+let is_only_printing mods =
+ let test = function SetOnlyPrinting -> true | _ -> false in
+ List.exists test mods
+
+let get_compat_version mods =
+ let test = function SetCompatVersion v -> Some v | _ -> None in
+ try Some (List.find_map test mods) with Not_found -> None
+
+(* Compute precedences from modifiers (or find default ones) *)
+
+let set_entry_type etyps (x,typ) =
+ let typ = try
+ match List.assoc x etyps, typ with
+ | ETConstr (n,()), (_,BorderProd (left,_)) ->
+ ETConstr (n,BorderProd (left,None))
+ | ETConstr (n,()), (_,InternalProd) -> ETConstr (n,InternalProd)
+ | (ETPattern | ETName | ETBigint | ETOther _ |
+ ETReference | ETBinder _ as t), _ -> t
+ | (ETBinderList _ |ETConstrList _), _ -> assert false
+ with Not_found -> ETConstr typ
+ in (x,typ)
+
+let join_auxiliary_recursive_types recvars etyps =
+ List.fold_right (fun (x,y) typs ->
+ let xtyp = try Some (List.assoc x etyps) with Not_found -> None in
+ let ytyp = try Some (List.assoc y etyps) with Not_found -> None in
+ match xtyp,ytyp with
+ | None, None -> typs
+ | Some _, None -> typs
+ | None, Some ytyp -> (x,ytyp)::typs
+ | Some xtyp, Some ytyp when Pervasives.(=) xtyp ytyp -> typs (* FIXME *)
+ | Some xtyp, Some ytyp ->
+ user_err
+ (strbrk "In " ++ pr_id x ++ str " .. " ++ pr_id y ++
+ strbrk ", both ends have incompatible types."))
+ recvars etyps
+
+let internalization_type_of_entry_type = function
+ | ETConstr _ -> NtnInternTypeConstr
+ | ETBigint | ETReference -> NtnInternTypeConstr
+ | ETBinder _ -> NtnInternTypeBinder
+ | ETName -> NtnInternTypeIdent
+ | ETPattern | ETOther _ -> error "Not supported."
+ | ETBinderList _ | ETConstrList _ -> assert false
+
+let set_internalization_type typs =
+ List.map (fun (_, e) -> internalization_type_of_entry_type e) typs
+
+let make_internalization_vars recvars mainvars typs =
+ let maintyps = List.combine mainvars typs in
+ let extratyps = List.map (fun (x,y) -> (y,List.assoc x maintyps)) recvars in
+ maintyps @ extratyps
+
+let make_interpretation_type isrec isonlybinding = function
+ | NtnInternTypeConstr when isrec -> NtnTypeConstrList
+ | NtnInternTypeConstr | NtnInternTypeIdent ->
+ if isonlybinding then NtnTypeOnlyBinder else NtnTypeConstr
+ | NtnInternTypeBinder when isrec -> NtnTypeBinderList
+ | NtnInternTypeBinder -> error "Type binder is only for use in recursive notations for binders."
+
+let make_interpretation_vars recvars allvars =
+ let eq_subscope (sc1, l1) (sc2, l2) =
+ Option.equal String.equal sc1 sc2 &&
+ List.equal String.equal l1 l2
+ in
+ let check (x, y) =
+ let (_,scope1, _) = Id.Map.find x allvars in
+ let (_,scope2, _) = Id.Map.find y allvars in
+ if not (eq_subscope scope1 scope2) then error_not_same_scope x y
+ in
+ let () = List.iter check recvars in
+ let useless_recvars = List.map snd recvars in
+ let mainvars =
+ Id.Map.filter (fun x _ -> not (Id.List.mem x useless_recvars)) allvars in
+ Id.Map.mapi (fun x (isonlybinding, sc, typ) ->
+ (sc, make_interpretation_type (Id.List.mem_assoc x recvars) isonlybinding typ)) mainvars
+
+let check_rule_productivity l =
+ if List.for_all (function NonTerminal _ | Break _ -> true | _ -> false) l then
+ error "A notation must include at least one symbol.";
+ if (match l with SProdList _ :: _ -> true | _ -> false) then
+ error "A recursive notation must start with at least one symbol."
+
+let warn_notation_bound_to_variable =
+ CWarnings.create ~name:"notation-bound-to-variable" ~category:"parsing"
+ (fun () ->
+ strbrk "This notation will not be used for printing as it is bound to a single variable.")
+
+let warn_non_reversible_notation =
+ CWarnings.create ~name:"non-reversible-notation" ~category:"parsing"
+ (fun () ->
+ strbrk "This notation will not be used for printing as it is not reversible.")
+
+let is_not_printable onlyparse nonreversible = function
+| NVar _ ->
+ if not onlyparse then warn_notation_bound_to_variable ();
+ true
+| _ ->
+ if not onlyparse && nonreversible then
+ (warn_non_reversible_notation (); true)
+ else onlyparse
+
+let find_precedence lev etyps symbols =
+ let first_symbol =
+ let rec aux = function
+ | Break _ :: t -> aux t
+ | h :: t -> h
+ | [] -> assert false (* rule is known to be productive *) in
+ aux symbols in
+ let last_is_terminal () =
+ let rec aux b = function
+ | Break _ :: t -> aux b t
+ | Terminal _ :: t -> aux true t
+ | _ :: t -> aux false t
+ | [] -> b in
+ aux false symbols in
+ match first_symbol with
+ | NonTerminal x ->
+ (try match List.assoc x etyps with
+ | ETConstr _ ->
+ error "The level of the leftmost non-terminal cannot be changed."
+ | ETName | ETBigint | ETReference ->
+ begin match lev with
+ | None ->
+ ([Feedback.msg_info ?loc:None ,strbrk "Setting notation at level 0."],0)
+ | Some 0 ->
+ ([],0)
+ | _ ->
+ error "A notation starting with an atomic expression must be at level 0."
+ end
+ | ETPattern | ETBinder _ | ETOther _ -> (* Give a default ? *)
+ if Option.is_empty lev then
+ error "Need an explicit level."
+ else [],Option.get lev
+ | ETConstrList _ | ETBinderList _ ->
+ assert false (* internally used in grammar only *)
+ with Not_found ->
+ if Option.is_empty lev then
+ error "A left-recursive notation must have an explicit level."
+ else [],Option.get lev)
+ | Terminal _ when last_is_terminal () ->
+ if Option.is_empty lev then
+ ([Feedback.msg_info ?loc:None ,strbrk "Setting notation at level 0."], 0)
+ else [],Option.get lev
+ | _ ->
+ if Option.is_empty lev then error "Cannot determine the level.";
+ [],Option.get lev
+
+let check_curly_brackets_notation_exists () =
+ try let _ = Notation.level_of_notation "{ _ }" in ()
+ with Not_found ->
+ error "Notations involving patterns of the form \"{ _ }\" are treated \n\
+specially and require that the notation \"{ _ }\" is already reserved."
+
+let warn_skip_spaces_curly =
+ CWarnings.create ~name:"skip-spaces-curly" ~category:"parsing"
+ (fun () ->strbrk "Skipping spaces inside curly brackets")
+
+(* Remove patterns of the form "{ _ }", unless it is the "{ _ }" notation *)
+let remove_curly_brackets l =
+ let rec skip_break acc = function
+ | Break _ as br :: l -> skip_break (br::acc) l
+ | l -> List.rev acc, l in
+ let rec aux deb = function
+ | [] -> []
+ | Terminal "{" as t1 :: l ->
+ let br,next = skip_break [] l in
+ (match next with
+ | NonTerminal _ as x :: l' as l0 ->
+ let br',next' = skip_break [] l' in
+ (match next' with
+ | Terminal "}" as t2 :: l'' as l1 ->
+ if not (List.equal Notation.symbol_eq l l0) ||
+ not (List.equal Notation.symbol_eq l' l1) then
+ warn_skip_spaces_curly ();
+ if deb && List.is_empty l'' then [t1;x;t2] else begin
+ check_curly_brackets_notation_exists ();
+ x :: aux false l''
+ end
+ | l1 -> t1 :: br @ x :: br' @ aux false l1)
+ | l0 -> t1 :: aux false l0)
+ | x :: l -> x :: aux false l
+ in aux true l
+
+module SynData = struct
+
+ (* XXX: Document *)
+ type syn_data = {
+
+ (* Notation name and location *)
+ info : notation * notation_location;
+
+ (* Fields coming from the vernac-level modifiers *)
+ only_parsing : bool;
+ only_printing : bool;
+ compat : compat_version option;
+ format : string Loc.located option;
+ extra : (string * string) list;
+
+ (* XXX: Callback to printing, must remove *)
+ msgs : ((std_ppcmds -> unit) * std_ppcmds) list;
+
+ (* Fields for internalization *)
+ recvars : (Id.t * Id.t) list;
+ mainvars : Id.List.elt list;
+ intern_typs : notation_var_internalization_type list;
+
+ (* Notation data for parsing *)
+
+ level : int;
+ syntax_data : (Id.t * (production_level, production_position) constr_entry_key_gen) list * (* typs *)
+ symbol list; (* symbols *)
+ not_data : notation * (* notation *)
+ (int * parenRelation) list * (* precedence *)
+ bool; (* needs_squash *)
+ }
+
+end
+
+let compute_syntax_data df modifiers =
+ let open SynData in
+ let open NotationMods in
+ let mods = interp_modifiers modifiers in
+ let assoc = Option.append mods.assoc (Some NonA) in
+ let toks = split_notation_string df in
+ let recvars,mainvars,symbols = analyze_notation_tokens toks in
+ let _ = check_useless_entry_types recvars mainvars mods.etyps in
+
+ (* Notations for interp and grammar *)
+ let ntn_for_interp = make_notation_key symbols in
+ let symbols' = remove_curly_brackets symbols in
+ let ntn_for_grammar = make_notation_key symbols' in
+ check_rule_productivity symbols';
+
+ (* Misc *)
+ let need_squash = not (List.equal Notation.symbol_eq symbols symbols') in
+ let msgs,n = find_precedence mods.level mods.etyps symbols' in
+ let innerlevel = NumLevel 200 in
+ let typs =
+ find_symbols
+ (NumLevel n,BorderProd(Left,assoc))
+ (innerlevel,InternalProd)
+ (NumLevel n,BorderProd(Right,assoc))
+ symbols' in
+ (* To globalize... *)
+ let etyps = join_auxiliary_recursive_types recvars mods.etyps in
+ let sy_typs = List.map (set_entry_type etyps) typs in
+ let prec = List.map (assoc_of_type n) sy_typs in
+ let i_typs = set_internalization_type sy_typs in
+ let sy_data = (sy_typs,symbols') in
+ let sy_fulldata = (ntn_for_grammar,prec,need_squash) in
+ let df' = ((Lib.library_dp(),Lib.current_dirpath true),df) in
+ let i_data = ntn_for_interp, df' in
+
+ (* Return relevant data for interpretation and for parsing/printing *)
+ { info = i_data;
+
+ only_parsing = mods.only_parsing;
+ only_printing = mods.only_printing;
+ compat = mods.compat;
+ format = mods.format;
+ extra = mods.extra;
+
+ msgs;
+
+ recvars;
+ mainvars;
+ intern_typs = i_typs;
+
+ level = n;
+ syntax_data = sy_data;
+ not_data = sy_fulldata;
+ }
+
+let compute_pure_syntax_data df mods =
+ let open SynData in
+ let sd = compute_syntax_data df mods in
+ let msgs =
+ if sd.only_parsing then
+ (Feedback.msg_warning ?loc:None,
+ strbrk "The only parsing modifier has no effect in Reserved Notation.")::sd.msgs
+ else sd.msgs in
+ { sd with msgs }
+
+(**********************************************************************)
+(* Registration of notations interpretation *)
+
+type notation_obj = {
+ notobj_local : bool;
+ notobj_scope : scope_name option;
+ notobj_interp : interpretation;
+ notobj_onlyparse : bool;
+ notobj_onlyprint : bool;
+ notobj_compat : Flags.compat_version option;
+ notobj_notation : notation * notation_location;
+}
+
+let load_notation _ (_, nobj) =
+ Option.iter Notation.declare_scope nobj.notobj_scope
+
+let open_notation i (_, nobj) =
+ let scope = nobj.notobj_scope in
+ let (ntn, df) = nobj.notobj_notation in
+ let pat = nobj.notobj_interp in
+ let fresh = not (Notation.exists_notation_in_scope scope ntn pat) in
+ let active = is_active_compat nobj.notobj_compat in
+ if Int.equal i 1 && fresh && active then begin
+ (* Declare the interpretation *)
+ let onlyprint = nobj.notobj_onlyprint in
+ let () = Notation.declare_notation_interpretation ntn scope pat df ~onlyprint in
+ (* Declare the uninterpretation *)
+ if not nobj.notobj_onlyparse then
+ Notation.declare_uninterpretation (NotationRule (scope, ntn)) pat
+ end
+
+let cache_notation o =
+ load_notation 1 o;
+ open_notation 1 o
+
+let subst_notation (subst, nobj) =
+ { nobj with notobj_interp = subst_interpretation subst nobj.notobj_interp; }
+
+let classify_notation nobj =
+ if nobj.notobj_local then Dispose else Substitute nobj
+
+let inNotation : notation_obj -> obj =
+ declare_object {(default_object "NOTATION") with
+ open_function = open_notation;
+ cache_function = cache_notation;
+ subst_function = subst_notation;
+ load_function = load_notation;
+ classify_function = classify_notation}
+
+(**********************************************************************)
+
+let with_lib_stk_protection f x =
+ let fs = Lib.freeze `No in
+ try let a = f x in Lib.unfreeze fs; a
+ with reraise ->
+ let reraise = CErrors.push reraise in
+ let () = Lib.unfreeze fs in
+ iraise reraise
+
+let with_syntax_protection f x =
+ with_lib_stk_protection
+ (Pcoq.with_grammar_rule_protection
+ (with_notation_protection f)) x
+
+(**********************************************************************)
+(* Recovering existing syntax *)
+
+let contract_notation ntn =
+ if String.equal ntn "{ _ }" then ntn else
+ let rec aux ntn i =
+ if i <= String.length ntn - 5 then
+ let ntn' =
+ if String.is_sub "{ _ }" ntn i &&
+ (i = 0 || ntn.[i-1] = ' ') &&
+ (i = String.length ntn - 5 || ntn.[i+5] = ' ')
+ then
+ String.sub ntn 0 i ^ "_" ^
+ String.sub ntn (i+5) (String.length ntn -i-5)
+ else ntn in
+ aux ntn' (i+1)
+ else ntn in
+ aux ntn 0
+
+exception NoSyntaxRule
+
+let recover_syntax ntn =
+ try
+ let prec = Notation.level_of_notation ntn in
+ let pp_rule,_ = Notation.find_notation_printing_rule ntn in
+ let pp_extra_rules = Notation.find_notation_extra_printing_rules ntn in
+ let pa_rule = Notation.find_notation_parsing_rules ntn in
+ { synext_level = prec;
+ synext_notation = ntn;
+ synext_notgram = pa_rule;
+ synext_unparsing = pp_rule;
+ synext_extra = pp_extra_rules;
+ synext_compat = None;
+ }
+ with Not_found ->
+ raise NoSyntaxRule
+
+let recover_squash_syntax sy =
+ let sq = recover_syntax "{ _ }" in
+ [sy; sq]
+
+let recover_notation_syntax rawntn =
+ let ntn = contract_notation rawntn in
+ let sy = recover_syntax ntn in
+ let need_squash = not (String.equal ntn rawntn) in
+ let rules = if need_squash then recover_squash_syntax sy else [sy] in
+ sy.synext_notgram.notgram_typs, rules, sy.synext_notgram.notgram_onlyprinting
+
+(**********************************************************************)
+(* Main entry point for building parsing and printing rules *)
+
+let make_pa_rule i_typs level (typs,symbols) ntn onlyprint =
+ let assoc = recompute_assoc typs in
+ let prod = make_production typs symbols in
+ { notgram_level = level;
+ notgram_assoc = assoc;
+ notgram_notation = ntn;
+ notgram_prods = prod;
+ notgram_typs = i_typs;
+ notgram_onlyprinting = onlyprint;
+ }
+
+let make_pp_rule level (typs,symbols) fmt =
+ match fmt with
+ | None -> [UnpBox (PpHOVB 0, make_hunks typs symbols level)]
+ | Some fmt -> hunks_of_format (level, List.split typs) (symbols, parse_format fmt)
+
+(* let make_syntax_rules i_typs (ntn,prec,need_squash) sy_data fmt extra onlyprint compat = *)
+let make_syntax_rules (sd : SynData.syn_data) = let open SynData in
+ let ntn, prec, need_squash = sd.not_data in
+ let pa_rule = make_pa_rule sd.intern_typs sd.level sd.syntax_data ntn sd.only_printing in
+ let pp_rule = make_pp_rule sd.level sd.syntax_data sd.format in
+ let sy = {
+ synext_level = (sd.level, prec);
+ synext_notation = ntn;
+ synext_notgram = pa_rule;
+ synext_unparsing = pp_rule;
+ synext_extra = sd.extra;
+ synext_compat = sd.compat;
+ } in
+ (* By construction, the rule for "{ _ }" is declared, but we need to
+ redeclare it because the file where it is declared needs not be open
+ when the current file opens (especially in presence of -nois) *)
+ if need_squash then recover_squash_syntax sy else [sy]
+
+(**********************************************************************)
+(* Main functions about notations *)
+
+let to_map l =
+ let fold accu (x, v) = Id.Map.add x v accu in
+ List.fold_left fold Id.Map.empty l
+
+let add_notation_in_scope local df c mods scope =
+ let open SynData in
+ let sd = compute_syntax_data df mods in
+ (* Prepare the interpretation *)
+ (* Prepare the parsing and printing rules *)
+ let sy_rules = make_syntax_rules sd in
+ let i_vars = make_internalization_vars sd.recvars sd.mainvars sd.intern_typs in
+ let nenv = {
+ ninterp_var_type = to_map i_vars;
+ ninterp_rec_vars = to_map sd.recvars;
+ } in
+ let (acvars, ac, reversible) = interp_notation_constr nenv c in
+ let interp = make_interpretation_vars sd.recvars acvars in
+ let map (x, _) = try Some (x, Id.Map.find x interp) with Not_found -> None in
+ let onlyparse = is_not_printable sd.only_parsing (not reversible) ac in
+ let notation = {
+ notobj_local = local;
+ notobj_scope = scope;
+ notobj_interp = (List.map_filter map i_vars, ac);
+ (** Order is important here! *)
+ notobj_onlyparse = onlyparse;
+ notobj_onlyprint = sd.only_printing;
+ notobj_compat = sd.compat;
+ notobj_notation = sd.info;
+ } in
+ (* Ready to change the global state *)
+ Flags.if_verbose (List.iter (fun (f,x) -> f x)) sd.msgs;
+ Lib.add_anonymous_leaf (inSyntaxExtension (local, sy_rules));
+ Lib.add_anonymous_leaf (inNotation notation);
+ sd.info
+
+let add_notation_interpretation_core local df ?(impls=empty_internalization_env) c scope onlyparse onlyprint compat =
+ let dfs = split_notation_string df in
+ let recvars,mainvars,symbs = analyze_notation_tokens dfs in
+ (* Recover types of variables and pa/pp rules; redeclare them if needed *)
+ let i_typs, onlyprint = if not (is_numeral symbs) then begin
+ let i_typs,sy_rules,onlyprint' = recover_notation_syntax (make_notation_key symbs) in
+ let () = Lib.add_anonymous_leaf (inSyntaxExtension (local,sy_rules)) in
+ (** If the only printing flag has been explicitly requested, put it back *)
+ let onlyprint = onlyprint || onlyprint' in
+ i_typs, onlyprint
+ end else [], false in
+ (* Declare interpretation *)
+ let path = (Lib.library_dp(), Lib.current_dirpath true) in
+ let df' = (make_notation_key symbs, (path,df)) in
+ let i_vars = make_internalization_vars recvars mainvars i_typs in
+ let nenv = {
+ ninterp_var_type = to_map i_vars;
+ ninterp_rec_vars = to_map recvars;
+ } in
+ let (acvars, ac, reversible) = interp_notation_constr ~impls nenv c in
+ let interp = make_interpretation_vars recvars acvars in
+ let map (x, _) = try Some (x, Id.Map.find x interp) with Not_found -> None in
+ let onlyparse = is_not_printable onlyparse (not reversible) ac in
+ let notation = {
+ notobj_local = local;
+ notobj_scope = scope;
+ notobj_interp = (List.map_filter map i_vars, ac);
+ (** Order is important here! *)
+ notobj_onlyparse = onlyparse;
+ notobj_onlyprint = onlyprint;
+ notobj_compat = compat;
+ notobj_notation = df';
+ } in
+ Lib.add_anonymous_leaf (inNotation notation);
+ df'
+
+(* Notations without interpretation (Reserved Notation) *)
+
+let add_syntax_extension local ((loc,df),mods) = let open SynData in
+ let psd = compute_pure_syntax_data df mods in
+ let sy_rules = make_syntax_rules {psd with compat = None} in
+ Flags.if_verbose (List.iter (fun (f,x) -> f x)) psd.msgs;
+ Lib.add_anonymous_leaf (inSyntaxExtension(local,sy_rules))
+
+(* Notations with only interpretation *)
+
+let add_notation_interpretation ((loc,df),c,sc) =
+ let df' = add_notation_interpretation_core false df c sc false false None in
+ Dumpglob.dump_notation (loc,df') sc true
+
+let set_notation_for_interpretation impls ((_,df),c,sc) =
+ (try ignore
+ (silently (fun () -> add_notation_interpretation_core false df ~impls c sc false false None) ());
+ with NoSyntaxRule ->
+ error "Parsing rule for this notation has to be previously declared.");
+ Option.iter (fun sc -> Notation.open_close_scope (false,true,sc)) sc
+
+(* Main entry point *)
+
+let add_notation local c ((loc,df),modifiers) sc =
+ let df' =
+ if no_syntax_modifiers modifiers then
+ (* No syntax data: try to rely on a previously declared rule *)
+ let onlyparse = is_only_parsing modifiers in
+ let onlyprint = is_only_printing modifiers in
+ let compat = get_compat_version modifiers in
+ try add_notation_interpretation_core local df c sc onlyparse onlyprint compat
+ with NoSyntaxRule ->
+ (* Try to determine a default syntax rule *)
+ add_notation_in_scope local df c modifiers sc
+ else
+ (* Declare both syntax and interpretation *)
+ add_notation_in_scope local df c modifiers sc
+ in
+ Dumpglob.dump_notation (loc,df') sc true
+
+let add_notation_extra_printing_rule df k v =
+ let notk =
+ let dfs = split_notation_string df in
+ let _,_, symbs = analyze_notation_tokens dfs in
+ make_notation_key symbs in
+ Notation.add_notation_extra_printing_rule notk k v
+
+(* Infix notations *)
+
+let inject_var x = CRef (Ident (Loc.ghost, Id.of_string x),None)
+
+let add_infix local ((loc,inf),modifiers) pr sc =
+ check_infix_modifiers modifiers;
+ (* check the precedence *)
+ let metas = [inject_var "x"; inject_var "y"] in
+ let c = mkAppC (pr,metas) in
+ let df = "x "^(quote_notation_token inf)^" y" in
+ add_notation local c ((loc,df),modifiers) sc
+
+(**********************************************************************)
+(* Delimiters and classes bound to scopes *)
+
+type scope_command =
+ | ScopeDelim of string
+ | ScopeClasses of scope_class list
+ | ScopeRemove
+
+let load_scope_command _ (_,(scope,dlm)) =
+ Notation.declare_scope scope
+
+let open_scope_command i (_,(scope,o)) =
+ if Int.equal i 1 then
+ match o with
+ | ScopeDelim dlm -> Notation.declare_delimiters scope dlm
+ | ScopeClasses cl -> List.iter (Notation.declare_scope_class scope) cl
+ | ScopeRemove -> Notation.remove_delimiters scope
+
+let cache_scope_command o =
+ load_scope_command 1 o;
+ open_scope_command 1 o
+
+let subst_scope_command (subst,(scope,o as x)) = match o with
+ | ScopeClasses cl ->
+ let cl' = List.map_filter (subst_scope_class subst) cl in
+ let cl' =
+ if List.for_all2eq (==) cl cl' then cl
+ else cl' in
+ scope, ScopeClasses cl'
+ | _ -> x
+
+let inScopeCommand : scope_name * scope_command -> obj =
+ declare_object {(default_object "DELIMITERS") with
+ cache_function = cache_scope_command;
+ open_function = open_scope_command;
+ load_function = load_scope_command;
+ subst_function = subst_scope_command;
+ classify_function = (fun obj -> Substitute obj)}
+
+let add_delimiters scope key =
+ Lib.add_anonymous_leaf (inScopeCommand(scope,ScopeDelim key))
+
+let remove_delimiters scope =
+ Lib.add_anonymous_leaf (inScopeCommand(scope,ScopeRemove))
+
+let add_class_scope scope cl =
+ Lib.add_anonymous_leaf (inScopeCommand(scope,ScopeClasses cl))
+
+(* Check if abbreviation to a name and avoid early insertion of
+ maximal implicit arguments *)
+let try_interp_name_alias = function
+ | [], CRef (ref,_) -> intern_reference ref
+ | _ -> raise Not_found
+
+let add_syntactic_definition ident (vars,c) local onlyparse =
+ let nonprintable = ref false in
+ let vars,pat =
+ try [], NRef (try_interp_name_alias (vars,c))
+ with Not_found ->
+ let fold accu id = Id.Map.add id NtnInternTypeConstr accu in
+ let i_vars = List.fold_left fold Id.Map.empty vars in
+ let nenv = {
+ ninterp_var_type = i_vars;
+ ninterp_rec_vars = Id.Map.empty;
+ } in
+ let nvars, pat, reversible = interp_notation_constr nenv c in
+ let () = nonprintable := not reversible in
+ let map id = let (_,sc,_) = Id.Map.find id nvars in (id, sc) in
+ List.map map vars, pat
+ in
+ let onlyparse = match onlyparse with
+ | None when (is_not_printable false !nonprintable pat) -> Some Flags.Current
+ | p -> p
+ in
+ Syntax_def.declare_syntactic_definition local ident onlyparse (vars,pat)
diff --git a/vernac/metasyntax.mli b/vernac/metasyntax.mli
new file mode 100644
index 000000000..57c120402
--- /dev/null
+++ b/vernac/metasyntax.mli
@@ -0,0 +1,63 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open Vernacexpr
+open Notation
+open Constrexpr
+open Notation_term
+
+val add_token_obj : string -> unit
+
+(** Adding a (constr) notation in the environment*)
+
+val add_infix : locality_flag -> (lstring * syntax_modifier list) ->
+ constr_expr -> scope_name option -> unit
+
+val add_notation : locality_flag -> constr_expr ->
+ (lstring * syntax_modifier list) -> scope_name option -> unit
+
+val add_notation_extra_printing_rule : string -> string -> string -> unit
+
+(** Declaring delimiter keys and default scopes *)
+
+val add_delimiters : scope_name -> string -> unit
+val remove_delimiters : scope_name -> unit
+val add_class_scope : scope_name -> scope_class list -> unit
+
+(** Add only the interpretation of a notation that already has pa/pp rules *)
+
+val add_notation_interpretation :
+ (lstring * constr_expr * scope_name option) -> unit
+
+(** Add a notation interpretation for supporting the "where" clause *)
+
+val set_notation_for_interpretation : Constrintern.internalization_env ->
+ (lstring * constr_expr * scope_name option) -> unit
+
+(** Add only the parsing/printing rule of a notation *)
+
+val add_syntax_extension :
+ locality_flag -> (lstring * syntax_modifier list) -> unit
+
+(** Add a syntactic definition (as in "Notation f := ...") *)
+
+val add_syntactic_definition : Id.t -> Id.t list * constr_expr ->
+ bool -> Flags.compat_version option -> unit
+
+(** Print the Camlp4 state of a grammar *)
+
+val pr_grammar : string -> Pp.std_ppcmds
+
+type any_entry = AnyEntry : 'a Pcoq.Gram.entry -> any_entry
+
+val register_grammar : string -> any_entry list -> unit
+
+val check_infix_modifiers : syntax_modifier list -> unit
+
+val with_syntax_protection : ('a -> 'b) -> 'a -> 'b
diff --git a/vernac/mltop.ml b/vernac/mltop.ml
new file mode 100644
index 000000000..2396cf04a
--- /dev/null
+++ b/vernac/mltop.ml
@@ -0,0 +1,447 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open CErrors
+open Util
+open Pp
+open Flags
+open Libobject
+open System
+
+(* Code to hook Coq into the ML toplevel -- depends on having the
+ objective-caml compiler mostly visible. The functions implemented here are:
+ \begin{itemize}
+ \item [dir_ml_load name]: Loads the ML module fname from the current ML
+ path.
+ \item [dir_ml_use]: Directive #use of Ocaml toplevel
+ \item [add_ml_dir]: Directive #directory of Ocaml toplevel
+ \end{itemize}
+
+ How to build an ML module interface with these functions.
+ The idea is that the ML directory path is like the Coq directory
+ path. So we can maintain the two in parallel.
+ In the same way, we can use the "ml_env" as a kind of ML
+ environment, which we freeze, unfreeze, and add things to just like
+ to the other environments.
+ Finally, we can create an object which is an ML module, and require
+ that the "caching" of the ML module cause the loading of the
+ associated ML file, if that file has not been yet loaded. Of
+ course, the problem is how to record dependencies between ML
+ modules.
+ (I do not know of a solution to this problem, other than to
+ put all the needed names into the ML Module object.) *)
+
+
+(* NB: this module relies on OCaml's Dynlink library. The bytecode version
+ of Dynlink is always available, but there are some architectures
+ with native compilation and no dynlink.cmxa. Instead of nasty IFDEF's
+ here for hiding the calls to Dynlink, we now simply reject this rather
+ rare situation during ./configure, and give instructions there about how
+ to build a dummy dynlink.cmxa, cf. dev/dynlink.ml. *)
+
+(* This path is where we look for .cmo *)
+let coq_mlpath_copy = ref [Sys.getcwd ()]
+let keep_copy_mlpath path =
+ let cpath = CUnix.canonical_path_name path in
+ let filter path' = not (String.equal cpath path')
+ in
+ coq_mlpath_copy := cpath :: List.filter filter !coq_mlpath_copy
+
+(* If there is a toplevel under Coq *)
+type toplevel = {
+ load_obj : string -> unit;
+ use_file : string -> unit;
+ add_dir : string -> unit;
+ ml_loop : unit -> unit }
+
+(* Determines the behaviour of Coq with respect to ML files (compiled
+ or not) *)
+type kind_load =
+ | WithTop of toplevel
+ | WithoutTop
+
+(* Must be always initialized *)
+let load = ref WithoutTop
+
+(* Are we in a native version of Coq? *)
+let is_native = Dynlink.is_native
+
+(* Sets and initializes a toplevel (if any) *)
+let set_top toplevel = load :=
+ WithTop toplevel;
+ Nativelib.load_obj := toplevel.load_obj
+
+(* Removes the toplevel (if any) *)
+let remove () =
+ load := WithoutTop;
+ Nativelib.load_obj := (fun x -> () : string -> unit)
+
+(* Tests if an Ocaml toplevel runs under Coq *)
+let is_ocaml_top () =
+ match !load with
+ | WithTop _ -> true
+ |_ -> false
+
+(* Tests if we can load ML files *)
+let has_dynlink = Coq_config.has_natdynlink || not is_native
+
+(* Runs the toplevel loop of Ocaml *)
+let ocaml_toploop () =
+ match !load with
+ | WithTop t -> Printexc.catch t.ml_loop ()
+ | _ -> ()
+
+(* Try to interpret load_obj's (internal) errors *)
+let report_on_load_obj_error exc =
+ let x = Obj.repr exc in
+ (* Try an horrible (fragile) hack to report on Symtable dynlink errors *)
+ (* (we follow ocaml's Printexc.to_string decoding of exceptions) *)
+ if Obj.is_block x && String.equal (Obj.magic (Obj.field (Obj.field x 0) 0)) "Symtable.Error"
+ then
+ let err_block = Obj.field x 1 in
+ if Int.equal (Obj.tag err_block) 0 then
+ (* Symtable.Undefined_global of string *)
+ str "reference to undefined global " ++
+ str (Obj.magic (Obj.field err_block 0))
+ else str (Printexc.to_string exc)
+ else str (Printexc.to_string exc)
+
+(* Dynamic loading of .cmo/.cma *)
+
+let ml_load s =
+ match !load with
+ | WithTop t ->
+ (try t.load_obj s; s
+ with
+ | e when CErrors.noncritical e ->
+ let e = CErrors.push e in
+ match fst e with
+ | (UserError _ | Failure _ | Not_found as u) -> Exninfo.iraise (u, snd e)
+ | exc ->
+ let msg = report_on_load_obj_error exc in
+ user_err ~hdr:"Mltop.load_object" (str"Cannot link ml-object " ++
+ str s ++ str" to Coq code (" ++ msg ++ str ")."))
+ | WithoutTop ->
+ try
+ Dynlink.loadfile s; s
+ with Dynlink.Error a ->
+ user_err ~hdr:"Mltop.load_object"
+ (strbrk "while loading " ++ str s ++
+ strbrk ": " ++ str (Dynlink.error_message a))
+
+let dir_ml_load s =
+ match !load with
+ | WithTop _ -> ml_load s
+ | WithoutTop ->
+ let warn = Flags.is_verbose() in
+ let _,gname = find_file_in_path ~warn !coq_mlpath_copy s in
+ ml_load gname
+
+(* Dynamic interpretation of .ml *)
+let dir_ml_use s =
+ match !load with
+ | WithTop t -> t.use_file s
+ | _ ->
+ let moreinfo =
+ if Dynlink.is_native then " Loading ML code works only in bytecode."
+ else ""
+ in
+ user_err ~hdr:"Mltop.dir_ml_use" (str "Could not load ML code." ++ str moreinfo)
+
+(* Adds a path to the ML paths *)
+let add_ml_dir s =
+ match !load with
+ | WithTop t -> t.add_dir s; keep_copy_mlpath s
+ | WithoutTop when has_dynlink -> keep_copy_mlpath s
+ | _ -> ()
+
+(* For Rec Add ML Path (-R) *)
+let add_rec_ml_dir unix_path =
+ List.iter (fun (lp,_) -> add_ml_dir lp) (all_subdirs ~unix_path)
+
+(* Adding files to Coq and ML loadpath *)
+
+let warn_cannot_use_directory =
+ CWarnings.create ~name:"cannot-use-directory" ~category:"filesystem"
+ (fun d ->
+ str "Directory " ++ str d ++
+ strbrk " cannot be used as a Coq identifier (skipped)")
+
+let convert_string d =
+ try Names.Id.of_string d
+ with UserError _ ->
+ warn_cannot_use_directory d;
+ raise Exit
+
+let warn_cannot_open_path =
+ CWarnings.create ~name:"cannot-open-path" ~category:"filesystem"
+ (fun unix_path -> str "Cannot open " ++ str unix_path)
+
+type add_ml = AddNoML | AddTopML | AddRecML
+
+let add_rec_path add_ml ~unix_path ~coq_root ~implicit =
+ if exists_dir unix_path then
+ let dirs = all_subdirs ~unix_path in
+ let prefix = Names.DirPath.repr coq_root in
+ let convert_dirs (lp, cp) =
+ try
+ let path = List.rev_map convert_string cp @ prefix in
+ Some (lp, Names.DirPath.make path)
+ with Exit -> None
+ in
+ let dirs = List.map_filter convert_dirs dirs in
+ let () = match add_ml with
+ | AddNoML -> ()
+ | AddTopML -> add_ml_dir unix_path
+ | AddRecML -> List.iter (fun (lp,_) -> add_ml_dir lp) dirs in
+ let add (path, dir) =
+ Loadpath.add_load_path path ~implicit dir in
+ let () = List.iter add dirs in
+ Loadpath.add_load_path unix_path ~implicit coq_root
+ else
+ warn_cannot_open_path unix_path
+
+(* convertit un nom quelconque en nom de fichier ou de module *)
+let mod_of_name name =
+ if Filename.check_suffix name ".cmo" then
+ Filename.chop_suffix name ".cmo"
+ else
+ name
+
+let get_ml_object_suffix name =
+ if Filename.check_suffix name ".cmo" then
+ Some ".cmo"
+ else if Filename.check_suffix name ".cma" then
+ Some ".cma"
+ else if Filename.check_suffix name ".cmxs" then
+ Some ".cmxs"
+ else
+ None
+
+let file_of_name name =
+ let suffix = get_ml_object_suffix name in
+ let fail s =
+ user_err ~hdr:"Mltop.load_object"
+ (str"File not found on loadpath : " ++ str s ++ str"\n" ++
+ str"Loadpath: " ++ str(String.concat ":" !coq_mlpath_copy)) in
+ if not (Filename.is_relative name) then
+ if Sys.file_exists name then name else fail name
+ else if is_native then
+ let name = match suffix with
+ | Some ((".cmo"|".cma") as suffix) ->
+ (Filename.chop_suffix name suffix) ^ ".cmxs"
+ | Some ".cmxs" -> name
+ | _ -> name ^ ".cmxs"
+ in
+ if is_in_path !coq_mlpath_copy name then name else fail name
+ else
+ let (full, base) = match suffix with
+ | Some ".cmo" | Some ".cma" -> true, name
+ | Some ".cmxs" -> false, Filename.chop_suffix name ".cmxs"
+ | _ -> false, name
+ in
+ if full then
+ if is_in_path !coq_mlpath_copy base then base else fail base
+ else
+ let name = base ^ ".cma" in
+ if is_in_path !coq_mlpath_copy name then name else
+ let name = base ^ ".cmo" in
+ if is_in_path !coq_mlpath_copy name then name else
+ fail (base ^ ".cm[ao]")
+
+(** Is the ML code of the standard library placed into loadable plugins
+ or statically compiled into coqtop ? For the moment this choice is
+ made according to the presence of native dynlink : even if bytecode
+ coqtop could always load plugins, we prefer to have uniformity between
+ bytecode and native versions. *)
+
+(* [known_loaded_module] contains the names of the loaded ML modules
+ * (linked or loaded with load_object). It is used not to load a
+ * module twice. It is NOT the list of ML modules Coq knows. *)
+
+let known_loaded_modules = ref String.Map.empty
+
+let add_known_module mname path =
+ if not (String.Map.mem mname !known_loaded_modules) ||
+ String.Map.find mname !known_loaded_modules = None then
+ known_loaded_modules := String.Map.add mname path !known_loaded_modules
+
+let module_is_known mname =
+ String.Map.mem mname !known_loaded_modules
+
+let known_module_path mname =
+ String.Map.find mname !known_loaded_modules
+
+(** A plugin is just an ML module with an initialization function. *)
+
+let known_loaded_plugins = ref String.Map.empty
+
+let add_known_plugin init name =
+ add_known_module name None;
+ known_loaded_plugins := String.Map.add name init !known_loaded_plugins
+
+let init_known_plugins () =
+ String.Map.iter (fun _ f -> f()) !known_loaded_plugins
+
+(** Registering functions to be used at caching time, that is when the Declare
+ ML module command is issued. *)
+
+let cache_objs = ref String.Map.empty
+
+let declare_cache_obj f name =
+ let objs = try String.Map.find name !cache_objs with Not_found -> [] in
+ let objs = f :: objs in
+ cache_objs := String.Map.add name objs !cache_objs
+
+let perform_cache_obj name =
+ let objs = try String.Map.find name !cache_objs with Not_found -> [] in
+ let objs = List.rev objs in
+ List.iter (fun f -> f ()) objs
+
+(** ml object = ml module or plugin *)
+
+let init_ml_object mname =
+ try String.Map.find mname !known_loaded_plugins ()
+ with Not_found -> ()
+
+let load_ml_object mname ?path fname=
+ let path = match path with
+ | None -> dir_ml_load fname
+ | Some p -> ml_load p in
+ add_known_module mname (Some path);
+ init_ml_object mname;
+ path
+
+let dir_ml_load m = ignore(dir_ml_load m)
+let add_known_module m = add_known_module m None
+let load_ml_object_raw fname = dir_ml_load (file_of_name fname)
+let load_ml_objects_raw_rex rex =
+ List.iter (fun (_,fp) ->
+ let name = file_of_name (Filename.basename fp) in
+ try dir_ml_load name
+ with e -> prerr_endline (Printexc.to_string e))
+ (System.where_in_path_rex !coq_mlpath_copy rex)
+
+(* Summary of declared ML Modules *)
+
+(* List and not String.Set because order is important: most recent first. *)
+
+let loaded_modules = ref []
+let get_loaded_modules () = List.rev !loaded_modules
+let add_loaded_module md path =
+ if not (List.mem_assoc md !loaded_modules) then
+ loaded_modules := (md,path) :: !loaded_modules
+let reset_loaded_modules () = loaded_modules := []
+
+let if_verbose_load verb f name ?path fname =
+ if not verb then f name ?path fname
+ else
+ let info = str "[Loading ML file " ++ str fname ++ str " ..." in
+ try
+ let path = f name ?path fname in
+ Feedback.msg_info (info ++ str " done]");
+ path
+ with reraise ->
+ Feedback.msg_info (info ++ str " failed]");
+ raise reraise
+
+(** Load a module for the first time (i.e. dynlink it)
+ or simulate its reload (i.e. doing nothing except maybe
+ an initialization function). *)
+
+let trigger_ml_object verb cache reinit ?path name =
+ if module_is_known name then begin
+ if reinit then init_ml_object name;
+ add_loaded_module name (known_module_path name);
+ if cache then perform_cache_obj name
+ end else if not has_dynlink then
+ user_err ~hdr:"Mltop.trigger_ml_object"
+ (str "Dynamic link not supported (module " ++ str name ++ str ")")
+ else begin
+ let file = file_of_name (Option.default name path) in
+ let path =
+ if_verbose_load (verb && is_verbose ()) load_ml_object name ?path file in
+ add_loaded_module name (Some path);
+ if cache then perform_cache_obj name
+ end
+
+let load_ml_object n m = ignore(load_ml_object n m)
+
+let unfreeze_ml_modules x =
+ reset_loaded_modules ();
+ List.iter
+ (fun (name,path) -> trigger_ml_object false false false ?path name) x
+
+let _ =
+ Summary.declare_summary Summary.ml_modules
+ { Summary.freeze_function = (fun _ -> get_loaded_modules ());
+ Summary.unfreeze_function = unfreeze_ml_modules;
+ Summary.init_function = reset_loaded_modules }
+
+(* Liboject entries of declared ML Modules *)
+
+type ml_module_object = {
+ mlocal : Vernacexpr.locality_flag;
+ mnames : string list
+}
+
+let cache_ml_objects (_,{mnames=mnames}) =
+ let iter obj = trigger_ml_object true true true obj in
+ List.iter iter mnames
+
+let load_ml_objects _ (_,{mnames=mnames}) =
+ let iter obj = trigger_ml_object true false true obj in
+ List.iter iter mnames
+
+let classify_ml_objects ({mlocal=mlocal} as o) =
+ if mlocal then Dispose else Substitute o
+
+let inMLModule : ml_module_object -> obj =
+ declare_object
+ {(default_object "ML-MODULE") with
+ cache_function = cache_ml_objects;
+ load_function = load_ml_objects;
+ subst_function = (fun (_,o) -> o);
+ classify_function = classify_ml_objects }
+
+let declare_ml_modules local l =
+ let l = List.map mod_of_name l in
+ Lib.add_anonymous_leaf ~cache_first:false (inMLModule {mlocal=local; mnames=l})
+
+let print_ml_path () =
+ let l = !coq_mlpath_copy in
+ str"ML Load Path:" ++ fnl () ++ str" " ++
+ hv 0 (prlist_with_sep fnl str l)
+
+(* Printing of loaded ML modules *)
+
+let print_ml_modules () =
+ let l = get_loaded_modules () in
+ str"Loaded ML Modules: " ++ pr_vertical_list str (List.map fst l)
+
+let print_gc () =
+ let stat = Gc.stat () in
+ let msg =
+ str "minor words: " ++ real stat.Gc.minor_words ++ fnl () ++
+ str "promoted words: " ++ real stat.Gc.promoted_words ++ fnl () ++
+ str "major words: " ++ real stat.Gc.major_words ++ fnl () ++
+ str "minor_collections: " ++ int stat.Gc.minor_collections ++ fnl () ++
+ str "major_collections: " ++ int stat.Gc.major_collections ++ fnl () ++
+ str "heap_words: " ++ int stat.Gc.heap_words ++ fnl () ++
+ str "heap_chunks: " ++ int stat.Gc.heap_chunks ++ fnl () ++
+ str "live_words: " ++ int stat.Gc.live_words ++ fnl () ++
+ str "live_blocks: " ++ int stat.Gc.live_blocks ++ fnl () ++
+ str "free_words: " ++ int stat.Gc.free_words ++ fnl () ++
+ str "free_blocks: " ++ int stat.Gc.free_blocks ++ fnl () ++
+ str "largest_free: " ++ int stat.Gc.largest_free ++ fnl () ++
+ str "fragments: " ++ int stat.Gc.fragments ++ fnl () ++
+ str "compactions: " ++ int stat.Gc.compactions ++ fnl () ++
+ str "top_heap_words: " ++ int stat.Gc.top_heap_words ++ fnl () ++
+ str "stack_size: " ++ int stat.Gc.stack_size
+ in
+ hv 0 msg
diff --git a/vernac/mltop.mli b/vernac/mltop.mli
new file mode 100644
index 000000000..6633cb937
--- /dev/null
+++ b/vernac/mltop.mli
@@ -0,0 +1,88 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** {5 Toplevel management} *)
+
+(** If there is a toplevel under Coq, it is described by the following
+ record. *)
+type toplevel = {
+ load_obj : string -> unit;
+ use_file : string -> unit;
+ add_dir : string -> unit;
+ ml_loop : unit -> unit }
+
+(** Sets and initializes a toplevel (if any) *)
+val set_top : toplevel -> unit
+
+(** Are we in a native version of Coq? *)
+val is_native : bool
+
+(** Removes the toplevel (if any) *)
+val remove : unit -> unit
+
+(** Tests if an Ocaml toplevel runs under Coq *)
+val is_ocaml_top : unit -> bool
+
+(** Starts the Ocaml toplevel loop *)
+val ocaml_toploop : unit -> unit
+
+(** {5 ML Dynlink} *)
+
+(** Tests if we can load ML files *)
+val has_dynlink : bool
+
+(** Dynamic loading of .cmo *)
+val dir_ml_load : string -> unit
+
+(** Dynamic interpretation of .ml *)
+val dir_ml_use : string -> unit
+
+(** Adds a path to the ML paths *)
+val add_ml_dir : string -> unit
+val add_rec_ml_dir : string -> unit
+
+type add_ml = AddNoML | AddTopML | AddRecML
+
+(** Adds a path to the Coq and ML paths *)
+val add_rec_path : add_ml -> unix_path:string -> coq_root:Names.DirPath.t -> implicit:bool -> unit
+
+(** List of modules linked to the toplevel *)
+val add_known_module : string -> unit
+val module_is_known : string -> bool
+val load_ml_object : string -> string -> unit
+val load_ml_object_raw : string -> unit
+val load_ml_objects_raw_rex : Str.regexp -> unit
+
+(** {5 Initialization functions} *)
+
+(** Declare a plugin and its initialization function.
+ A plugin is just an ML module with an initialization function.
+ Adding a known plugin implies adding it as a known ML module.
+ The initialization function is granted to be called after Coq is fully
+ bootstrapped, even if the plugin is statically linked with the toplevel *)
+val add_known_plugin : (unit -> unit) -> string -> unit
+
+(** Calls all initialization functions in a non-specified order *)
+val init_known_plugins : unit -> unit
+
+(** Register a callback that will be called when the module is declared with
+ the Declare ML Module command. This is useful to define Coq objects at that
+ time only. Several functions can be defined for one module; they will be
+ called in the order of declaration, and after the ML module has been
+ properly initialized. *)
+val declare_cache_obj : (unit -> unit) -> string -> unit
+
+(** {5 Declaring modules} *)
+
+val declare_ml_modules : Vernacexpr.locality_flag -> string list -> unit
+
+(** {5 Utilities} *)
+
+val print_ml_path : unit -> Pp.std_ppcmds
+val print_ml_modules : unit -> Pp.std_ppcmds
+val print_gc : unit -> Pp.std_ppcmds
diff --git a/vernac/obligations.ml b/vernac/obligations.ml
new file mode 100644
index 000000000..c1d9ae48a
--- /dev/null
+++ b/vernac/obligations.ml
@@ -0,0 +1,1179 @@
+open Printf
+open Globnames
+open Libobject
+open Entries
+open Decl_kinds
+open Declare
+
+(**
+ - Get types of existentials ;
+ - Flatten dependency tree (prefix order) ;
+ - Replace existentials by De Bruijn indices in term, applied to the right arguments ;
+ - Apply term prefixed by quantification on "existentials".
+*)
+
+open Term
+open Vars
+open Names
+open Evd
+open Pp
+open CErrors
+open Util
+
+(* EJGA: This should go away, no more need for fix_exn *)
+module Stm = struct
+ let u = (fun x -> x)
+ let get_fix_exn () = u
+end
+
+module NamedDecl = Context.Named.Declaration
+
+let declare_fix_ref = ref (fun ?opaque _ _ _ _ _ _ -> assert false)
+let declare_definition_ref = ref (fun _ _ _ _ _ -> assert false)
+
+let succfix (depth, fixrels) =
+ (succ depth, List.map succ fixrels)
+
+let check_evars env evm =
+ Evar.Map.iter
+ (fun key evi ->
+ let (loc,k) = evar_source key evm in
+ match k with
+ | Evar_kinds.QuestionMark _
+ | Evar_kinds.ImplicitArg (_,_,false) -> ()
+ | _ ->
+ Pretype_errors.error_unsolvable_implicit ~loc env evm key None)
+ (Evd.undefined_map evm)
+
+type oblinfo =
+ { ev_name: int * Id.t;
+ ev_hyps: Context.Named.t;
+ ev_status: bool * Evar_kinds.obligation_definition_status;
+ ev_chop: int option;
+ ev_src: Evar_kinds.t Loc.located;
+ ev_typ: types;
+ ev_tac: unit Proofview.tactic option;
+ ev_deps: Int.Set.t }
+
+(** Substitute evar references in t using De Bruijn indices,
+ where n binders were passed through. *)
+
+let subst_evar_constr evs n idf t =
+ let seen = ref Int.Set.empty in
+ let transparent = ref Id.Set.empty in
+ let evar_info id = List.assoc_f Evar.equal id evs in
+ let rec substrec (depth, fixrels) c = match kind_of_term c with
+ | Evar (k, args) ->
+ let { ev_name = (id, idstr) ;
+ ev_hyps = hyps ; ev_chop = chop } =
+ try evar_info k
+ with Not_found ->
+ anomaly ~label:"eterm" (Pp.str "existential variable " ++ int (Evar.repr k) ++ str " not found")
+ in
+ seen := Int.Set.add id !seen;
+ (* Evar arguments are created in inverse order,
+ and we must not apply to defined ones (i.e. LetIn's)
+ *)
+ let args =
+ let n = match chop with None -> 0 | Some c -> c in
+ let (l, r) = List.chop n (List.rev (Array.to_list args)) in
+ List.rev r
+ in
+ let args =
+ let rec aux hyps args acc =
+ let open Context.Named.Declaration in
+ match hyps, args with
+ (LocalAssum _ :: tlh), (c :: tla) ->
+ aux tlh tla ((substrec (depth, fixrels) c) :: acc)
+ | (LocalDef _ :: tlh), (_ :: tla) ->
+ aux tlh tla acc
+ | [], [] -> acc
+ | _, _ -> acc (*failwith "subst_evars: invalid argument"*)
+ in aux hyps args []
+ in
+ if List.exists
+ (fun x -> match kind_of_term x with
+ | Rel n -> Int.List.mem n fixrels
+ | _ -> false) args
+ then
+ transparent := Id.Set.add idstr !transparent;
+ mkApp (idf idstr, Array.of_list args)
+ | Fix _ ->
+ map_constr_with_binders succfix substrec (depth, 1 :: fixrels) c
+ | _ -> map_constr_with_binders succfix substrec (depth, fixrels) c
+ in
+ let t' = substrec (0, []) t in
+ t', !seen, !transparent
+
+
+(** Substitute variable references in t using De Bruijn indices,
+ where n binders were passed through. *)
+let subst_vars acc n t =
+ let var_index id = Util.List.index Id.equal id acc in
+ let rec substrec depth c = match kind_of_term c with
+ | Var v -> (try mkRel (depth + (var_index v)) with Not_found -> c)
+ | _ -> map_constr_with_binders succ substrec depth c
+ in
+ substrec 0 t
+
+(** Rewrite type of an evar ([ H1 : t1, ... Hn : tn |- concl ])
+ to a product : forall H1 : t1, ..., forall Hn : tn, concl.
+ Changes evars and hypothesis references to variable references.
+*)
+let etype_of_evar evs hyps concl =
+ let open Context.Named.Declaration in
+ let rec aux acc n = function
+ decl :: tl ->
+ let t', s, trans = subst_evar_constr evs n mkVar (NamedDecl.get_type decl) in
+ let t'' = subst_vars acc 0 t' in
+ let rest, s', trans' = aux (NamedDecl.get_id decl :: acc) (succ n) tl in
+ let s' = Int.Set.union s s' in
+ let trans' = Id.Set.union trans trans' in
+ (match decl with
+ | LocalDef (id,c,_) ->
+ let c', s'', trans'' = subst_evar_constr evs n mkVar c in
+ let c' = subst_vars acc 0 c' in
+ mkNamedProd_or_LetIn (LocalDef (id, c', t'')) rest,
+ Int.Set.union s'' s',
+ Id.Set.union trans'' trans'
+ | LocalAssum (id,_) ->
+ mkNamedProd_or_LetIn (LocalAssum (id, t'')) rest, s', trans')
+ | [] ->
+ let t', s, trans = subst_evar_constr evs n mkVar concl in
+ subst_vars acc 0 t', s, trans
+ in aux [] 0 (List.rev hyps)
+
+let trunc_named_context n ctx =
+ let len = List.length ctx in
+ List.firstn (len - n) ctx
+
+let rec chop_product n t =
+ if Int.equal n 0 then Some t
+ else
+ match kind_of_term t with
+ | Prod (_, _, b) -> if noccurn 1 b then chop_product (pred n) (Termops.pop b) else None
+ | _ -> None
+
+let evar_dependencies evm oev =
+ let one_step deps =
+ Evar.Set.fold (fun ev s ->
+ let evi = Evd.find evm ev in
+ let deps' = evars_of_filtered_evar_info evi in
+ if Evar.Set.mem oev deps' then
+ invalid_arg ("Ill-formed evar map: cycle detected for evar " ^ string_of_existential oev)
+ else Evar.Set.union deps' s)
+ deps deps
+ in
+ let rec aux deps =
+ let deps' = one_step deps in
+ if Evar.Set.equal deps deps' then deps
+ else aux deps'
+ in aux (Evar.Set.singleton oev)
+
+let move_after (id, ev, deps as obl) l =
+ let rec aux restdeps = function
+ | (id', _, _) as obl' :: tl ->
+ let restdeps' = Evar.Set.remove id' restdeps in
+ if Evar.Set.is_empty restdeps' then
+ obl' :: obl :: tl
+ else obl' :: aux restdeps' tl
+ | [] -> [obl]
+ in aux (Evar.Set.remove id deps) l
+
+let sort_dependencies evl =
+ let rec aux l found list =
+ match l with
+ | (id, ev, deps) as obl :: tl ->
+ let found' = Evar.Set.union found (Evar.Set.singleton id) in
+ if Evar.Set.subset deps found' then
+ aux tl found' (obl :: list)
+ else aux (move_after obl tl) found list
+ | [] -> List.rev list
+ in aux evl Evar.Set.empty []
+
+open Environ
+
+let eterm_obligations env name evm fs ?status t ty =
+ (* 'Serialize' the evars *)
+ let nc = Environ.named_context env in
+ let nc_len = Context.Named.length nc in
+ let evm = Evarutil.nf_evar_map_undefined evm in
+ let evl = Evarutil.non_instantiated evm in
+ let evl = Evar.Map.bindings evl in
+ let evl = List.map (fun (id, ev) -> (id, ev, evar_dependencies evm id)) evl in
+ let sevl = sort_dependencies evl in
+ let evl = List.map (fun (id, ev, _) -> id, ev) sevl in
+ let evn =
+ let i = ref (-1) in
+ List.rev_map (fun (id, ev) -> incr i;
+ (id, (!i, Id.of_string
+ (Id.to_string name ^ "_obligation_" ^ string_of_int (succ !i))),
+ ev)) evl
+ in
+ let evts =
+ (* Remove existential variables in types and build the corresponding products *)
+ List.fold_right
+ (fun (id, (n, nstr), ev) l ->
+ let hyps = Evd.evar_filtered_context ev in
+ let hyps = trunc_named_context nc_len hyps in
+ let evtyp, deps, transp = etype_of_evar l hyps ev.evar_concl in
+ let evtyp, hyps, chop =
+ match chop_product fs evtyp with
+ | Some t -> t, trunc_named_context fs hyps, fs
+ | None -> evtyp, hyps, 0
+ in
+ let loc, k = evar_source id evm in
+ let status = match k with
+ | Evar_kinds.QuestionMark o -> o
+ | _ -> match status with
+ | Some o -> o
+ | None -> Evar_kinds.Define (not (Program.get_proofs_transparency ()))
+ in
+ let force_status, status, chop = match status with
+ | Evar_kinds.Define true as stat ->
+ if not (Int.equal chop fs) then true, Evar_kinds.Define false, None
+ else false, stat, Some chop
+ | s -> false, s, None
+ in
+ let info = { ev_name = (n, nstr);
+ ev_hyps = hyps; ev_status = force_status, status; ev_chop = chop;
+ ev_src = loc, k; ev_typ = evtyp ; ev_deps = deps; ev_tac = None }
+ in (id, info) :: l)
+ evn []
+ in
+ let t', _, transparent = (* Substitute evar refs in the term by variables *)
+ subst_evar_constr evts 0 mkVar t
+ in
+ let ty, _, _ = subst_evar_constr evts 0 mkVar ty in
+ let evars =
+ List.map (fun (ev, info) ->
+ let { ev_name = (_, name); ev_status = force_status, status;
+ ev_src = src; ev_typ = typ; ev_deps = deps; ev_tac = tac } = info
+ in
+ let force_status, status = match status with
+ | Evar_kinds.Define true when Id.Set.mem name transparent ->
+ true, Evar_kinds.Define false
+ | _ -> force_status, status
+ in name, typ, src, (force_status, status), deps, tac) evts
+ in
+ let evnames = List.map (fun (ev, info) -> ev, snd info.ev_name) evts in
+ let evmap f c = pi1 (subst_evar_constr evts 0 f c) in
+ Array.of_list (List.rev evars), (evnames, evmap), t', ty
+
+let tactics_module = ["Program";"Tactics"]
+let safe_init_constant md name () =
+ Coqlib.check_required_library ("Coq"::md);
+ Coqlib.gen_constant "Obligations" md name
+let hide_obligation = safe_init_constant tactics_module "obligation"
+
+let pperror cmd = CErrors.user_err ~hdr:"Program" cmd
+let error s = pperror (str s)
+
+let reduce c =
+ Reductionops.clos_norm_flags CClosure.betaiota (Global.env ()) Evd.empty c
+
+exception NoObligations of Id.t option
+
+let explain_no_obligations = function
+ Some ident -> str "No obligations for program " ++ Id.print ident
+ | None -> str "No obligations remaining"
+
+type obligation_info =
+ (Names.Id.t * Term.types * Evar_kinds.t Loc.located *
+ (bool * Evar_kinds.obligation_definition_status)
+ * Int.Set.t * unit Proofview.tactic option) array
+
+type 'a obligation_body =
+ | DefinedObl of 'a
+ | TermObl of constr
+
+type obligation =
+ { obl_name : Id.t;
+ obl_type : types;
+ obl_location : Evar_kinds.t Loc.located;
+ obl_body : constant obligation_body option;
+ obl_status : bool * Evar_kinds.obligation_definition_status;
+ obl_deps : Int.Set.t;
+ obl_tac : unit Proofview.tactic option;
+ }
+
+type obligations = (obligation array * int)
+
+type fixpoint_kind =
+ | IsFixpoint of (Id.t Loc.located option * Constrexpr.recursion_order_expr) list
+ | IsCoFixpoint
+
+type notations = (Vernacexpr.lstring * Constrexpr.constr_expr * Notation_term.scope_name option) list
+
+type program_info_aux = {
+ prg_name: Id.t;
+ prg_body: constr;
+ prg_type: constr;
+ prg_ctx: Evd.evar_universe_context;
+ prg_pl: Id.t Loc.located list option;
+ prg_obligations: obligations;
+ prg_deps : Id.t list;
+ prg_fixkind : fixpoint_kind option ;
+ prg_implicits : (Constrexpr.explicitation * (bool * bool * bool)) list;
+ prg_notations : notations ;
+ prg_kind : definition_kind;
+ prg_reduce : constr -> constr;
+ prg_hook : (Evd.evar_universe_context -> unit) Lemmas.declaration_hook;
+ prg_opaque : bool;
+ prg_sign: named_context_val;
+}
+
+type program_info = program_info_aux CEphemeron.key
+
+let get_info x =
+ try CEphemeron.get x
+ with CEphemeron.InvalidKey ->
+ CErrors.anomaly Pp.(str "Program obligation can't be accessed by a worker")
+
+let assumption_message = Declare.assumption_message
+
+let default_tactic = ref (Proofview.tclUNIT ())
+
+(* true = hide obligations *)
+let hide_obligations = ref false
+
+let set_hide_obligations = (:=) hide_obligations
+let get_hide_obligations () = !hide_obligations
+
+open Goptions
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "Hidding of Program obligations";
+ optkey = ["Hide";"Obligations"];
+ optread = get_hide_obligations;
+ optwrite = set_hide_obligations; }
+
+let shrink_obligations = ref true
+
+let set_shrink_obligations = (:=) shrink_obligations
+let get_shrink_obligations () = !shrink_obligations
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = true;
+ optname = "Shrinking of Program obligations";
+ optkey = ["Shrink";"Obligations"];
+ optread = get_shrink_obligations;
+ optwrite = set_shrink_obligations; }
+
+let evar_of_obligation o = make_evar (Global.named_context_val ()) o.obl_type
+
+let get_body obl =
+ match obl.obl_body with
+ | None -> None
+ | Some (DefinedObl c) ->
+ let ctx = Environ.constant_context (Global.env ()) c in
+ let pc = (c, Univ.UContext.instance ctx) in
+ Some (DefinedObl pc)
+ | Some (TermObl c) ->
+ Some (TermObl c)
+
+let get_obligation_body expand obl =
+ match get_body obl with
+ | None -> None
+ | Some c ->
+ if expand && snd obl.obl_status == Evar_kinds.Expand then
+ match c with
+ | DefinedObl pc -> Some (constant_value_in (Global.env ()) pc)
+ | TermObl c -> Some c
+ else (match c with
+ | DefinedObl pc -> Some (mkConstU pc)
+ | TermObl c -> Some c)
+
+let obl_substitution expand obls deps =
+ Int.Set.fold
+ (fun x acc ->
+ let xobl = obls.(x) in
+ match get_obligation_body expand xobl with
+ | None -> acc
+ | Some oblb -> (xobl.obl_name, (xobl.obl_type, oblb)) :: acc)
+ deps []
+
+let subst_deps expand obls deps t =
+ let osubst = obl_substitution expand obls deps in
+ (Vars.replace_vars (List.map (fun (n, (_, b)) -> n, b) osubst) t)
+
+let rec prod_app t n =
+ match kind_of_term (Termops.strip_outer_cast t) with
+ | Prod (_,_,b) -> subst1 n b
+ | LetIn (_, b, t, b') -> prod_app (subst1 b b') n
+ | _ ->
+ user_err ~hdr:"prod_app"
+ (str"Needed a product, but didn't find one" ++ fnl ())
+
+
+(* prod_appvect T [| a1 ; ... ; an |] -> (T a1 ... an) *)
+let prod_applist t nL = List.fold_left prod_app t nL
+
+let replace_appvars subst =
+ let rec aux c =
+ let f, l = decompose_app c in
+ if isVar f then
+ try
+ let c' = List.map (map_constr aux) l in
+ let (t, b) = Id.List.assoc (destVar f) subst in
+ mkApp (delayed_force hide_obligation,
+ [| prod_applist t c'; applistc b c' |])
+ with Not_found -> map_constr aux c
+ else map_constr aux c
+ in map_constr aux
+
+let subst_prog expand obls ints prg =
+ let subst = obl_substitution expand obls ints in
+ if get_hide_obligations () then
+ (replace_appvars subst prg.prg_body,
+ replace_appvars subst ((* Termops.refresh_universes *) prg.prg_type))
+ else
+ let subst' = List.map (fun (n, (_, b)) -> n, b) subst in
+ (Vars.replace_vars subst' prg.prg_body,
+ Vars.replace_vars subst' ((* Termops.refresh_universes *) prg.prg_type))
+
+let subst_deps_obl obls obl =
+ let t' = subst_deps true obls obl.obl_deps obl.obl_type in
+ { obl with obl_type = t' }
+
+module ProgMap = Id.Map
+
+let map_replace k v m = ProgMap.add k (CEphemeron.create v) (ProgMap.remove k m)
+
+let map_keys m = ProgMap.fold (fun k _ l -> k :: l) m []
+
+let from_prg : program_info ProgMap.t ref =
+ Summary.ref ProgMap.empty ~name:"program-tcc-table"
+
+let close sec =
+ if not (ProgMap.is_empty !from_prg) then
+ let keys = map_keys !from_prg in
+ user_err ~hdr:"Program"
+ (str "Unsolved obligations when closing " ++ str sec ++ str":" ++ spc () ++
+ prlist_with_sep spc (fun x -> Nameops.pr_id x) keys ++
+ (str (if Int.equal (List.length keys) 1 then " has " else " have ") ++
+ str "unsolved obligations"))
+
+let input : program_info ProgMap.t -> obj =
+ declare_object
+ { (default_object "Program state") with
+ cache_function = (fun (na, pi) -> from_prg := pi);
+ load_function = (fun _ (_, pi) -> from_prg := pi);
+ discharge_function = (fun _ -> close "section"; None);
+ classify_function = (fun _ -> close "module"; Dispose) }
+
+open Evd
+
+let progmap_remove prg =
+ Lib.add_anonymous_leaf (input (ProgMap.remove prg.prg_name !from_prg))
+
+let progmap_add n prg =
+ Lib.add_anonymous_leaf (input (ProgMap.add n prg !from_prg))
+
+let progmap_replace prg' =
+ Lib.add_anonymous_leaf (input (map_replace prg'.prg_name prg' !from_prg))
+
+let rec intset_to = function
+ -1 -> Int.Set.empty
+ | n -> Int.Set.add n (intset_to (pred n))
+
+let subst_body expand prg =
+ let obls, _ = prg.prg_obligations in
+ let ints = intset_to (pred (Array.length obls)) in
+ subst_prog expand obls ints prg
+
+let declare_definition prg =
+ let body, typ = subst_body true prg in
+ let nf = Universes.nf_evars_and_universes_opt_subst (fun x -> None)
+ (Evd.evar_universe_context_subst prg.prg_ctx) in
+ let opaque = prg.prg_opaque in
+ let fix_exn = Stm.get_fix_exn () in
+ let pl, ctx =
+ Evd.universe_context ?names:prg.prg_pl (Evd.from_ctx prg.prg_ctx) in
+ let ce =
+ definition_entry ~fix_exn
+ ~opaque ~types:(nf typ) ~poly:(pi2 prg.prg_kind)
+ ~univs:(Evd.evar_context_universe_context prg.prg_ctx) (nf body)
+ in
+ let () = progmap_remove prg in
+ let cst =
+ !declare_definition_ref prg.prg_name
+ prg.prg_kind ce prg.prg_implicits
+ (Lemmas.mk_hook (fun l r -> Lemmas.call_hook fix_exn prg.prg_hook l r prg.prg_ctx; r))
+ in
+ Universes.register_universe_binders cst pl;
+ cst
+
+open Pp
+
+let rec lam_index n t acc =
+ match kind_of_term t with
+ | Lambda (Name n', _, _) when Id.equal n n' ->
+ acc
+ | Lambda (_, _, b) ->
+ lam_index n b (succ acc)
+ | _ -> raise Not_found
+
+let compute_possible_guardness_evidences (n,_) fixbody fixtype =
+ match n with
+ | Some (loc, n) -> [lam_index n fixbody 0]
+ | None ->
+ (* If recursive argument was not given by user, we try all args.
+ An earlier approach was to look only for inductive arguments,
+ but doing it properly involves delta-reduction, and it finally
+ doesn't seem to worth the effort (except for huge mutual
+ fixpoints ?) *)
+ let m = Termops.nb_prod fixtype in
+ let ctx = fst (decompose_prod_n_assum m fixtype) in
+ List.map_i (fun i _ -> i) 0 ctx
+
+let mk_proof c = ((c, Univ.ContextSet.empty), Safe_typing.empty_private_constants)
+
+let declare_mutual_definition l =
+ let len = List.length l in
+ let first = List.hd l in
+ let fixdefs, fixtypes, fiximps =
+ List.split3
+ (List.map (fun x ->
+ let subs, typ = (subst_body true x) in
+ let term = snd (Reductionops.splay_lam_n (Global.env ()) Evd.empty len subs) in
+ let typ = snd (Reductionops.splay_prod_n (Global.env ()) Evd.empty len typ) in
+ x.prg_reduce term, x.prg_reduce typ, x.prg_implicits) l)
+ in
+(* let fixdefs = List.map reduce_fix fixdefs in *)
+ let fixkind = Option.get first.prg_fixkind in
+ let arrrec, recvec = Array.of_list fixtypes, Array.of_list fixdefs in
+ let fixdecls = (Array.of_list (List.map (fun x -> Name x.prg_name) l), arrrec, recvec) in
+ let (local,poly,kind) = first.prg_kind in
+ let fixnames = first.prg_deps in
+ let opaque = first.prg_opaque in
+ let kind = if fixkind != IsCoFixpoint then Fixpoint else CoFixpoint in
+ let indexes, fixdecls =
+ match fixkind with
+ | IsFixpoint wfl ->
+ let possible_indexes =
+ List.map3 compute_possible_guardness_evidences
+ wfl fixdefs fixtypes in
+ let indexes =
+ Pretyping.search_guard
+ Loc.ghost (Global.env())
+ possible_indexes fixdecls in
+ Some indexes,
+ List.map_i (fun i _ ->
+ mk_proof (mkFix ((indexes,i),fixdecls))) 0 l
+ | IsCoFixpoint ->
+ None,
+ List.map_i (fun i _ ->
+ mk_proof (mkCoFix (i,fixdecls))) 0 l
+ in
+ (* Declare the recursive definitions *)
+ let ctx = Evd.evar_context_universe_context first.prg_ctx in
+ let fix_exn = Stm.get_fix_exn () in
+ let kns = List.map4 (!declare_fix_ref ~opaque (local, poly, kind) ctx)
+ fixnames fixdecls fixtypes fiximps in
+ (* Declare notations *)
+ List.iter Metasyntax.add_notation_interpretation first.prg_notations;
+ Declare.recursive_message (fixkind != IsCoFixpoint) indexes fixnames;
+ let gr = List.hd kns in
+ let kn = match gr with ConstRef kn -> kn | _ -> assert false in
+ Lemmas.call_hook fix_exn first.prg_hook local gr first.prg_ctx;
+ List.iter progmap_remove l; kn
+
+let decompose_lam_prod c ty =
+ let open Context.Rel.Declaration in
+ let rec aux ctx c ty =
+ match kind_of_term c, kind_of_term ty with
+ | LetIn (x, b, t, c), LetIn (x', b', t', ty)
+ when eq_constr b b' && eq_constr t t' ->
+ let ctx' = Context.Rel.add (LocalDef (x,b',t')) ctx in
+ aux ctx' c ty
+ | _, LetIn (x', b', t', ty) ->
+ let ctx' = Context.Rel.add (LocalDef (x',b',t')) ctx in
+ aux ctx' (lift 1 c) ty
+ | LetIn (x, b, t, c), _ ->
+ let ctx' = Context.Rel.add (LocalDef (x,b,t)) ctx in
+ aux ctx' c (lift 1 ty)
+ | Lambda (x, b, t), Prod (x', b', t')
+ (* By invariant, must be convertible *) ->
+ let ctx' = Context.Rel.add (LocalAssum (x,b')) ctx in
+ aux ctx' t t'
+ | Cast (c, _, _), _ -> aux ctx c ty
+ | _, _ -> ctx, c, ty
+ in aux Context.Rel.empty c ty
+
+let shrink_body c ty =
+ let ctx, b, ty =
+ match ty with
+ | None ->
+ let ctx, b = decompose_lam_assum c in
+ ctx, b, None
+ | Some ty ->
+ let ctx, b, ty = decompose_lam_prod c ty in
+ ctx, b, Some ty
+ in
+ let b', ty', n, args =
+ List.fold_left (fun (b, ty, i, args) decl ->
+ if noccurn 1 b && Option.cata (noccurn 1) true ty then
+ subst1 mkProp b, Option.map (subst1 mkProp) ty, succ i, args
+ else
+ let open Context.Rel.Declaration in
+ let args = if is_local_assum decl then mkRel i :: args else args in
+ mkLambda_or_LetIn decl b, Option.map (mkProd_or_LetIn decl) ty,
+ succ i, args)
+ (b, ty, 1, []) ctx
+ in ctx, b', ty', Array.of_list args
+
+let unfold_entry cst = Hints.HintsUnfoldEntry [EvalConstRef cst]
+
+let add_hint local prg cst =
+ Hints.add_hints local [Id.to_string prg.prg_name] (unfold_entry cst)
+
+let declare_obligation prg obl body ty uctx =
+ let body = prg.prg_reduce body in
+ let ty = Option.map prg.prg_reduce ty in
+ match obl.obl_status with
+ | _, Evar_kinds.Expand -> false, { obl with obl_body = Some (TermObl body) }
+ | force, Evar_kinds.Define opaque ->
+ let opaque = not force && opaque in
+ let poly = pi2 prg.prg_kind in
+ let ctx, body, ty, args =
+ if get_shrink_obligations () && not poly then
+ shrink_body body ty else [], body, ty, [||]
+ in
+ let body = ((body,Univ.ContextSet.empty),Safe_typing.empty_private_constants) in
+ let ce =
+ { const_entry_body = Future.from_val ~fix_exn:(fun x -> x) body;
+ const_entry_secctx = None;
+ const_entry_type = ty;
+ const_entry_polymorphic = poly;
+ const_entry_universes = uctx;
+ const_entry_opaque = opaque;
+ const_entry_inline_code = false;
+ const_entry_feedback = None;
+ } in
+ (** ppedrot: seems legit to have obligations as local *)
+ let constant = Declare.declare_constant obl.obl_name ~local:true
+ (DefinitionEntry ce,IsProof Property)
+ in
+ if not opaque then add_hint false prg constant;
+ definition_message obl.obl_name;
+ true, { obl with obl_body =
+ if poly then
+ Some (DefinedObl constant)
+ else
+ Some (TermObl (it_mkLambda_or_LetIn (mkApp (mkConst constant, args)) ctx)) }
+
+let init_prog_info ?(opaque = false) sign n pl b t ctx deps fixkind
+ notations obls impls kind reduce hook =
+ let obls', b =
+ match b with
+ | None ->
+ assert(Int.equal (Array.length obls) 0);
+ let n = Nameops.add_suffix n "_obligation" in
+ [| { obl_name = n; obl_body = None;
+ obl_location = Loc.ghost, Evar_kinds.InternalHole; obl_type = t;
+ obl_status = false, Evar_kinds.Expand; obl_deps = Int.Set.empty;
+ obl_tac = None } |],
+ mkVar n
+ | Some b ->
+ Array.mapi
+ (fun i (n, t, l, o, d, tac) ->
+ { obl_name = n ; obl_body = None;
+ obl_location = l; obl_type = t; obl_status = o;
+ obl_deps = d; obl_tac = tac })
+ obls, b
+ in
+ { prg_name = n ; prg_body = b; prg_type = reduce t;
+ prg_ctx = ctx; prg_pl = pl;
+ prg_obligations = (obls', Array.length obls');
+ prg_deps = deps; prg_fixkind = fixkind ; prg_notations = notations ;
+ prg_implicits = impls; prg_kind = kind; prg_reduce = reduce;
+ prg_hook = hook; prg_opaque = opaque;
+ prg_sign = sign }
+
+let map_cardinal m =
+ let i = ref 0 in
+ ProgMap.iter (fun _ v ->
+ if snd (CEphemeron.get v).prg_obligations > 0 then incr i) m;
+ !i
+
+exception Found of program_info
+
+let map_first m =
+ try
+ ProgMap.iter (fun _ v ->
+ if snd (CEphemeron.get v).prg_obligations > 0 then
+ raise (Found v)) m;
+ assert(false)
+ with Found x -> x
+
+let get_prog name =
+ let prg_infos = !from_prg in
+ match name with
+ Some n ->
+ (try get_info (ProgMap.find n prg_infos)
+ with Not_found -> raise (NoObligations (Some n)))
+ | None ->
+ (let n = map_cardinal prg_infos in
+ match n with
+ 0 -> raise (NoObligations None)
+ | 1 -> get_info (map_first prg_infos)
+ | _ ->
+ let progs = Id.Set.elements (ProgMap.domain prg_infos) in
+ let prog = List.hd progs in
+ let progs = prlist_with_sep pr_comma Nameops.pr_id progs in
+ user_err
+ (str "More than one program with unsolved obligations: " ++ progs
+ ++ str "; use the \"of\" clause to specify, as in \"Obligation 1 of " ++ Nameops.pr_id prog ++ str "\""))
+
+let get_any_prog () =
+ let prg_infos = !from_prg in
+ let n = map_cardinal prg_infos in
+ if n > 0 then get_info (map_first prg_infos)
+ else raise (NoObligations None)
+
+let get_prog_err n =
+ try get_prog n with NoObligations id -> pperror (explain_no_obligations id)
+
+let get_any_prog_err () =
+ try get_any_prog () with NoObligations id -> pperror (explain_no_obligations id)
+
+let obligations_solved prg = Int.equal (snd prg.prg_obligations) 0
+
+let all_programs () =
+ ProgMap.fold (fun k p l -> p :: l) !from_prg []
+
+type progress =
+ | Remain of int
+ | Dependent
+ | Defined of global_reference
+
+let obligations_message rem =
+ if rem > 0 then
+ if Int.equal rem 1 then
+ Flags.if_verbose Feedback.msg_info (int rem ++ str " obligation remaining")
+ else
+ Flags.if_verbose Feedback.msg_info (int rem ++ str " obligations remaining")
+ else
+ Flags.if_verbose Feedback.msg_info (str "No more obligations remaining")
+
+let update_obls prg obls rem =
+ let prg' = { prg with prg_obligations = (obls, rem) } in
+ progmap_replace prg';
+ obligations_message rem;
+ if rem > 0 then Remain rem
+ else (
+ match prg'.prg_deps with
+ | [] ->
+ let kn = declare_definition prg' in
+ progmap_remove prg';
+ Defined kn
+ | l ->
+ let progs = List.map (fun x -> get_info (ProgMap.find x !from_prg)) prg'.prg_deps in
+ if List.for_all (fun x -> obligations_solved x) progs then
+ let kn = declare_mutual_definition progs in
+ Defined (ConstRef kn)
+ else Dependent)
+
+let is_defined obls x = not (Option.is_empty obls.(x).obl_body)
+
+let deps_remaining obls deps =
+ Int.Set.fold
+ (fun x acc ->
+ if is_defined obls x then acc
+ else x :: acc)
+ deps []
+
+let dependencies obls n =
+ let res = ref Int.Set.empty in
+ Array.iteri
+ (fun i obl ->
+ if not (Int.equal i n) && Int.Set.mem n obl.obl_deps then
+ res := Int.Set.add i !res)
+ obls;
+ !res
+
+let goal_kind poly = Decl_kinds.Local, poly, Decl_kinds.DefinitionBody Decl_kinds.Definition
+
+let goal_proof_kind poly = Decl_kinds.Local, poly, Decl_kinds.Proof Decl_kinds.Lemma
+
+let kind_of_obligation poly o =
+ match o with
+ | Evar_kinds.Define false | Evar_kinds.Expand -> goal_kind poly
+ | _ -> goal_proof_kind poly
+
+let not_transp_msg =
+ str "Obligation should be transparent but was declared opaque." ++ spc () ++
+ str"Use 'Defined' instead."
+
+let err_not_transp () = pperror not_transp_msg
+
+let rec string_of_list sep f = function
+ [] -> ""
+ | x :: [] -> f x
+ | x :: ((y :: _) as tl) -> f x ^ sep ^ string_of_list sep f tl
+
+(* Solve an obligation using tactics, return the corresponding proof term *)
+
+let solve_by_tac name evi t poly ctx =
+ let id = name in
+ let concl = evi.evar_concl in
+ (* spiwack: the status is dropped. *)
+ let (entry,_,ctx') = Pfedit.build_constant_by_tactic
+ id ~goal_kind:(goal_kind poly) ctx evi.evar_hyps concl (Tacticals.New.tclCOMPLETE t) in
+ let env = Global.env () in
+ let entry = Safe_typing.inline_private_constants_in_definition_entry env entry in
+ let body, eff = Future.force entry.const_entry_body in
+ assert(Safe_typing.empty_private_constants = eff);
+ let ctx' = Evd.merge_context_set ~sideff:true Evd.univ_rigid (Evd.from_ctx ctx') (snd body) in
+ Inductiveops.control_only_guard (Global.env ()) (fst body);
+ (fst body), entry.const_entry_type, Evd.evar_universe_context ctx'
+
+let obligation_terminator name num guard hook auto pf =
+ let open Proof_global in
+ let term = Lemmas.universe_proof_terminator guard hook in
+ match pf with
+ | Admitted _ -> apply_terminator term pf
+ | Proved (opq, id, proof) ->
+ if not !shrink_obligations then apply_terminator term pf
+ else
+ let (_, (entry, uctx, _)) = Pfedit.cook_this_proof proof in
+ let env = Global.env () in
+ let entry = Safe_typing.inline_private_constants_in_definition_entry env entry in
+ let ty = entry.Entries.const_entry_type in
+ let (body, cstr), eff = Future.force entry.Entries.const_entry_body in
+ assert(Safe_typing.empty_private_constants = eff);
+ let sigma = Evd.from_ctx (fst uctx) in
+ let sigma = Evd.merge_context_set ~sideff:true Evd.univ_rigid sigma cstr in
+ Inductiveops.control_only_guard (Global.env ()) body;
+ (** Declare the obligation ourselves and drop the hook *)
+ let prg = get_info (ProgMap.find name !from_prg) in
+ let ctx = Evd.evar_universe_context sigma in
+ let prg = { prg with prg_ctx = ctx } in
+ let obls, rem = prg.prg_obligations in
+ let obl = obls.(num) in
+ let status =
+ match obl.obl_status, opq with
+ | (_, Evar_kinds.Expand), Vernacexpr.Opaque _ -> err_not_transp ()
+ | (true, _), Vernacexpr.Opaque _ -> err_not_transp ()
+ | (false, _), Vernacexpr.Opaque _ -> Evar_kinds.Define true
+ | (_, Evar_kinds.Define true), Vernacexpr.Transparent -> Evar_kinds.Define false
+ | (_, status), Vernacexpr.Transparent -> status
+ in
+ let obl = { obl with obl_status = false, status } in
+ let uctx = Evd.evar_context_universe_context ctx in
+ let (_, obl) = declare_obligation prg obl body ty uctx in
+ let obls = Array.copy obls in
+ let _ = obls.(num) <- obl in
+ try
+ ignore (update_obls prg obls (pred rem));
+ if pred rem > 0 then
+ begin
+ let deps = dependencies obls num in
+ if not (Int.Set.is_empty deps) then
+ ignore (auto (Some name) None deps)
+ end
+ with e when CErrors.noncritical e ->
+ let e = CErrors.push e in
+ pperror (CErrors.iprint (ExplainErr.process_vernac_interp_error e))
+
+let obligation_hook prg obl num auto ctx' _ gr =
+ let obls, rem = prg.prg_obligations in
+ let cst = match gr with ConstRef cst -> cst | _ -> assert false in
+ let transparent = evaluable_constant cst (Global.env ()) in
+ let () = match obl.obl_status with
+ (true, Evar_kinds.Expand)
+ | (true, Evar_kinds.Define true) ->
+ if not transparent then err_not_transp ()
+ | _ -> ()
+in
+ let obl = { obl with obl_body = Some (DefinedObl cst) } in
+ let () = if transparent then add_hint true prg cst in
+ let obls = Array.copy obls in
+ let _ = obls.(num) <- obl in
+ let ctx' = match ctx' with None -> prg.prg_ctx | Some ctx' -> ctx' in
+ let ctx' =
+ if not (pi2 prg.prg_kind) (* Not polymorphic *) then
+ (* The universe context was declared globally, we continue
+ from the new global environment. *)
+ let evd = Evd.from_env (Global.env ()) in
+ let ctx' = Evd.merge_universe_subst evd (Evd.universe_subst (Evd.from_ctx ctx')) in
+ Evd.evar_universe_context ctx'
+ else ctx'
+ in
+ let prg = { prg with prg_ctx = ctx' } in
+ let () =
+ try ignore (update_obls prg obls (pred rem))
+ with e when CErrors.noncritical e ->
+ let e = CErrors.push e in
+ pperror (CErrors.iprint (ExplainErr.process_vernac_interp_error e))
+ in
+ if pred rem > 0 then begin
+ let deps = dependencies obls num in
+ if not (Int.Set.is_empty deps) then
+ ignore (auto (Some prg.prg_name) None deps)
+ end
+
+let rec solve_obligation prg num tac =
+ let user_num = succ num in
+ let obls, rem = prg.prg_obligations in
+ let obl = obls.(num) in
+ let remaining = deps_remaining obls obl.obl_deps in
+ let () =
+ if not (Option.is_empty obl.obl_body) then
+ pperror (str "Obligation" ++ spc () ++ int user_num ++ str "already" ++ spc() ++ str "solved.");
+ if not (List.is_empty remaining) then
+ pperror (str "Obligation " ++ int user_num ++ str " depends on obligation(s) "
+ ++ str (string_of_list ", " (fun x -> string_of_int (succ x)) remaining));
+ in
+ let obl = subst_deps_obl obls obl in
+ let kind = kind_of_obligation (pi2 prg.prg_kind) (snd obl.obl_status) in
+ let evd = Evd.from_ctx prg.prg_ctx in
+ let evd = Evd.update_sigma_env evd (Global.env ()) in
+ let auto n tac oblset = auto_solve_obligations n ~oblset tac in
+ let terminator guard hook =
+ Proof_global.make_terminator
+ (obligation_terminator prg.prg_name num guard hook auto) in
+ let hook ctx = Lemmas.mk_hook (obligation_hook prg obl num auto ctx) in
+ let () = Lemmas.start_proof_univs ~sign:prg.prg_sign obl.obl_name kind evd obl.obl_type ~terminator hook in
+ let _ = Pfedit.by !default_tactic in
+ Option.iter (fun tac -> Pfedit.set_end_tac tac) tac
+
+and obligation (user_num, name, typ) tac =
+ let num = pred user_num in
+ let prg = get_prog_err name in
+ let obls, rem = prg.prg_obligations in
+ if num < Array.length obls then
+ let obl = obls.(num) in
+ match obl.obl_body with
+ None -> solve_obligation prg num tac
+ | Some r -> error "Obligation already solved"
+ else error (sprintf "Unknown obligation number %i" (succ num))
+
+
+and solve_obligation_by_tac prg obls i tac =
+ let obl = obls.(i) in
+ match obl.obl_body with
+ | Some _ -> None
+ | None ->
+ try
+ if List.is_empty (deps_remaining obls obl.obl_deps) then
+ let obl = subst_deps_obl obls obl in
+ let tac =
+ match tac with
+ | Some t -> t
+ | None ->
+ match obl.obl_tac with
+ | Some t -> t
+ | None -> !default_tactic
+ in
+ let evd = Evd.from_ctx prg.prg_ctx in
+ let evd = Evd.update_sigma_env evd (Global.env ()) in
+ let t, ty, ctx =
+ solve_by_tac obl.obl_name (evar_of_obligation obl) tac
+ (pi2 prg.prg_kind) (Evd.evar_universe_context evd)
+ in
+ let uctx = Evd.evar_context_universe_context ctx in
+ let prg = {prg with prg_ctx = ctx} in
+ let def, obl' = declare_obligation prg obl t ty uctx in
+ obls.(i) <- obl';
+ if def && not (pi2 prg.prg_kind) then (
+ (* Declare the term constraints with the first obligation only *)
+ let evd = Evd.from_env (Global.env ()) in
+ let evd = Evd.merge_universe_subst evd (Evd.universe_subst (Evd.from_ctx ctx)) in
+ let ctx' = Evd.evar_universe_context evd in
+ Some {prg with prg_ctx = ctx'})
+ else Some prg
+ else None
+ with e when CErrors.noncritical e ->
+ let (e, _) = CErrors.push e in
+ match e with
+ | Refiner.FailError (_, s) ->
+ user_err ~loc:(fst obl.obl_location) ~hdr:"solve_obligation" (Lazy.force s)
+ | e -> None (* FIXME really ? *)
+
+and solve_prg_obligations prg ?oblset tac =
+ let obls, rem = prg.prg_obligations in
+ let rem = ref rem in
+ let obls' = Array.copy obls in
+ let set = ref Int.Set.empty in
+ let p = match oblset with
+ | None -> (fun _ -> true)
+ | Some s -> set := s;
+ (fun i -> Int.Set.mem i !set)
+ in
+ let prgref = ref prg in
+ let _ =
+ Array.iteri (fun i x ->
+ if p i then
+ match solve_obligation_by_tac !prgref obls' i tac with
+ | None -> ()
+ | Some prg' ->
+ prgref := prg';
+ let deps = dependencies obls i in
+ (set := Int.Set.union !set deps;
+ decr rem))
+ obls'
+ in
+ update_obls !prgref obls' !rem
+
+and solve_obligations n tac =
+ let prg = get_prog_err n in
+ solve_prg_obligations prg tac
+
+and solve_all_obligations tac =
+ ProgMap.iter (fun k v -> ignore(solve_prg_obligations (get_info v) tac)) !from_prg
+
+and try_solve_obligation n prg tac =
+ let prg = get_prog prg in
+ let obls, rem = prg.prg_obligations in
+ let obls' = Array.copy obls in
+ match solve_obligation_by_tac prg obls' n tac with
+ | Some prg' -> ignore(update_obls prg' obls' (pred rem))
+ | None -> ()
+
+and try_solve_obligations n tac =
+ try ignore (solve_obligations n tac) with NoObligations _ -> ()
+
+and auto_solve_obligations n ?oblset tac : progress =
+ Flags.if_verbose Feedback.msg_info (str "Solving obligations automatically...");
+ try solve_prg_obligations (get_prog_err n) ?oblset tac with NoObligations _ -> Dependent
+
+open Pp
+let show_obligations_of_prg ?(msg=true) prg =
+ let n = prg.prg_name in
+ let obls, rem = prg.prg_obligations in
+ let showed = ref 5 in
+ if msg then Feedback.msg_info (int rem ++ str " obligation(s) remaining: ");
+ Array.iteri (fun i x ->
+ match x.obl_body with
+ | None ->
+ if !showed > 0 then (
+ decr showed;
+ let x = subst_deps_obl obls x in
+ Feedback.msg_info (str "Obligation" ++ spc() ++ int (succ i) ++ spc () ++
+ str "of" ++ spc() ++ Id.print n ++ str ":" ++ spc () ++
+ hov 1 (Printer.pr_constr_env (Global.env ()) Evd.empty x.obl_type ++
+ str "." ++ fnl ())))
+ | Some _ -> ())
+ obls
+
+let show_obligations ?(msg=true) n =
+ let progs = match n with
+ | None -> all_programs ()
+ | Some n ->
+ try [ProgMap.find n !from_prg]
+ with Not_found -> raise (NoObligations (Some n))
+ in List.iter (fun x -> show_obligations_of_prg ~msg (get_info x)) progs
+
+let show_term n =
+ let prg = get_prog_err n in
+ let n = prg.prg_name in
+ (Id.print n ++ spc () ++ str":" ++ spc () ++
+ Printer.pr_constr_env (Global.env ()) Evd.empty prg.prg_type ++ spc () ++ str ":=" ++ fnl ()
+ ++ Printer.pr_constr_env (Global.env ()) Evd.empty prg.prg_body)
+
+let add_definition n ?term t ctx ?pl ?(implicits=[]) ?(kind=Global,false,Definition) ?tactic
+ ?(reduce=reduce) ?(hook=Lemmas.mk_hook (fun _ _ _ -> ())) ?(opaque = false) obls =
+ let sign = Decls.initialize_named_context_for_proof () in
+ let info = Id.print n ++ str " has type-checked" in
+ let prg = init_prog_info sign ~opaque n pl term t ctx [] None [] obls implicits kind reduce hook in
+ let obls,_ = prg.prg_obligations in
+ if Int.equal (Array.length obls) 0 then (
+ Flags.if_verbose Feedback.msg_info (info ++ str ".");
+ let cst = declare_definition prg in
+ Defined cst)
+ else (
+ let len = Array.length obls in
+ let _ = Flags.if_verbose Feedback.msg_info (info ++ str ", generating " ++ int len ++ str " obligation(s)") in
+ progmap_add n (CEphemeron.create prg);
+ let res = auto_solve_obligations (Some n) tactic in
+ match res with
+ | Remain rem -> Flags.if_verbose (fun () -> show_obligations ~msg:false (Some n)) (); res
+ | _ -> res)
+
+let add_mutual_definitions l ctx ?pl ?tactic ?(kind=Global,false,Definition) ?(reduce=reduce)
+ ?(hook=Lemmas.mk_hook (fun _ _ _ -> ())) ?(opaque = false) notations fixkind =
+ let sign = Decls.initialize_named_context_for_proof () in
+ let deps = List.map (fun (n, b, t, imps, obls) -> n) l in
+ List.iter
+ (fun (n, b, t, imps, obls) ->
+ let prg = init_prog_info sign ~opaque n pl (Some b) t ctx deps (Some fixkind)
+ notations obls imps kind reduce hook
+ in progmap_add n (CEphemeron.create prg)) l;
+ let _defined =
+ List.fold_left (fun finished x ->
+ if finished then finished
+ else
+ let res = auto_solve_obligations (Some x) tactic in
+ match res with
+ | Defined _ ->
+ (* If one definition is turned into a constant,
+ the whole block is defined. *) true
+ | _ -> false)
+ false deps
+ in ()
+
+let admit_prog prg =
+ let obls, rem = prg.prg_obligations in
+ let obls = Array.copy obls in
+ Array.iteri
+ (fun i x ->
+ match x.obl_body with
+ | None ->
+ let x = subst_deps_obl obls x in
+ let ctx = Evd.evar_context_universe_context prg.prg_ctx in
+ let kn = Declare.declare_constant x.obl_name ~local:true
+ (ParameterEntry (None,false,(x.obl_type,ctx),None), IsAssumption Conjectural)
+ in
+ assumption_message x.obl_name;
+ obls.(i) <- { x with obl_body = Some (DefinedObl kn) }
+ | Some _ -> ())
+ obls;
+ ignore(update_obls prg obls 0)
+
+let rec admit_all_obligations () =
+ let prg = try Some (get_any_prog ()) with NoObligations _ -> None in
+ match prg with
+ | None -> ()
+ | Some prg ->
+ admit_prog prg;
+ admit_all_obligations ()
+
+let admit_obligations n =
+ match n with
+ | None -> admit_all_obligations ()
+ | Some _ ->
+ let prg = get_prog_err n in
+ admit_prog prg
+
+let next_obligation n tac =
+ let prg = match n with
+ | None -> get_any_prog_err ()
+ | Some _ -> get_prog_err n
+ in
+ let obls, rem = prg.prg_obligations in
+ let is_open _ x = Option.is_empty x.obl_body && List.is_empty (deps_remaining obls x.obl_deps) in
+ let i = match Array.findi is_open obls with
+ | Some i -> i
+ | None -> anomaly (Pp.str "Could not find a solvable obligation.")
+ in
+ solve_obligation prg i tac
+
+let init_program () =
+ Coqlib.check_required_library Coqlib.datatypes_module_name;
+ Coqlib.check_required_library ["Coq";"Init";"Specif"];
+ Coqlib.check_required_library ["Coq";"Program";"Tactics"]
+
+
+let set_program_mode c =
+ if c then
+ if !Flags.program_mode then ()
+ else begin
+ init_program ();
+ Flags.program_mode := true;
+ end
diff --git a/vernac/obligations.mli b/vernac/obligations.mli
new file mode 100644
index 000000000..80b689144
--- /dev/null
+++ b/vernac/obligations.mli
@@ -0,0 +1,113 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Environ
+open Term
+open Evd
+open Names
+open Pp
+open Globnames
+open Vernacexpr
+open Decl_kinds
+
+(** Forward declaration. *)
+val declare_fix_ref : (?opaque:bool -> definition_kind -> Univ.universe_context -> Id.t ->
+ Safe_typing.private_constants Entries.proof_output -> types -> Impargs.manual_implicits -> global_reference) ref
+
+val declare_definition_ref :
+ (Id.t -> definition_kind ->
+ Safe_typing.private_constants Entries.definition_entry -> Impargs.manual_implicits
+ -> global_reference Lemmas.declaration_hook -> global_reference) ref
+
+val check_evars : env -> evar_map -> unit
+
+val evar_dependencies : evar_map -> Evar.t -> Evar.Set.t
+val sort_dependencies : (Evar.t * evar_info * Evar.Set.t) list -> (Evar.t * evar_info * Evar.Set.t) list
+
+(* env, id, evars, number of function prototypes to try to clear from
+ evars contexts, object and type *)
+val eterm_obligations : env -> Id.t -> evar_map -> int ->
+ ?status:Evar_kinds.obligation_definition_status -> constr -> types ->
+ (Id.t * types * Evar_kinds.t Loc.located *
+ (bool * Evar_kinds.obligation_definition_status) * Int.Set.t *
+ unit Proofview.tactic option) array
+ (* Existential key, obl. name, type as product,
+ location of the original evar, associated tactic,
+ status and dependencies as indexes into the array *)
+ * ((existential_key * Id.t) list * ((Id.t -> constr) -> constr -> constr)) *
+ constr * types
+ (* Translations from existential identifiers to obligation identifiers
+ and for terms with existentials to closed terms, given a
+ translation from obligation identifiers to constrs, new term, new type *)
+
+type obligation_info =
+ (Id.t * Term.types * Evar_kinds.t Loc.located *
+ (bool * Evar_kinds.obligation_definition_status) * Int.Set.t * unit Proofview.tactic option) array
+ (* ident, type, location, (opaque or transparent, expand or define),
+ dependencies, tactic to solve it *)
+
+type progress = (* Resolution status of a program *)
+ | Remain of int (* n obligations remaining *)
+ | Dependent (* Dependent on other definitions *)
+ | Defined of global_reference (* Defined as id *)
+
+val default_tactic : unit Proofview.tactic ref
+
+val add_definition : Names.Id.t -> ?term:Term.constr -> Term.types ->
+ Evd.evar_universe_context ->
+ ?pl:(Id.t Loc.located list) -> (* Universe binders *)
+ ?implicits:(Constrexpr.explicitation * (bool * bool * bool)) list ->
+ ?kind:Decl_kinds.definition_kind ->
+ ?tactic:unit Proofview.tactic ->
+ ?reduce:(Term.constr -> Term.constr) ->
+ ?hook:(Evd.evar_universe_context -> unit) Lemmas.declaration_hook -> ?opaque:bool -> obligation_info -> progress
+
+type notations =
+ (Vernacexpr.lstring * Constrexpr.constr_expr * Notation_term.scope_name option) list
+
+type fixpoint_kind =
+ | IsFixpoint of (Id.t Loc.located option * Constrexpr.recursion_order_expr) list
+ | IsCoFixpoint
+
+val add_mutual_definitions :
+ (Names.Id.t * Term.constr * Term.types *
+ (Constrexpr.explicitation * (bool * bool * bool)) list * obligation_info) list ->
+ Evd.evar_universe_context ->
+ ?pl:(Id.t Loc.located list) -> (* Universe binders *)
+ ?tactic:unit Proofview.tactic ->
+ ?kind:Decl_kinds.definition_kind ->
+ ?reduce:(Term.constr -> Term.constr) ->
+ ?hook:(Evd.evar_universe_context -> unit) Lemmas.declaration_hook -> ?opaque:bool ->
+ notations ->
+ fixpoint_kind -> unit
+
+val obligation : int * Names.Id.t option * Constrexpr.constr_expr option ->
+ Genarg.glob_generic_argument option -> unit
+
+val next_obligation : Names.Id.t option -> Genarg.glob_generic_argument option -> unit
+
+val solve_obligations : Names.Id.t option -> unit Proofview.tactic option -> progress
+(* Number of remaining obligations to be solved for this program *)
+
+val solve_all_obligations : unit Proofview.tactic option -> unit
+
+val try_solve_obligation : int -> Names.Id.t option -> unit Proofview.tactic option -> unit
+
+val try_solve_obligations : Names.Id.t option -> unit Proofview.tactic option -> unit
+
+val show_obligations : ?msg:bool -> Names.Id.t option -> unit
+
+val show_term : Names.Id.t option -> std_ppcmds
+
+val admit_obligations : Names.Id.t option -> unit
+
+exception NoObligations of Names.Id.t option
+
+val explain_no_obligations : Names.Id.t option -> Pp.std_ppcmds
+
+val set_program_mode : bool -> unit
diff --git a/vernac/record.ml b/vernac/record.ml
new file mode 100644
index 000000000..d5faafaf8
--- /dev/null
+++ b/vernac/record.ml
@@ -0,0 +1,583 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Pp
+open CErrors
+open Util
+open Names
+open Globnames
+open Nameops
+open Term
+open Vars
+open Environ
+open Declarations
+open Entries
+open Declare
+open Constrintern
+open Decl_kinds
+open Type_errors
+open Constrexpr
+open Constrexpr_ops
+open Goptions
+open Sigma.Notations
+open Context.Rel.Declaration
+
+module RelDecl = Context.Rel.Declaration
+
+(********** definition d'un record (structure) **************)
+
+(** Flag governing use of primitive projections. Disabled by default. *)
+let primitive_flag = ref false
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "use of primitive projections";
+ optkey = ["Primitive";"Projections"];
+ optread = (fun () -> !primitive_flag) ;
+ optwrite = (fun b -> primitive_flag := b) }
+
+let typeclasses_strict = ref false
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "strict typeclass resolution";
+ optkey = ["Typeclasses";"Strict";"Resolution"];
+ optread = (fun () -> !typeclasses_strict);
+ optwrite = (fun b -> typeclasses_strict := b); }
+
+let typeclasses_unique = ref false
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "unique typeclass instances";
+ optkey = ["Typeclasses";"Unique";"Instances"];
+ optread = (fun () -> !typeclasses_unique);
+ optwrite = (fun b -> typeclasses_unique := b); }
+
+let interp_fields_evars env evars impls_env nots l =
+ List.fold_left2
+ (fun (env, uimpls, params, impls) no ((loc, i), b, t) ->
+ let t', impl = interp_type_evars_impls env evars ~impls t in
+ let b' = Option.map (fun x -> fst (interp_casted_constr_evars_impls env evars ~impls x t')) b in
+ let impls =
+ match i with
+ | Anonymous -> impls
+ | Name id -> Id.Map.add id (compute_internalization_data env Constrintern.Method t' impl) impls
+ in
+ let d = match b' with
+ | None -> LocalAssum (i,t')
+ | Some b' -> LocalDef (i,b',t')
+ in
+ List.iter (Metasyntax.set_notation_for_interpretation impls) no;
+ (push_rel d env, impl :: uimpls, d::params, impls))
+ (env, [], [], impls_env) nots l
+
+let compute_constructor_level evars env l =
+ List.fold_right (fun d (env, univ) ->
+ let univ =
+ if is_local_assum d then
+ let s = Retyping.get_sort_of env evars (RelDecl.get_type d) in
+ Univ.sup (univ_of_sort s) univ
+ else univ
+ in (push_rel d env, univ))
+ l (env, Univ.type0m_univ)
+
+let binder_of_decl = function
+ | Vernacexpr.AssumExpr(n,t) -> (n,None,t)
+ | Vernacexpr.DefExpr(n,c,t) -> (n,Some c, match t with Some c -> c | None -> CHole (fst n, None, Misctypes.IntroAnonymous, None))
+
+let binders_of_decls = List.map binder_of_decl
+
+let typecheck_params_and_fields def id pl t ps nots fs =
+ let env0 = Global.env () in
+ let ctx = Evd.make_evar_universe_context env0 pl in
+ let evars = ref (Evd.from_ctx ctx) in
+ let _ =
+ let error bk (loc, name) =
+ match bk, name with
+ | Default _, Anonymous ->
+ user_err ~loc ~hdr:"record" (str "Record parameters must be named")
+ | _ -> ()
+ in
+ List.iter
+ (function LocalRawDef (b, _) -> error default_binder_kind b
+ | LocalRawAssum (ls, bk, ce) -> List.iter (error bk) ls
+ | LocalPattern _ -> assert false) ps
+ in
+ let impls_env, ((env1,newps), imps) = interp_context_evars env0 evars ps in
+ let t', template = match t with
+ | Some t ->
+ let env = push_rel_context newps env0 in
+ let poly =
+ match t with
+ | CSort (_, Misctypes.GType []) -> true | _ -> false in
+ let s = interp_type_evars env evars ~impls:empty_internalization_env t in
+ let sred = Reductionops.whd_all env !evars s in
+ (match kind_of_term sred with
+ | Sort s' ->
+ (if poly then
+ match Evd.is_sort_variable !evars s' with
+ | Some l -> evars := Evd.make_flexible_variable !evars true l;
+ sred, true
+ | None -> s, false
+ else s, false)
+ | _ -> user_err ~loc:(constr_loc t) (str"Sort expected."))
+ | None ->
+ let uvarkind = Evd.univ_flexible_alg in
+ mkSort (Evarutil.evd_comb0 (Evd.new_sort_variable uvarkind) evars), true
+ in
+ let fullarity = it_mkProd_or_LetIn t' newps in
+ let env_ar = push_rel_context newps (push_rel (LocalAssum (Name id,fullarity)) env0) in
+ let env2,impls,newfs,data =
+ interp_fields_evars env_ar evars impls_env nots (binders_of_decls fs)
+ in
+ let sigma =
+ Pretyping.solve_remaining_evars Pretyping.all_and_fail_flags env_ar !evars (Evd.empty,!evars) in
+ let evars, nf = Evarutil.nf_evars_and_universes sigma in
+ let arity = nf t' in
+ let arity, evars =
+ let _, univ = compute_constructor_level evars env_ar newfs in
+ let ctx, aritysort = Reduction.dest_arity env0 arity in
+ assert(List.is_empty ctx); (* Ensured by above analysis *)
+ if not def && (Sorts.is_prop aritysort ||
+ (Sorts.is_set aritysort && is_impredicative_set env0)) then
+ arity, evars
+ else
+ let evars = Evd.set_leq_sort env_ar evars (Type univ) aritysort in
+ if Univ.is_small_univ univ &&
+ Option.cata (Evd.is_flexible_level evars) false (Evd.is_sort_variable evars aritysort) then
+ (* We can assume that the level in aritysort is not constrained
+ and clear it, if it is flexible *)
+ mkArity (ctx, Sorts.sort_of_univ univ),
+ Evd.set_eq_sort env_ar evars (Prop Pos) aritysort
+ else arity, evars
+ in
+ let evars, nf = Evarutil.nf_evars_and_universes evars in
+ let newps = Context.Rel.map nf newps in
+ let newfs = Context.Rel.map nf newfs in
+ let ce t = Pretyping.check_evars env0 Evd.empty evars t in
+ List.iter (iter_constr ce) (List.rev newps);
+ List.iter (iter_constr ce) (List.rev newfs);
+ Evd.universe_context ?names:pl evars, nf arity, template, imps, newps, impls, newfs
+
+let degenerate_decl decl =
+ let id = match RelDecl.get_name decl with
+ | Name id -> id
+ | Anonymous -> anomaly (Pp.str "Unnamed record variable") in
+ match decl with
+ | LocalAssum (_,t) -> (id, LocalAssumEntry t)
+ | LocalDef (_,b,_) -> (id, LocalDefEntry b)
+
+type record_error =
+ | MissingProj of Id.t * Id.t list
+ | BadTypedProj of Id.t * env * Type_errors.type_error
+
+let warn_cannot_define_projection =
+ CWarnings.create ~name:"cannot-define-projection" ~category:"records"
+ (fun msg -> hov 0 msg)
+
+(* If a projection is not definable, we throw an error if the user
+asked it to be a coercion. Otherwise, we just print an info
+message. The user might still want to name the field of the record. *)
+let warning_or_error coe indsp err =
+ let st = match err with
+ | MissingProj (fi,projs) ->
+ let s,have = if List.length projs > 1 then "s","were" else "","was" in
+ (pr_id fi ++
+ strbrk" cannot be defined because the projection" ++ str s ++ spc () ++
+ prlist_with_sep pr_comma pr_id projs ++ spc () ++ str have ++
+ strbrk " not defined.")
+ | BadTypedProj (fi,ctx,te) ->
+ match te with
+ | ElimArity (_,_,_,_,Some (_,_,NonInformativeToInformative)) ->
+ (pr_id fi ++
+ strbrk" cannot be defined because it is informative and " ++
+ Printer.pr_inductive (Global.env()) indsp ++
+ strbrk " is not.")
+ | ElimArity (_,_,_,_,Some (_,_,StrongEliminationOnNonSmallType)) ->
+ (pr_id fi ++
+ strbrk" cannot be defined because it is large and " ++
+ Printer.pr_inductive (Global.env()) indsp ++
+ strbrk " is not.")
+ | _ ->
+ (pr_id fi ++ strbrk " cannot be defined because it is not typable.")
+ in
+ if coe then user_err ~hdr:"structure" st;
+ Flags.if_verbose Feedback.msg_info (hov 0 st)
+
+type field_status =
+ | NoProjection of Name.t
+ | Projection of constr
+
+exception NotDefinable of record_error
+
+(* This replaces previous projection bodies in current projection *)
+(* Undefined projs are collected and, at least one undefined proj occurs *)
+(* in the body of current projection then the latter can not be defined *)
+(* [c] is defined in ctxt [[params;fields]] and [l] is an instance of *)
+(* [[fields]] defined in ctxt [[params;x:ind]] *)
+let subst_projection fid l c =
+ let lv = List.length l in
+ let bad_projs = ref [] in
+ let rec substrec depth c = match kind_of_term c with
+ | Rel k ->
+ (* We are in context [[params;fields;x:ind;...depth...]] *)
+ if k <= depth+1 then
+ c
+ else if k-depth-1 <= lv then
+ match List.nth l (k-depth-2) with
+ | Projection t -> lift depth t
+ | NoProjection (Name id) -> bad_projs := id :: !bad_projs; mkRel k
+ | NoProjection Anonymous ->
+ user_err (str "Field " ++ pr_id fid ++
+ str " depends on the " ++ pr_nth (k-depth-1) ++ str
+ " field which has no name.")
+ else
+ mkRel (k-lv)
+ | _ -> map_constr_with_binders succ substrec depth c
+ in
+ let c' = lift 1 c in (* to get [c] defined in ctxt [[params;fields;x:ind]] *)
+ let c'' = substrec 0 c' in
+ if not (List.is_empty !bad_projs) then
+ raise (NotDefinable (MissingProj (fid,List.rev !bad_projs)));
+ c''
+
+let instantiate_possibly_recursive_type indu paramdecls fields =
+ let subst = List.map_i (fun i _ -> mkRel i) 1 paramdecls in
+ Termops.substl_rel_context (subst@[mkIndU indu]) fields
+
+let warn_non_primitive_record =
+ CWarnings.create ~name:"non-primitive-record" ~category:"record"
+ (fun (env,indsp) ->
+ (hov 0 (str "The record " ++ Printer.pr_inductive env indsp ++
+ strbrk" could not be defined as a primitive record")))
+
+(* We build projections *)
+let declare_projections indsp ?(kind=StructureComponent) binder_name coers fieldimpls fields =
+ let env = Global.env() in
+ let (mib,mip) = Global.lookup_inductive indsp in
+ let u = Declareops.inductive_instance mib in
+ let paramdecls = Inductive.inductive_paramdecls (mib, u) in
+ let poly = mib.mind_polymorphic in
+ let ctx = Univ.instantiate_univ_context mib.mind_universes in
+ let indu = indsp, u in
+ let r = mkIndU (indsp,u) in
+ let rp = applist (r, Context.Rel.to_extended_list 0 paramdecls) in
+ let paramargs = Context.Rel.to_extended_list 1 paramdecls in (*def in [[params;x:rp]]*)
+ let x = Name binder_name in
+ let fields = instantiate_possibly_recursive_type indu paramdecls fields in
+ let lifted_fields = Termops.lift_rel_context 1 fields in
+ let primitive =
+ if !primitive_flag then
+ let is_primitive =
+ match mib.mind_record with
+ | Some (Some _) -> true
+ | Some None | None -> false
+ in
+ if not is_primitive then
+ warn_non_primitive_record (env,indsp);
+ is_primitive
+ else false
+ in
+ let (_,_,kinds,sp_projs,_) =
+ List.fold_left3
+ (fun (nfi,i,kinds,sp_projs,subst) coe decl impls ->
+ let fi = RelDecl.get_name decl in
+ let ti = RelDecl.get_type decl in
+ let (sp_projs,i,subst) =
+ match fi with
+ | Anonymous ->
+ (None::sp_projs,i,NoProjection fi::subst)
+ | Name fid -> try
+ let kn, term =
+ if is_local_assum decl && primitive then
+ (** Already defined in the kernel silently *)
+ let kn = destConstRef (Nametab.locate (Libnames.qualid_of_ident fid)) in
+ Declare.definition_message fid;
+ kn, mkProj (Projection.make kn false,mkRel 1)
+ else
+ let ccl = subst_projection fid subst ti in
+ let body = match decl with
+ | LocalDef (_,ci,_) -> subst_projection fid subst ci
+ | LocalAssum _ ->
+ (* [ccl] is defined in context [params;x:rp] *)
+ (* [ccl'] is defined in context [params;x:rp;x:rp] *)
+ let ccl' = liftn 1 2 ccl in
+ let p = mkLambda (x, lift 1 rp, ccl') in
+ let branch = it_mkLambda_or_LetIn (mkRel nfi) lifted_fields in
+ let ci = Inductiveops.make_case_info env indsp LetStyle in
+ mkCase (ci, p, mkRel 1, [|branch|])
+ in
+ let proj =
+ it_mkLambda_or_LetIn (mkLambda (x,rp,body)) paramdecls in
+ let projtyp =
+ it_mkProd_or_LetIn (mkProd (x,rp,ccl)) paramdecls in
+ try
+ let entry = {
+ const_entry_body =
+ Future.from_val (Safe_typing.mk_pure_proof proj);
+ const_entry_secctx = None;
+ const_entry_type = Some projtyp;
+ const_entry_polymorphic = poly;
+ const_entry_universes =
+ if poly then ctx else Univ.UContext.empty;
+ const_entry_opaque = false;
+ const_entry_inline_code = false;
+ const_entry_feedback = None } in
+ let k = (DefinitionEntry entry,IsDefinition kind) in
+ let kn = declare_constant ~internal:InternalTacticRequest fid k in
+ let constr_fip =
+ let proj_args = (*Rel 1 refers to "x"*) paramargs@[mkRel 1] in
+ applist (mkConstU (kn,u),proj_args)
+ in
+ Declare.definition_message fid;
+ kn, constr_fip
+ with Type_errors.TypeError (ctx,te) ->
+ raise (NotDefinable (BadTypedProj (fid,ctx,te)))
+ in
+ let refi = ConstRef kn in
+ Impargs.maybe_declare_manual_implicits false refi impls;
+ if coe then begin
+ let cl = Class.class_of_global (IndRef indsp) in
+ Class.try_add_new_coercion_with_source refi ~local:false poly ~source:cl
+ end;
+ let i = if is_local_assum decl then i+1 else i in
+ (Some kn::sp_projs, i, Projection term::subst)
+ with NotDefinable why ->
+ warning_or_error coe indsp why;
+ (None::sp_projs,i,NoProjection fi::subst) in
+ (nfi-1,i,(fi, is_local_assum decl)::kinds,sp_projs,subst))
+ (List.length fields,0,[],[],[]) coers (List.rev fields) (List.rev fieldimpls)
+ in (kinds,sp_projs)
+
+let structure_signature ctx =
+ let rec deps_to_evar evm l =
+ match l with [] -> Evd.empty
+ | [decl] ->
+ let env = Environ.empty_named_context_val in
+ let evm = Sigma.Unsafe.of_evar_map evm in
+ let Sigma (_, evm, _) = Evarutil.new_pure_evar env evm (RelDecl.get_type decl) in
+ let evm = Sigma.to_evar_map evm in
+ evm
+ | decl::tl ->
+ let env = Environ.empty_named_context_val in
+ let evm = Sigma.Unsafe.of_evar_map evm in
+ let Sigma (ev, evm, _) = Evarutil.new_pure_evar env evm (RelDecl.get_type decl) in
+ let evm = Sigma.to_evar_map evm in
+ let new_tl = Util.List.map_i
+ (fun pos decl ->
+ RelDecl.map_type (fun t -> Termops.replace_term (mkRel pos) (mkEvar(ev,[||])) t) decl) 1 tl in
+ deps_to_evar evm new_tl in
+ deps_to_evar Evd.empty (List.rev ctx)
+
+open Typeclasses
+
+let declare_structure finite poly ctx id idbuild paramimpls params arity template
+ fieldimpls fields ?(kind=StructureComponent) ?name is_coe coers sign =
+ let nparams = List.length params and nfields = List.length fields in
+ let args = Context.Rel.to_extended_list nfields params in
+ let ind = applist (mkRel (1+nparams+nfields), args) in
+ let type_constructor = it_mkProd_or_LetIn ind fields in
+ let binder_name =
+ match name with
+ | None -> Id.of_string (Unicode.lowercase_first_char (Id.to_string id))
+ | Some n -> n
+ in
+ let mie_ind =
+ { mind_entry_typename = id;
+ mind_entry_arity = arity;
+ mind_entry_template = not poly && template;
+ mind_entry_consnames = [idbuild];
+ mind_entry_lc = [type_constructor] }
+ in
+ let mie =
+ { mind_entry_params = List.map degenerate_decl params;
+ mind_entry_record = Some (if !primitive_flag then Some binder_name else None);
+ mind_entry_finite = finite;
+ mind_entry_inds = [mie_ind];
+ mind_entry_polymorphic = poly;
+ mind_entry_private = None;
+ mind_entry_universes = ctx;
+ }
+ in
+ let kn = Command.declare_mutual_inductive_with_eliminations mie [] [(paramimpls,[])] in
+ let rsp = (kn,0) in (* This is ind path of idstruc *)
+ let cstr = (rsp,1) in
+ let kinds,sp_projs = declare_projections rsp ~kind binder_name coers fieldimpls fields in
+ let build = ConstructRef cstr in
+ let () = if is_coe then Class.try_add_new_coercion build ~local:false poly in
+ Recordops.declare_structure(rsp,cstr,List.rev kinds,List.rev sp_projs);
+ rsp
+
+let implicits_of_context ctx =
+ List.map_i (fun i name ->
+ let explname =
+ match name with
+ | Name n -> Some n
+ | Anonymous -> None
+ in ExplByPos (i, explname), (true, true, true))
+ 1 (List.rev (Anonymous :: (List.map RelDecl.get_name ctx)))
+
+let declare_class finite def poly ctx id idbuild paramimpls params arity
+ template fieldimpls fields ?(kind=StructureComponent) is_coe coers priorities sign =
+ let fieldimpls =
+ (* Make the class implicit in the projections, and the params if applicable. *)
+ let len = List.length params in
+ let impls = implicits_of_context params in
+ List.map (fun x -> impls @ Impargs.lift_implicits (succ len) x) fieldimpls
+ in
+ let binder_name = Namegen.next_ident_away (snd id) (Termops.ids_of_context (Global.env())) in
+ let impl, projs =
+ match fields with
+ | [LocalAssum (Name proj_name, field) | LocalDef (Name proj_name, _, field)] when def ->
+ let class_body = it_mkLambda_or_LetIn field params in
+ let class_type = it_mkProd_or_LetIn arity params in
+ let class_entry =
+ Declare.definition_entry ~types:class_type ~poly ~univs:ctx class_body in
+ let cst = Declare.declare_constant (snd id)
+ (DefinitionEntry class_entry, IsDefinition Definition)
+ in
+ let cstu = (cst, if poly then Univ.UContext.instance ctx else Univ.Instance.empty) in
+ let inst_type = appvectc (mkConstU cstu)
+ (Termops.rel_vect 0 (List.length params)) in
+ let proj_type =
+ it_mkProd_or_LetIn (mkProd(Name binder_name, inst_type, lift 1 field)) params in
+ let proj_body =
+ it_mkLambda_or_LetIn (mkLambda (Name binder_name, inst_type, mkRel 1)) params in
+ let proj_entry =
+ Declare.definition_entry ~types:proj_type ~poly
+ ~univs:(if poly then ctx else Univ.UContext.empty) proj_body
+ in
+ let proj_cst = Declare.declare_constant proj_name
+ (DefinitionEntry proj_entry, IsDefinition Definition)
+ in
+ let cref = ConstRef cst in
+ Impargs.declare_manual_implicits false cref [paramimpls];
+ Impargs.declare_manual_implicits false (ConstRef proj_cst) [List.hd fieldimpls];
+ Classes.set_typeclass_transparency (EvalConstRef cst) false false;
+ let sub = match List.hd coers with
+ | Some b -> Some ((if b then Backward else Forward), List.hd priorities)
+ | None -> None
+ in
+ cref, [Name proj_name, sub, Some proj_cst]
+ | _ ->
+ let ind = declare_structure BiFinite poly ctx (snd id) idbuild paramimpls
+ params arity template fieldimpls fields
+ ~kind:Method ~name:binder_name false (List.map (fun _ -> false) fields) sign
+ in
+ let coers = List.map2 (fun coe pri ->
+ Option.map (fun b ->
+ if b then Backward, pri else Forward, pri) coe)
+ coers priorities
+ in
+ let l = List.map3 (fun decl b y -> RelDecl.get_name decl, b, y)
+ (List.rev fields) coers (Recordops.lookup_projections ind)
+ in IndRef ind, l
+ in
+ let ctx_context =
+ List.map (fun decl ->
+ match Typeclasses.class_of_constr (RelDecl.get_type decl) with
+ | Some (_, ((cl,_), _)) -> Some (cl.cl_impl, true)
+ | None -> None)
+ params, params
+ in
+ let k =
+ { cl_impl = impl;
+ cl_strict = !typeclasses_strict;
+ cl_unique = !typeclasses_unique;
+ cl_context = ctx_context;
+ cl_props = fields;
+ cl_projs = projs }
+ in
+ add_class k; impl
+
+
+let add_constant_class cst =
+ let ty = Universes.unsafe_type_of_global (ConstRef cst) in
+ let ctx, arity = decompose_prod_assum ty in
+ let tc =
+ { cl_impl = ConstRef cst;
+ cl_context = (List.map (const None) ctx, ctx);
+ cl_props = [LocalAssum (Anonymous, arity)];
+ cl_projs = [];
+ cl_strict = !typeclasses_strict;
+ cl_unique = !typeclasses_unique
+ }
+ in add_class tc;
+ set_typeclass_transparency (EvalConstRef cst) false false
+
+let add_inductive_class ind =
+ let mind, oneind = Global.lookup_inductive ind in
+ let k =
+ let ctx = oneind.mind_arity_ctxt in
+ let inst = Univ.UContext.instance mind.mind_universes in
+ let ty = Inductive.type_of_inductive
+ (push_rel_context ctx (Global.env ()))
+ ((mind,oneind),inst)
+ in
+ { cl_impl = IndRef ind;
+ cl_context = List.map (const None) ctx, ctx;
+ cl_props = [LocalAssum (Anonymous, ty)];
+ cl_projs = [];
+ cl_strict = !typeclasses_strict;
+ cl_unique = !typeclasses_unique }
+ in add_class k
+
+let declare_existing_class g =
+ match g with
+ | ConstRef x -> add_constant_class x
+ | IndRef x -> add_inductive_class x
+ | _ -> user_err ~hdr:"declare_existing_class"
+ (Pp.str"Unsupported class type, only constants and inductives are allowed")
+
+open Vernacexpr
+
+(* [fs] corresponds to fields and [ps] to parameters; [coers] is a
+ list telling if the corresponding fields must me declared as coercions
+ or subinstances. *)
+let definition_structure (kind,poly,finite,(is_coe,((loc,idstruc),pl)),ps,cfs,idbuild,s) =
+ let cfs,notations = List.split cfs in
+ let cfs,priorities = List.split cfs in
+ let coers,fs = List.split cfs in
+ let extract_name acc = function
+ Vernacexpr.AssumExpr((_,Name id),_) -> id::acc
+ | Vernacexpr.DefExpr ((_,Name id),_,_) -> id::acc
+ | _ -> acc in
+ let allnames = idstruc::(List.fold_left extract_name [] fs) in
+ let () = match List.duplicates Id.equal allnames with
+ | [] -> ()
+ | id :: _ -> user_err (str "Two objects have the same name" ++ spc () ++ quote (Id.print id))
+ in
+ let isnot_class = match kind with Class false -> false | _ -> true in
+ if isnot_class && List.exists (fun opt -> not (Option.is_empty opt)) priorities then
+ error "Priorities only allowed for type class substructures";
+ (* Now, younger decl in params and fields is on top *)
+ let (pl, ctx), arity, template, implpars, params, implfs, fields =
+ States.with_state_protection (fun () ->
+ typecheck_params_and_fields (kind = Class true) idstruc pl s ps notations fs) () in
+ let sign = structure_signature (fields@params) in
+ let gr = match kind with
+ | Class def ->
+ let priorities = List.map (fun id -> {hint_priority = id; hint_pattern = None}) priorities in
+ let gr = declare_class finite def poly ctx (loc,idstruc) idbuild
+ implpars params arity template implfs fields is_coe coers priorities sign in
+ gr
+ | _ ->
+ let implfs = List.map
+ (fun impls -> implpars @ Impargs.lift_implicits
+ (succ (List.length params)) impls) implfs in
+ let ind = declare_structure finite poly ctx idstruc
+ idbuild implpars params arity template implfs
+ fields is_coe (List.map (fun coe -> not (Option.is_empty coe)) coers) sign in
+ IndRef ind
+ in
+ Universes.register_universe_binders gr pl;
+ gr
diff --git a/vernac/record.mli b/vernac/record.mli
new file mode 100644
index 000000000..c50e57786
--- /dev/null
+++ b/vernac/record.mli
@@ -0,0 +1,46 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open Term
+open Vernacexpr
+open Constrexpr
+open Impargs
+open Globnames
+
+val primitive_flag : bool ref
+
+(** [declare_projections ref name coers params fields] declare projections of
+ record [ref] (if allowed) using the given [name] as argument, and put them
+ as coercions accordingly to [coers]; it returns the absolute names of projections *)
+
+val declare_projections :
+ inductive -> ?kind:Decl_kinds.definition_object_kind -> Id.t ->
+ coercion_flag list -> manual_explicitation list list -> Context.Rel.t ->
+ (Name.t * bool) list * constant option list
+
+val declare_structure :
+ Decl_kinds.recursivity_kind ->
+ bool (** polymorphic?*) -> Univ.universe_context ->
+ Id.t -> Id.t ->
+ manual_explicitation list -> Context.Rel.t -> (** params *) constr -> (** arity *)
+ bool (** template arity ? *) ->
+ Impargs.manual_explicitation list list -> Context.Rel.t -> (** fields *)
+ ?kind:Decl_kinds.definition_object_kind -> ?name:Id.t ->
+ bool -> (** coercion? *)
+ bool list -> (** field coercions *)
+ Evd.evar_map ->
+ inductive
+
+val definition_structure :
+ inductive_kind * Decl_kinds.polymorphic * Decl_kinds.recursivity_kind *
+ plident with_coercion * local_binder list *
+ (local_decl_expr with_instance with_priority with_notation) list *
+ Id.t * constr_expr option -> global_reference
+
+val declare_existing_class : global_reference -> unit
diff --git a/vernac/search.ml b/vernac/search.ml
new file mode 100644
index 000000000..e1b56b131
--- /dev/null
+++ b/vernac/search.ml
@@ -0,0 +1,381 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Pp
+open Util
+open Names
+open Term
+open Declarations
+open Libobject
+open Environ
+open Pattern
+open Printer
+open Libnames
+open Globnames
+open Nametab
+open Goptions
+
+module NamedDecl = Context.Named.Declaration
+
+type filter_function = global_reference -> env -> constr -> bool
+type display_function = global_reference -> env -> constr -> unit
+
+(* This option restricts the output of [SearchPattern ...],
+[SearchAbout ...], etc. to the names of the symbols matching the
+query, separated by a newline. This type of output is useful for
+editors (like emacs), to generate a list of completion candidates
+without having to parse thorugh the types of all symbols. *)
+
+type glob_search_about_item =
+ | GlobSearchSubPattern of constr_pattern
+ | GlobSearchString of string
+
+module SearchBlacklist =
+ Goptions.MakeStringTable
+ (struct
+ let key = ["Search";"Blacklist"]
+ let title = "Current search blacklist : "
+ let member_message s b =
+ str "Search blacklist does " ++ (if b then mt () else str "not ") ++ str "include " ++ str s
+ let synchronous = true
+ end)
+
+(* The functions iter_constructors and iter_declarations implement the behavior
+ needed for the Coq searching commands.
+ These functions take as first argument the procedure
+ that will be called to treat each entry. This procedure receives the name
+ of the object, the assumptions that will make it possible to print its type,
+ and the constr term that represent its type. *)
+
+let iter_constructors indsp u fn env nconstr =
+ for i = 1 to nconstr do
+ let typ = Inductiveops.type_of_constructor env ((indsp, i), u) in
+ fn (ConstructRef (indsp, i)) env typ
+ done
+
+let iter_named_context_name_type f =
+ List.iter (fun decl -> f (NamedDecl.get_id decl) (NamedDecl.get_type decl))
+
+(* General search over hypothesis of a goal *)
+let iter_hypothesis glnum (fn : global_reference -> env -> constr -> unit) =
+ let env = Global.env () in
+ let iter_hyp idh typ = fn (VarRef idh) env typ in
+ let evmap,e = Pfedit.get_goal_context glnum in
+ let pfctxt = named_context e in
+ iter_named_context_name_type iter_hyp pfctxt
+
+(* General search over declarations *)
+let iter_declarations (fn : global_reference -> env -> constr -> unit) =
+ let env = Global.env () in
+ let iter_obj (sp, kn) lobj = match object_tag lobj with
+ | "VARIABLE" ->
+ begin try
+ let decl = Global.lookup_named (basename sp) in
+ fn (VarRef (NamedDecl.get_id decl)) env (NamedDecl.get_type decl)
+ with Not_found -> (* we are in a section *) () end
+ | "CONSTANT" ->
+ let cst = Global.constant_of_delta_kn kn in
+ let gr = ConstRef cst in
+ let typ = Global.type_of_global_unsafe gr in
+ fn gr env typ
+ | "INDUCTIVE" ->
+ let mind = Global.mind_of_delta_kn kn in
+ let mib = Global.lookup_mind mind in
+ let iter_packet i mip =
+ let ind = (mind, i) in
+ let u = Declareops.inductive_instance mib in
+ let i = (ind, u) in
+ let typ = Inductiveops.type_of_inductive env i in
+ let () = fn (IndRef ind) env typ in
+ let len = Array.length mip.mind_user_lc in
+ iter_constructors ind u fn env len
+ in
+ Array.iteri iter_packet mib.mind_packets
+ | _ -> ()
+ in
+ try Declaremods.iter_all_segments iter_obj
+ with Not_found -> ()
+
+let generic_search glnumopt fn =
+ (match glnumopt with
+ | None -> ()
+ | Some glnum -> iter_hypothesis glnum fn);
+ iter_declarations fn
+
+(** This module defines a preference on constrs in the form of a
+ [compare] function (preferred constr must be big for this
+ functions, so preferences such as small constr must use a reversed
+ order). This priority will be used to order search results and
+ propose first results which are more likely to be relevant to the
+ query, this is why the type [t] contains the other elements
+ required of a search. *)
+module ConstrPriority = struct
+
+ (* The priority is memoised here. Because of the very localised use
+ of this module, it is not worth it making a convenient interface. *)
+ type t =
+ Globnames.global_reference * Environ.env * Constr.t * priority
+ and priority = int
+
+ module ConstrSet = CSet.Make(Constr)
+
+ (** A measure of the size of a term *)
+ let rec size t =
+ Constr.fold (fun s t -> 1 + s + size t) 0 t
+
+ (** Set of the "symbols" (definitions, inductives, constructors)
+ which appear in a term. *)
+ let rec symbols acc t =
+ let open Constr in
+ match kind t with
+ | Const _ | Ind _ | Construct _ -> ConstrSet.add t acc
+ | _ -> Constr.fold symbols acc t
+
+ (** The number of distinct "symbols" (see {!symbols}) which appear
+ in a term. *)
+ let num_symbols t =
+ ConstrSet.(cardinal (symbols empty t))
+
+ let priority t : priority =
+ -(3*(num_symbols t) + size t)
+
+ let compare (_,_,_,p1) (_,_,_,p2) =
+ compare p1 p2
+end
+
+module PriorityQueue = Heap.Functional(ConstrPriority)
+
+let rec iter_priority_queue q fn =
+ (* use an option to make the function tail recursive. Will be
+ obsoleted with Ocaml 4.02 with the [match … with | exception …]
+ syntax. *)
+ let next = begin
+ try Some (PriorityQueue.maximum q)
+ with Heap.EmptyHeap -> None
+ end in
+ match next with
+ | Some (gref,env,t,_) ->
+ fn gref env t;
+ iter_priority_queue (PriorityQueue.remove q) fn
+ | None -> ()
+
+let prioritize_search seq fn =
+ let acc = ref PriorityQueue.empty in
+ let iter gref env t =
+ let p = ConstrPriority.priority t in
+ acc := PriorityQueue.add (gref,env,t,p) !acc
+ in
+ let () = seq iter in
+ iter_priority_queue !acc fn
+
+(** Filters *)
+
+(** This function tries to see whether the conclusion matches a pattern. *)
+(** FIXME: this is quite dummy, we may find a more efficient algorithm. *)
+let rec pattern_filter pat ref env typ =
+ let typ = Termops.strip_outer_cast typ in
+ if Constr_matching.is_matching env Evd.empty pat typ then true
+ else match kind_of_term typ with
+ | Prod (_, _, typ)
+ | LetIn (_, _, _, typ) -> pattern_filter pat ref env typ
+ | _ -> false
+
+let rec head_filter pat ref env typ =
+ let typ = Termops.strip_outer_cast typ in
+ if Constr_matching.is_matching_head env Evd.empty pat typ then true
+ else match kind_of_term typ with
+ | Prod (_, _, typ)
+ | LetIn (_, _, _, typ) -> head_filter pat ref env typ
+ | _ -> false
+
+let full_name_of_reference ref =
+ let (dir,id) = repr_path (path_of_global ref) in
+ DirPath.to_string dir ^ "." ^ Id.to_string id
+
+(** Whether a reference is blacklisted *)
+let blacklist_filter_aux () =
+ let l = SearchBlacklist.elements () in
+ fun ref env typ ->
+ let name = full_name_of_reference ref in
+ let is_not_bl str = not (String.string_contains ~where:name ~what:str) in
+ List.for_all is_not_bl l
+
+let module_filter (mods, outside) ref env typ =
+ let sp = path_of_global ref in
+ let sl = dirpath sp in
+ let is_outside md = not (is_dirpath_prefix_of md sl) in
+ let is_inside md = is_dirpath_prefix_of md sl in
+ if outside then List.for_all is_outside mods
+ else List.is_empty mods || List.exists is_inside mods
+
+let name_of_reference ref = Id.to_string (basename_of_global ref)
+
+let search_about_filter query gr env typ = match query with
+| GlobSearchSubPattern pat ->
+ Constr_matching.is_matching_appsubterm ~closed:false env Evd.empty pat typ
+| GlobSearchString s ->
+ String.string_contains ~where:(name_of_reference gr) ~what:s
+
+
+(** SearchPattern *)
+
+let search_pattern gopt pat mods pr_search =
+ let blacklist_filter = blacklist_filter_aux () in
+ let filter ref env typ =
+ module_filter mods ref env typ &&
+ pattern_filter pat ref env typ &&
+ blacklist_filter ref env typ
+ in
+ let iter ref env typ =
+ if filter ref env typ then pr_search ref env typ
+ in
+ generic_search gopt iter
+
+(** SearchRewrite *)
+
+let eq = Coqlib.glob_eq
+
+let rewrite_pat1 pat =
+ PApp (PRef eq, [| PMeta None; pat; PMeta None |])
+
+let rewrite_pat2 pat =
+ PApp (PRef eq, [| PMeta None; PMeta None; pat |])
+
+let search_rewrite gopt pat mods pr_search =
+ let pat1 = rewrite_pat1 pat in
+ let pat2 = rewrite_pat2 pat in
+ let blacklist_filter = blacklist_filter_aux () in
+ let filter ref env typ =
+ module_filter mods ref env typ &&
+ (pattern_filter pat1 ref env typ ||
+ pattern_filter pat2 ref env typ) &&
+ blacklist_filter ref env typ
+ in
+ let iter ref env typ =
+ if filter ref env typ then pr_search ref env typ
+ in
+ generic_search gopt iter
+
+(** Search *)
+
+let search_by_head gopt pat mods pr_search =
+ let blacklist_filter = blacklist_filter_aux () in
+ let filter ref env typ =
+ module_filter mods ref env typ &&
+ head_filter pat ref env typ &&
+ blacklist_filter ref env typ
+ in
+ let iter ref env typ =
+ if filter ref env typ then pr_search ref env typ
+ in
+ generic_search gopt iter
+
+(** SearchAbout *)
+
+let search_about gopt items mods pr_search =
+ let blacklist_filter = blacklist_filter_aux () in
+ let filter ref env typ =
+ let eqb b1 b2 = if b1 then b2 else not b2 in
+ module_filter mods ref env typ &&
+ List.for_all
+ (fun (b,i) -> eqb b (search_about_filter i ref env typ)) items &&
+ blacklist_filter ref env typ
+ in
+ let iter ref env typ =
+ if filter ref env typ then pr_search ref env typ
+ in
+ generic_search gopt iter
+
+type search_constraint =
+ | Name_Pattern of Str.regexp
+ | Type_Pattern of Pattern.constr_pattern
+ | SubType_Pattern of Pattern.constr_pattern
+ | In_Module of Names.DirPath.t
+ | Include_Blacklist
+
+type 'a coq_object = {
+ coq_object_prefix : string list;
+ coq_object_qualid : string list;
+ coq_object_object : 'a;
+}
+
+let interface_search =
+ let rec extract_flags name tpe subtpe mods blacklist = function
+ | [] -> (name, tpe, subtpe, mods, blacklist)
+ | (Name_Pattern regexp, b) :: l ->
+ extract_flags ((regexp, b) :: name) tpe subtpe mods blacklist l
+ | (Type_Pattern pat, b) :: l ->
+ extract_flags name ((pat, b) :: tpe) subtpe mods blacklist l
+ | (SubType_Pattern pat, b) :: l ->
+ extract_flags name tpe ((pat, b) :: subtpe) mods blacklist l
+ | (In_Module id, b) :: l ->
+ extract_flags name tpe subtpe ((id, b) :: mods) blacklist l
+ | (Include_Blacklist, b) :: l ->
+ extract_flags name tpe subtpe mods b l
+ in
+ fun ?glnum flags ->
+ let (name, tpe, subtpe, mods, blacklist) =
+ extract_flags [] [] [] [] false flags
+ in
+ let blacklist_filter = blacklist_filter_aux () in
+ let filter_function ref env constr =
+ let id = Names.Id.to_string (Nametab.basename_of_global ref) in
+ let path = Libnames.dirpath (Nametab.path_of_global ref) in
+ let toggle x b = if x then b else not b in
+ let match_name (regexp, flag) =
+ toggle (Str.string_match regexp id 0) flag
+ in
+ let match_type (pat, flag) =
+ toggle (Constr_matching.is_matching env Evd.empty pat constr) flag
+ in
+ let match_subtype (pat, flag) =
+ toggle
+ (Constr_matching.is_matching_appsubterm ~closed:false
+ env Evd.empty pat constr) flag
+ in
+ let match_module (mdl, flag) =
+ toggle (Libnames.is_dirpath_prefix_of mdl path) flag
+ in
+ List.for_all match_name name &&
+ List.for_all match_type tpe &&
+ List.for_all match_subtype subtpe &&
+ List.for_all match_module mods &&
+ (blacklist || blacklist_filter ref env constr)
+ in
+ let ans = ref [] in
+ let print_function ref env constr =
+ let fullpath = DirPath.repr (Nametab.dirpath_of_global ref) in
+ let qualid = Nametab.shortest_qualid_of_global Id.Set.empty ref in
+ let (shortpath, basename) = Libnames.repr_qualid qualid in
+ let shortpath = DirPath.repr shortpath in
+ (* [shortpath] is a suffix of [fullpath] and we're looking for the missing
+ prefix *)
+ let rec prefix full short accu = match full, short with
+ | _, [] ->
+ let full = List.rev_map Id.to_string full in
+ (full, accu)
+ | _ :: full, m :: short ->
+ prefix full short (Id.to_string m :: accu)
+ | _ -> assert false
+ in
+ let (prefix, qualid) = prefix fullpath shortpath [Id.to_string basename] in
+ let answer = {
+ coq_object_prefix = prefix;
+ coq_object_qualid = qualid;
+ coq_object_object = string_of_ppcmds (pr_lconstr_env env Evd.empty constr);
+ } in
+ ans := answer :: !ans;
+ in
+ let iter ref env typ =
+ if filter_function ref env typ then print_function ref env typ
+ in
+ let () = generic_search glnum iter in
+ !ans
+
+let blacklist_filter ref env typ =
+ blacklist_filter_aux () ref env typ
diff --git a/vernac/search.mli b/vernac/search.mli
new file mode 100644
index 000000000..c9167c485
--- /dev/null
+++ b/vernac/search.mli
@@ -0,0 +1,84 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Pp
+open Names
+open Term
+open Environ
+open Pattern
+open Globnames
+
+(** {6 Search facilities. } *)
+
+type glob_search_about_item =
+ | GlobSearchSubPattern of constr_pattern
+ | GlobSearchString of string
+
+type filter_function = global_reference -> env -> constr -> bool
+type display_function = global_reference -> env -> constr -> unit
+
+(** {6 Generic filter functions} *)
+
+val blacklist_filter : filter_function
+(** Check whether a reference is blacklisted. *)
+
+val module_filter : DirPath.t list * bool -> filter_function
+(** Check whether a reference pertains or not to a set of modules *)
+
+val search_about_filter : glob_search_about_item -> filter_function
+(** Check whether a reference matches a SearchAbout query. *)
+
+(** {6 Specialized search functions}
+
+[search_xxx gl pattern modinout] searches the hypothesis of the [gl]th
+goal and the global environment for things matching [pattern] and
+satisfying module exclude/include clauses of [modinout]. *)
+
+val search_by_head : int option -> constr_pattern -> DirPath.t list * bool
+ -> display_function -> unit
+val search_rewrite : int option -> constr_pattern -> DirPath.t list * bool
+ -> display_function -> unit
+val search_pattern : int option -> constr_pattern -> DirPath.t list * bool
+ -> display_function -> unit
+val search_about : int option -> (bool * glob_search_about_item) list
+ -> DirPath.t list * bool -> display_function -> unit
+
+type search_constraint =
+ (** Whether the name satisfies a regexp (uses Ocaml Str syntax) *)
+ | Name_Pattern of Str.regexp
+ (** Whether the object type satisfies a pattern *)
+ | Type_Pattern of Pattern.constr_pattern
+ (** Whether some subtype of object type satisfies a pattern *)
+ | SubType_Pattern of Pattern.constr_pattern
+ (** Whether the object pertains to a module *)
+ | In_Module of Names.DirPath.t
+ (** Bypass the Search blacklist *)
+ | Include_Blacklist
+
+type 'a coq_object = {
+ coq_object_prefix : string list;
+ coq_object_qualid : string list;
+ coq_object_object : 'a;
+}
+
+val interface_search : ?glnum:int -> (search_constraint * bool) list ->
+ string coq_object list
+
+(** {6 Generic search function} *)
+
+val generic_search : int option -> display_function -> unit
+(** This function iterates over all hypothesis of the goal numbered
+ [glnum] (if present) and all known declarations. *)
+
+(** {6 Search function modifiers} *)
+
+val prioritize_search : (display_function -> unit) -> display_function -> unit
+(** [prioritize_search iter] iterates over the values of [iter] (seen
+ as a sequence of declarations), in a relevance order. This requires to
+ perform the entire iteration of [iter] before starting streaming. So
+ [prioritize_search] should not be used for low-latency streaming. *)
diff --git a/vernac/vernac.mllib b/vernac/vernac.mllib
new file mode 100644
index 000000000..94ef54f70
--- /dev/null
+++ b/vernac/vernac.mllib
@@ -0,0 +1,17 @@
+Lemmas
+Himsg
+ExplainErr
+Class
+Locality
+Metasyntax
+Auto_ind_decl
+Search
+Indschemes
+Obligations
+Command
+Classes
+Record
+Assumptions
+Vernacinterp
+Mltop
+Vernacentries
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
new file mode 100644
index 000000000..15f89e4b8
--- /dev/null
+++ b/vernac/vernacentries.ml
@@ -0,0 +1,2271 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* Concrete syntax of the mathematical vernacular MV V2.6 *)
+
+open Pp
+open CErrors
+open Util
+open Flags
+open Names
+open Nameops
+open Term
+open Pfedit
+open Tacmach
+open Constrintern
+open Prettyp
+open Printer
+open Command
+open Goptions
+open Libnames
+open Globnames
+open Vernacexpr
+open Decl_kinds
+open Constrexpr
+open Redexpr
+open Lemmas
+open Misctypes
+open Locality
+open Sigma.Notations
+
+module NamedDecl = Context.Named.Declaration
+
+(** TODO: make this function independent of Ltac *)
+let (f_interp_redexp, interp_redexp_hook) = Hook.make ()
+
+let debug = false
+let prerr_endline x =
+ if debug then prerr_endline (x ()) else ()
+
+(* Misc *)
+
+let cl_of_qualid = function
+ | FunClass -> Classops.CL_FUN
+ | SortClass -> Classops.CL_SORT
+ | RefClass r -> Class.class_of_global (Smartlocate.smart_global ~head:true r)
+
+let scope_class_of_qualid qid =
+ Notation.scope_class_of_class (cl_of_qualid qid)
+
+(*******************)
+(* "Show" commands *)
+
+let show_proof () =
+ (* spiwack: this would probably be cooler with a bit of polishing. *)
+ let p = Proof_global.give_me_the_proof () in
+ let pprf = Proof.partial_proof p in
+ Feedback.msg_notice (Pp.prlist_with_sep Pp.fnl Printer.pr_constr pprf)
+
+let show_node () =
+ (* spiwack: I'm have little clue what this function used to do. I deactivated it,
+ could, possibly, be cleaned away. (Feb. 2010) *)
+ ()
+
+let show_thesis () =
+ Feedback.msg_error (anomaly (Pp.str "TODO") )
+
+let show_top_evars () =
+ (* spiwack: new as of Feb. 2010: shows goal evars in addition to non-goal evars. *)
+ let pfts = get_pftreestate () in
+ let gls = Proof.V82.subgoals pfts in
+ let sigma = gls.Evd.sigma in
+ Feedback.msg_notice (pr_evars_int sigma 1 (Evarutil.non_instantiated sigma))
+
+let show_universes () =
+ let pfts = get_pftreestate () in
+ let gls = Proof.V82.subgoals pfts in
+ let sigma = gls.Evd.sigma in
+ let ctx = Evd.universe_context_set (Evd.nf_constraints sigma) in
+ Feedback.msg_notice (Evd.pr_evar_universe_context (Evd.evar_universe_context sigma));
+ Feedback.msg_notice (str"Normalized constraints: " ++ Univ.pr_universe_context_set (Evd.pr_evd_level sigma) ctx)
+
+let show_prooftree () =
+ (* Spiwack: proof tree is currently not working *)
+ ()
+
+let enable_goal_printing = ref true
+
+let print_subgoals () =
+ if !enable_goal_printing && is_verbose ()
+ then begin
+ Feedback.msg_notice (pr_open_subgoals ())
+ end
+
+let try_print_subgoals () =
+ try print_subgoals () with Proof_global.NoCurrentProof | UserError _ -> ()
+
+
+ (* Simulate the Intro(s) tactic *)
+
+let show_intro all =
+ let pf = get_pftreestate() in
+ let {Evd.it=gls ; sigma=sigma; } = Proof.V82.subgoals pf in
+ if not (List.is_empty gls) then begin
+ let gl = {Evd.it=List.hd gls ; sigma = sigma; } in
+ let l,_= decompose_prod_assum (Termops.strip_outer_cast (pf_concl gl)) in
+ if all then
+ let lid = Tactics.find_intro_names l gl in
+ Feedback.msg_notice (hov 0 (prlist_with_sep spc pr_id lid))
+ else if not (List.is_empty l) then
+ let n = List.last l in
+ Feedback.msg_notice (pr_id (List.hd (Tactics.find_intro_names [n] gl)))
+ end
+
+(** Prepare a "match" template for a given inductive type.
+ For each branch of the match, we list the constructor name
+ followed by enough pattern variables.
+ [Not_found] is raised if the given string isn't the qualid of
+ a known inductive type. *)
+
+let make_cases_aux glob_ref =
+ match glob_ref with
+ | Globnames.IndRef i ->
+ let {Declarations.mind_nparams = np}
+ , {Declarations.mind_consnames = carr ; Declarations.mind_nf_lc = tarr }
+ = Global.lookup_inductive i in
+ Util.Array.fold_right2
+ (fun consname typ l ->
+ let al = List.rev (fst (decompose_prod typ)) in
+ let al = Util.List.skipn np al in
+ let rec rename avoid = function
+ | [] -> []
+ | (n,_)::l ->
+ let n' = Namegen.next_name_away_with_default (Id.to_string Namegen.default_dependent_ident) n avoid in
+ Id.to_string n' :: rename (n'::avoid) l in
+ let al' = rename [] al in
+ (Id.to_string consname :: al') :: l)
+ carr tarr []
+ | _ -> raise Not_found
+
+let make_cases s =
+ let qualified_name = Libnames.qualid_of_string s in
+ let glob_ref = Nametab.locate qualified_name in
+ make_cases_aux glob_ref
+
+(** Textual display of a generic "match" template *)
+
+let show_match id =
+ let patterns =
+ try make_cases_aux (Nametab.global id)
+ with Not_found -> error "Unknown inductive type."
+ in
+ let pr_branch l =
+ str "| " ++ hov 1 (prlist_with_sep spc str l) ++ str " =>"
+ in
+ Feedback.msg_notice (v 1 (str "match # with" ++ fnl () ++
+ prlist_with_sep fnl pr_branch patterns ++ fnl () ++ str "end" ++ fnl ()))
+
+(* "Print" commands *)
+
+let print_path_entry p =
+ let dir = pr_dirpath (Loadpath.logical p) in
+ let path = str (Loadpath.physical p) in
+ (dir ++ str " " ++ tbrk (0, 0) ++ path)
+
+let print_loadpath dir =
+ let l = Loadpath.get_load_paths () in
+ let l = match dir with
+ | None -> l
+ | Some dir ->
+ let filter p = is_dirpath_prefix_of dir (Loadpath.logical p) in
+ List.filter filter l
+ in
+ Pp.t (str "Logical Path: " ++
+ tab () ++ str "Physical path:" ++ fnl () ++
+ prlist_with_sep fnl print_path_entry l)
+
+let print_modules () =
+ let opened = Library.opened_libraries ()
+ and loaded = Library.loaded_libraries () in
+ (* we intersect over opened to preserve the order of opened since *)
+ (* non-commutative operations (e.g. visibility) are done at import time *)
+ let loaded_opened = List.intersect DirPath.equal opened loaded
+ and only_loaded = List.subtract DirPath.equal loaded opened in
+ str"Loaded and imported library files: " ++
+ pr_vertical_list pr_dirpath loaded_opened ++ fnl () ++
+ str"Loaded and not imported library files: " ++
+ pr_vertical_list pr_dirpath only_loaded
+
+
+let print_module r =
+ let (loc,qid) = qualid_of_reference r in
+ try
+ let globdir = Nametab.locate_dir qid in
+ match globdir with
+ DirModule (dirpath,(mp,_)) ->
+ Feedback.msg_notice (Printmod.print_module (Printmod.printable_body dirpath) mp)
+ | _ -> raise Not_found
+ with
+ Not_found -> Feedback.msg_error (str"Unknown Module " ++ pr_qualid qid)
+
+let print_modtype r =
+ let (loc,qid) = qualid_of_reference r in
+ try
+ let kn = Nametab.locate_modtype qid in
+ Feedback.msg_notice (Printmod.print_modtype kn)
+ with Not_found ->
+ (* Is there a module of this name ? If yes we display its type *)
+ try
+ let mp = Nametab.locate_module qid in
+ Feedback.msg_notice (Printmod.print_module false mp)
+ with Not_found ->
+ Feedback.msg_error (str"Unknown Module Type or Module " ++ pr_qualid qid)
+
+let print_namespace ns =
+ let ns = List.rev (Names.DirPath.repr ns) in
+ (* [match_dirpath], [match_modulpath] are helpers for [matches]
+ which checks whether a constant is in the namespace [ns]. *)
+ let rec match_dirpath ns = function
+ | [] -> Some ns
+ | id::dir ->
+ begin match match_dirpath ns dir with
+ | Some [] as y -> y
+ | Some (a::ns') ->
+ if Names.Id.equal a id then Some ns'
+ else None
+ | None -> None
+ end
+ in
+ let rec match_modulepath ns = function
+ | MPbound _ -> None (* Not a proper namespace. *)
+ | MPfile dir -> match_dirpath ns (Names.DirPath.repr dir)
+ | MPdot (mp,lbl) ->
+ let id = Names.Label.to_id lbl in
+ begin match match_modulepath ns mp with
+ | Some [] as y -> y
+ | Some (a::ns') ->
+ if Names.Id.equal a id then Some ns'
+ else None
+ | None -> None
+ end
+ in
+ (* [qualified_minus n mp] returns a list of qualifiers representing
+ [mp] except the [n] first (in the concrete syntax order). The
+ idea is that if [mp] matches [ns], then [qualified_minus mp
+ (length ns)] will be the correct representation of [mp] assuming
+ [ns] is imported. *)
+ (* precondition: [mp] matches some namespace of length [n] *)
+ let qualified_minus n mp =
+ let rec list_of_modulepath = function
+ | MPbound _ -> assert false (* MPbound never matches *)
+ | MPfile dir -> Names.DirPath.repr dir
+ | MPdot (mp,lbl) -> (Names.Label.to_id lbl)::(list_of_modulepath mp)
+ in
+ snd (Util.List.chop n (List.rev (list_of_modulepath mp)))
+ in
+ let print_list pr l = prlist_with_sep (fun () -> str".") pr l in
+ let print_kn kn =
+ (* spiwack: I'm ignoring the dirpath, is that bad? *)
+ let (mp,_,lbl) = Names.repr_kn kn in
+ let qn = (qualified_minus (List.length ns) mp)@[Names.Label.to_id lbl] in
+ print_list pr_id qn
+ in
+ let print_constant k body =
+ (* FIXME: universes *)
+ let t = Typeops.type_of_constant_type (Global.env ()) body.Declarations.const_type in
+ print_kn k ++ str":" ++ spc() ++ Printer.pr_type t
+ in
+ let matches mp = match match_modulepath ns mp with
+ | Some [] -> true
+ | _ -> false in
+ let constants = (Environ.pre_env (Global.env ())).Pre_env.env_globals.Pre_env.env_constants in
+ let constants_in_namespace =
+ Cmap_env.fold (fun c (body,_) acc ->
+ let kn = user_con c in
+ if matches (modpath kn) then
+ acc++fnl()++hov 2 (print_constant kn body)
+ else
+ acc
+ ) constants (str"")
+ in
+ Feedback.msg_notice ((print_list pr_id ns)++str":"++fnl()++constants_in_namespace)
+
+let print_strategy r =
+ let open Conv_oracle in
+ let pr_level = function
+ | Expand -> str "expand"
+ | Level 0 -> str "transparent"
+ | Level n -> str "level" ++ spc() ++ int n
+ | Opaque -> str "opaque"
+ in
+ let pr_strategy (ref, lvl) = pr_global ref ++ str " : " ++ pr_level lvl in
+ let oracle = Environ.oracle (Global.env ()) in
+ match r with
+ | None ->
+ let fold key lvl (vacc, cacc) = match key with
+ | VarKey id -> ((VarRef id, lvl) :: vacc, cacc)
+ | ConstKey cst -> (vacc, (ConstRef cst, lvl) :: cacc)
+ | RelKey _ -> (vacc, cacc)
+ in
+ let var_lvl, cst_lvl = fold_strategy fold oracle ([], []) in
+ let var_msg =
+ if List.is_empty var_lvl then mt ()
+ else str "Variable strategies" ++ fnl () ++
+ hov 0 (prlist_with_sep fnl pr_strategy var_lvl) ++ fnl ()
+ in
+ let cst_msg =
+ if List.is_empty cst_lvl then mt ()
+ else str "Constant strategies" ++ fnl () ++
+ hov 0 (prlist_with_sep fnl pr_strategy cst_lvl)
+ in
+ Feedback.msg_notice (var_msg ++ cst_msg)
+ | Some r ->
+ let r = Smartlocate.smart_global r in
+ let key = match r with
+ | VarRef id -> VarKey id
+ | ConstRef cst -> ConstKey cst
+ | IndRef _ | ConstructRef _ -> error "The reference is not unfoldable"
+ in
+ let lvl = get_strategy oracle key in
+ Feedback.msg_notice (pr_strategy (r, lvl))
+
+let dump_universes_gen g s =
+ let output = open_out s in
+ let output_constraint, close =
+ if Filename.check_suffix s ".dot" || Filename.check_suffix s ".gv" then begin
+ (* the lazy unit is to handle errors while printing the first line *)
+ let init = lazy (Printf.fprintf output "digraph universes {\n") in
+ begin fun kind left right ->
+ let () = Lazy.force init in
+ match kind with
+ | Univ.Lt ->
+ Printf.fprintf output " \"%s\" -> \"%s\" [style=bold];\n" right left
+ | Univ.Le ->
+ Printf.fprintf output " \"%s\" -> \"%s\" [style=solid];\n" right left
+ | Univ.Eq ->
+ Printf.fprintf output " \"%s\" -> \"%s\" [style=dashed];\n" left right
+ end, begin fun () ->
+ if Lazy.is_val init then Printf.fprintf output "}\n";
+ close_out output
+ end
+ end else begin
+ begin fun kind left right ->
+ let kind = match kind with
+ | Univ.Lt -> "<"
+ | Univ.Le -> "<="
+ | Univ.Eq -> "="
+ in Printf.fprintf output "%s %s %s ;\n" left kind right
+ end, (fun () -> close_out output)
+ end
+ in
+ try
+ UGraph.dump_universes output_constraint g;
+ close ();
+ Feedback.msg_info (str "Universes written to file \"" ++ str s ++ str "\".")
+ with reraise ->
+ let reraise = CErrors.push reraise in
+ close ();
+ iraise reraise
+
+(*********************)
+(* "Locate" commands *)
+
+let locate_file f =
+ let file = Flags.silently Loadpath.locate_file f in
+ str file
+
+let msg_found_library = function
+ | Library.LibLoaded, fulldir, file ->
+ Feedback.msg_info (hov 0
+ (pr_dirpath fulldir ++ strbrk " has been loaded from file " ++
+ str file))
+ | Library.LibInPath, fulldir, file ->
+ Feedback.msg_info (hov 0
+ (pr_dirpath fulldir ++ strbrk " is bound to file " ++ str file))
+
+let err_unmapped_library loc ?from qid =
+ let dir = fst (repr_qualid qid) in
+ let prefix = match from with
+ | None -> str "."
+ | Some from ->
+ str " and prefix " ++ pr_dirpath from ++ str "."
+ in
+ user_err ~loc
+ ~hdr:"locate_library"
+ (strbrk "Cannot find a physical path bound to logical path matching suffix " ++
+ pr_dirpath dir ++ prefix)
+
+let err_notfound_library loc ?from qid =
+ let prefix = match from with
+ | None -> str "."
+ | Some from ->
+ str " with prefix " ++ pr_dirpath from ++ str "."
+ in
+ user_err ~loc ~hdr:"locate_library"
+ (strbrk "Unable to locate library " ++ pr_qualid qid ++ prefix)
+
+let print_located_library r =
+ let (loc,qid) = qualid_of_reference r in
+ try msg_found_library (Library.locate_qualified_library ~warn:false qid)
+ with
+ | Library.LibUnmappedDir -> err_unmapped_library loc qid
+ | Library.LibNotFound -> err_notfound_library loc qid
+
+let smart_global r =
+ let gr = Smartlocate.smart_global r in
+ Dumpglob.add_glob (Stdarg.loc_of_or_by_notation loc_of_reference r) gr;
+ gr
+
+let dump_global r =
+ try
+ let gr = Smartlocate.smart_global r in
+ Dumpglob.add_glob (Stdarg.loc_of_or_by_notation loc_of_reference r) gr
+ with e when CErrors.noncritical e -> ()
+(**********)
+(* Syntax *)
+
+let vernac_syntax_extension locality local =
+ let local = enforce_module_locality locality local in
+ Metasyntax.add_syntax_extension local
+
+let vernac_delimiters sc = function
+ | Some lr -> Metasyntax.add_delimiters sc lr
+ | None -> Metasyntax.remove_delimiters sc
+
+let vernac_bind_scope sc cll =
+ Metasyntax.add_class_scope sc (List.map scope_class_of_qualid cll)
+
+let vernac_open_close_scope locality local (b,s) =
+ let local = enforce_section_locality locality local in
+ Notation.open_close_scope (local,b,s)
+
+let vernac_arguments_scope locality r scl =
+ let local = make_section_locality locality in
+ Notation.declare_arguments_scope local (smart_global r) scl
+
+let vernac_infix locality local =
+ let local = enforce_module_locality locality local in
+ Metasyntax.add_infix local
+
+let vernac_notation locality local =
+ let local = enforce_module_locality locality local in
+ Metasyntax.add_notation local
+
+(***********)
+(* Gallina *)
+
+let start_proof_and_print k l hook =
+ let inference_hook =
+ if Flags.is_program_mode () then
+ let hook env sigma ev =
+ let tac = !Obligations.default_tactic in
+ let evi = Evd.find sigma ev in
+ let env = Evd.evar_filtered_env evi in
+ try
+ let concl = Evarutil.nf_evars_universes sigma evi.Evd.evar_concl in
+ if Evarutil.has_undefined_evars sigma concl then raise Exit;
+ let c, _, ctx =
+ Pfedit.build_by_tactic env (Evd.evar_universe_context sigma)
+ concl (Tacticals.New.tclCOMPLETE tac)
+ in Evd.set_universe_context sigma ctx, c
+ with Logic_monad.TacticFailure e when Logic.catchable_exception e ->
+ error "The statement obligations could not be resolved \
+ automatically, write a statement definition first."
+ in Some hook
+ else None
+ in
+ start_proof_com ?inference_hook k l hook
+
+let no_hook = Lemmas.mk_hook (fun _ _ -> ())
+
+let vernac_definition_hook p = function
+| Coercion -> Class.add_coercion_hook p
+| CanonicalStructure ->
+ Lemmas.mk_hook (fun _ -> Recordops.declare_canonical_structure)
+| SubClass -> Class.add_subclass_hook p
+| _ -> no_hook
+
+let vernac_definition locality p (local,k) ((loc,id as lid),pl) def =
+ let local = enforce_locality_exp locality local in
+ let hook = vernac_definition_hook p k in
+ let () = match local with
+ | Discharge -> Dumpglob.dump_definition lid true "var"
+ | Local | Global -> Dumpglob.dump_definition lid false "def"
+ in
+ (match def with
+ | ProveBody (bl,t) -> (* local binders, typ *)
+ start_proof_and_print (local,p,DefinitionBody k)
+ [Some (lid,pl), (bl,t,None)] hook
+ | DefineBody (bl,red_option,c,typ_opt) ->
+ let red_option = match red_option with
+ | None -> None
+ | Some r ->
+ let (evc,env)= get_current_context () in
+ Some (snd (Hook.get f_interp_redexp env evc r)) in
+ do_definition id (local,p,k) pl bl red_option c typ_opt hook)
+
+let vernac_start_proof locality p kind l lettop =
+ let local = enforce_locality_exp locality None in
+ if Dumpglob.dump () then
+ List.iter (fun (id, _) ->
+ match id with
+ | Some (lid,_) -> Dumpglob.dump_definition lid false "prf"
+ | None -> ()) l;
+ if not(refining ()) then
+ if lettop then
+ user_err ~hdr:"Vernacentries.StartProof"
+ (str "Let declarations can only be used in proof editing mode.");
+ start_proof_and_print (local, p, Proof kind) l no_hook
+
+let qed_display_script = ref true
+
+let vernac_end_proof ?proof = function
+ | Admitted -> save_proof ?proof Admitted
+ | Proved (_,_) as e ->
+ (*
+ if is_verbose () && !qed_display_script && !Flags.coqtop_ui then
+ Stm.show_script ?proof ();
+ *)
+ save_proof ?proof e
+
+ (* A stupid macro that should be replaced by ``Exact c. Save.'' all along
+ the theories [??] *)
+
+let vernac_exact_proof c =
+ (* spiwack: for simplicity I do not enforce that "Proof proof_term" is
+ called only at the begining of a proof. *)
+ let status = by (Tactics.exact_proof c) in
+ save_proof (Vernacexpr.(Proved(Opaque None,None)));
+ if not status then Feedback.feedback Feedback.AddedAxiom
+
+let vernac_assumption locality poly (local, kind) l nl =
+ let local = enforce_locality_exp locality local in
+ let global = local == Global in
+ let kind = local, poly, kind in
+ List.iter (fun (is_coe,(idl,c)) ->
+ if Dumpglob.dump () then
+ List.iter (fun (lid, _) ->
+ if global then Dumpglob.dump_definition lid false "ax"
+ else Dumpglob.dump_definition lid true "var") idl) l;
+ let status = do_assumptions kind nl l in
+ if not status then Feedback.feedback Feedback.AddedAxiom
+
+let vernac_record k poly finite struc binders sort nameopt cfs =
+ let const = match nameopt with
+ | None -> add_prefix "Build_" (snd (fst (snd struc)))
+ | Some (_,id as lid) ->
+ Dumpglob.dump_definition lid false "constr"; id in
+ if Dumpglob.dump () then (
+ Dumpglob.dump_definition (fst (snd struc)) false "rec";
+ List.iter (fun (((_, x), _), _) ->
+ match x with
+ | Vernacexpr.AssumExpr ((loc, Name id), _) -> Dumpglob.dump_definition (loc,id) false "proj"
+ | _ -> ()) cfs);
+ ignore(Record.definition_structure (k,poly,finite,struc,binders,cfs,const,sort))
+
+(** When [poly] is true the type is declared polymorphic. When [lo] is true,
+ then the type is declared private (as per the [Private] keyword). [finite]
+ indicates whether the type is inductive, co-inductive or
+ neither. *)
+let vernac_inductive poly lo finite indl =
+ if Dumpglob.dump () then
+ List.iter (fun (((coe,(lid,_)), _, _, _, cstrs), _) ->
+ match cstrs with
+ | Constructors cstrs ->
+ Dumpglob.dump_definition lid false "ind";
+ List.iter (fun (_, (lid, _)) ->
+ Dumpglob.dump_definition lid false "constr") cstrs
+ | _ -> () (* dumping is done by vernac_record (called below) *) )
+ indl;
+ match indl with
+ | [ ( _ , _ , _ ,(Record|Structure), Constructors _ ),_ ] ->
+ CErrors.error "The Record keyword is for types defined using the syntax { ... }."
+ | [ (_ , _ , _ ,Variant, RecordDecl _),_ ] ->
+ CErrors.error "The Variant keyword does not support syntax { ... }."
+ | [ ( id , bl , c , b, RecordDecl (oc,fs) ), [] ] ->
+ vernac_record (match b with Class _ -> Class false | _ -> b)
+ poly finite id bl c oc fs
+ | [ ( id , bl , c , Class _, Constructors [l]), [] ] ->
+ let f =
+ let (coe, ((loc, id), ce)) = l in
+ let coe' = if coe then Some true else None in
+ (((coe', AssumExpr ((loc, Name id), ce)), None), [])
+ in vernac_record (Class true) poly finite id bl c None [f]
+ | [ ( _ , _, _, Class _, Constructors _), [] ] ->
+ CErrors.error "Inductive classes not supported"
+ | [ ( id , bl , c , Class _, _), _ :: _ ] ->
+ CErrors.error "where clause not supported for classes"
+ | [ ( _ , _ , _ , _, RecordDecl _ ) , _ ] ->
+ CErrors.error "where clause not supported for (co)inductive records"
+ | _ -> let unpack = function
+ | ( (false, id) , bl , c , _ , Constructors l ) , ntn -> ( id , bl , c , l ) , ntn
+ | ( (true,_),_,_,_,Constructors _),_ ->
+ CErrors.error "Variant types do not handle the \"> Name\" syntax, which is reserved for records. Use the \":>\" syntax on constructors instead."
+ | _ -> CErrors.error "Cannot handle mutually (co)inductive records."
+ in
+ let indl = List.map unpack indl in
+ do_mutual_inductive indl poly lo finite
+
+let vernac_fixpoint locality poly local l =
+ let local = enforce_locality_exp locality local in
+ if Dumpglob.dump () then
+ List.iter (fun (((lid,_), _, _, _, _), _) -> Dumpglob.dump_definition lid false "def") l;
+ do_fixpoint local poly l
+
+let vernac_cofixpoint locality poly local l =
+ let local = enforce_locality_exp locality local in
+ if Dumpglob.dump () then
+ List.iter (fun (((lid,_), _, _, _), _) -> Dumpglob.dump_definition lid false "def") l;
+ do_cofixpoint local poly l
+
+let vernac_scheme l =
+ if Dumpglob.dump () then
+ List.iter (fun (lid, s) ->
+ Option.iter (fun lid -> Dumpglob.dump_definition lid false "def") lid;
+ match s with
+ | InductionScheme (_, r, _)
+ | CaseScheme (_, r, _)
+ | EqualityScheme r -> dump_global r) l;
+ Indschemes.do_scheme l
+
+let vernac_combined_scheme lid l =
+ if Dumpglob.dump () then
+ (Dumpglob.dump_definition lid false "def";
+ List.iter (fun lid -> dump_global (Misctypes.AN (Ident lid))) l);
+ Indschemes.do_combined_scheme lid l
+
+let vernac_universe loc poly l =
+ if poly && not (Lib.sections_are_opened ()) then
+ user_err ~loc ~hdr:"vernac_universe"
+ (str"Polymorphic universes can only be declared inside sections, " ++
+ str "use Monomorphic Universe instead");
+ do_universe poly l
+
+let vernac_constraint loc poly l =
+ if poly && not (Lib.sections_are_opened ()) then
+ user_err ~loc ~hdr:"vernac_constraint"
+ (str"Polymorphic universe constraints can only be declared"
+ ++ str " inside sections, use Monomorphic Constraint instead");
+ do_constraint poly l
+
+(**********************)
+(* Modules *)
+
+let vernac_import export refl =
+ Library.import_module export (List.map qualid_of_reference refl);
+ Lib.add_frozen_state ()
+
+let vernac_declare_module export (loc, id) binders_ast mty_ast =
+ (* We check the state of the system (in section, in module type)
+ and what module information is supplied *)
+ if Lib.sections_are_opened () then
+ error "Modules and Module Types are not allowed inside sections.";
+ let binders_ast = List.map
+ (fun (export,idl,ty) ->
+ if not (Option.is_empty export) then
+ error "Arguments of a functor declaration cannot be exported. Remove the \"Export\" and \"Import\" keywords from every functor argument."
+ else (idl,ty)) binders_ast in
+ let mp =
+ Declaremods.declare_module Modintern.interp_module_ast
+ id binders_ast (Enforce mty_ast) []
+ in
+ Dumpglob.dump_moddef loc mp "mod";
+ if_verbose Feedback.msg_info (str "Module " ++ pr_id id ++ str " is declared");
+ Option.iter (fun export -> vernac_import export [Ident (Loc.ghost,id)]) export
+
+let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l =
+ (* We check the state of the system (in section, in module type)
+ and what module information is supplied *)
+ if Lib.sections_are_opened () then
+ error "Modules and Module Types are not allowed inside sections.";
+ match mexpr_ast_l with
+ | [] ->
+ check_no_pending_proofs ();
+ let binders_ast,argsexport =
+ List.fold_right
+ (fun (export,idl,ty) (args,argsexport) ->
+ (idl,ty)::args, (List.map (fun (_,i) -> export,i)idl)@argsexport) binders_ast
+ ([],[]) in
+ let mp =
+ Declaremods.start_module Modintern.interp_module_ast
+ export id binders_ast mty_ast_o
+ in
+ Dumpglob.dump_moddef loc mp "mod";
+ if_verbose Feedback.msg_info
+ (str "Interactive Module " ++ pr_id id ++ str " started");
+ List.iter
+ (fun (export,id) ->
+ Option.iter
+ (fun export -> vernac_import export [Ident (Loc.ghost,id)]) export
+ ) argsexport
+ | _::_ ->
+ let binders_ast = List.map
+ (fun (export,idl,ty) ->
+ if not (Option.is_empty export) then
+ error "Arguments of a functor definition can be imported only if the definition is interactive. Remove the \"Export\" and \"Import\" keywords from every functor argument."
+ else (idl,ty)) binders_ast in
+ let mp =
+ Declaremods.declare_module Modintern.interp_module_ast
+ id binders_ast mty_ast_o mexpr_ast_l
+ in
+ Dumpglob.dump_moddef loc mp "mod";
+ if_verbose Feedback.msg_info
+ (str "Module " ++ pr_id id ++ str " is defined");
+ Option.iter (fun export -> vernac_import export [Ident (Loc.ghost,id)])
+ export
+
+let vernac_end_module export (loc,id as lid) =
+ let mp = Declaremods.end_module () in
+ Dumpglob.dump_modref loc mp "mod";
+ if_verbose Feedback.msg_info (str "Module " ++ pr_id id ++ str " is defined");
+ Option.iter (fun export -> vernac_import export [Ident lid]) export
+
+let vernac_declare_module_type (loc,id) binders_ast mty_sign mty_ast_l =
+ if Lib.sections_are_opened () then
+ error "Modules and Module Types are not allowed inside sections.";
+
+ match mty_ast_l with
+ | [] ->
+ check_no_pending_proofs ();
+ let binders_ast,argsexport =
+ List.fold_right
+ (fun (export,idl,ty) (args,argsexport) ->
+ (idl,ty)::args, (List.map (fun (_,i) -> export,i)idl)@argsexport) binders_ast
+ ([],[]) in
+
+ let mp =
+ Declaremods.start_modtype Modintern.interp_module_ast
+ id binders_ast mty_sign
+ in
+ Dumpglob.dump_moddef loc mp "modtype";
+ if_verbose Feedback.msg_info
+ (str "Interactive Module Type " ++ pr_id id ++ str " started");
+ List.iter
+ (fun (export,id) ->
+ Option.iter
+ (fun export -> vernac_import export [Ident (Loc.ghost,id)]) export
+ ) argsexport
+
+ | _ :: _ ->
+ let binders_ast = List.map
+ (fun (export,idl,ty) ->
+ if not (Option.is_empty export) then
+ error "Arguments of a functor definition can be imported only if the definition is interactive. Remove the \"Export\" and \"Import\" keywords from every functor argument."
+ else (idl,ty)) binders_ast in
+ let mp =
+ Declaremods.declare_modtype Modintern.interp_module_ast
+ id binders_ast mty_sign mty_ast_l
+ in
+ Dumpglob.dump_moddef loc mp "modtype";
+ if_verbose Feedback.msg_info
+ (str "Module Type " ++ pr_id id ++ str " is defined")
+
+let vernac_end_modtype (loc,id) =
+ let mp = Declaremods.end_modtype () in
+ Dumpglob.dump_modref loc mp "modtype";
+ if_verbose Feedback.msg_info (str "Module Type " ++ pr_id id ++ str " is defined")
+
+let vernac_include l =
+ Declaremods.declare_include Modintern.interp_module_ast l
+
+(**********************)
+(* Gallina extensions *)
+
+(* Sections *)
+
+let vernac_begin_section (_, id as lid) =
+ check_no_pending_proofs ();
+ Dumpglob.dump_definition lid true "sec";
+ Lib.open_section id
+
+let vernac_end_section (loc,_) =
+ Dumpglob.dump_reference loc
+ (DirPath.to_string (Lib.current_dirpath true)) "<>" "sec";
+ Lib.close_section ()
+
+let vernac_name_sec_hyp (_,id) set = Proof_using.name_set id set
+
+(* Dispatcher of the "End" command *)
+
+let vernac_end_segment (_,id as lid) =
+ check_no_pending_proofs ();
+ match Lib.find_opening_node id with
+ | Lib.OpenedModule (false,export,_,_) -> vernac_end_module export lid
+ | Lib.OpenedModule (true,_,_,_) -> vernac_end_modtype lid
+ | Lib.OpenedSection _ -> vernac_end_section lid
+ | _ -> assert false
+
+(* Libraries *)
+
+let vernac_require from import qidl =
+ let qidl = List.map qualid_of_reference qidl in
+ let root = match from with
+ | None -> None
+ | Some from ->
+ let (_, qid) = Libnames.qualid_of_reference from in
+ let (hd, tl) = Libnames.repr_qualid qid in
+ Some (Libnames.add_dirpath_suffix hd tl)
+ in
+ let locate (loc, qid) =
+ try
+ let warn = Flags.is_verbose () in
+ let (_, dir, f) = Library.locate_qualified_library ?root ~warn qid in
+ (dir, f)
+ with
+ | Library.LibUnmappedDir -> err_unmapped_library loc ?from:root qid
+ | Library.LibNotFound -> err_notfound_library loc ?from:root qid
+ in
+ let modrefl = List.map locate qidl in
+ if Dumpglob.dump () then
+ List.iter2 (fun (loc, _) dp -> Dumpglob.dump_libref loc dp "lib") qidl (List.map fst modrefl);
+ Library.require_library_from_dirpath modrefl import
+
+(* Coercions and canonical structures *)
+
+let vernac_canonical r =
+ Recordops.declare_canonical_structure (smart_global r)
+
+let vernac_coercion locality poly local ref qids qidt =
+ let local = enforce_locality locality local in
+ let target = cl_of_qualid qidt in
+ let source = cl_of_qualid qids in
+ let ref' = smart_global ref in
+ Class.try_add_new_coercion_with_target ref' ~local poly ~source ~target;
+ if_verbose Feedback.msg_info (pr_global ref' ++ str " is now a coercion")
+
+let vernac_identity_coercion locality poly local id qids qidt =
+ let local = enforce_locality locality local in
+ let target = cl_of_qualid qidt in
+ let source = cl_of_qualid qids in
+ Class.try_add_new_identity_coercion id ~local poly ~source ~target
+
+(* Type classes *)
+
+let vernac_instance abst locality poly sup inst props pri =
+ let global = not (make_section_locality locality) in
+ Dumpglob.dump_constraint inst false "inst";
+ ignore(Classes.new_instance ~abstract:abst ~global poly sup inst props pri)
+
+let vernac_context poly l =
+ if not (Classes.context poly l) then Feedback.feedback Feedback.AddedAxiom
+
+let vernac_declare_instances locality insts =
+ let glob = not (make_section_locality locality) in
+ List.iter (fun (id, info) -> Classes.existing_instance glob id (Some info)) insts
+
+let vernac_declare_class id =
+ Record.declare_existing_class (Nametab.global id)
+
+(***********)
+(* Solving *)
+
+let command_focus = Proof.new_focus_kind ()
+let focus_command_cond = Proof.no_cond command_focus
+
+ (* A command which should be a tactic. It has been
+ added by Christine to patch an error in the design of the proof
+ machine, and enables to instantiate existential variables when
+ there are no more goals to solve. It cannot be a tactic since
+ all tactics fail if there are no further goals to prove. *)
+
+let vernac_solve_existential = instantiate_nth_evar_com
+
+let vernac_set_end_tac tac =
+ let open Genintern in
+ let env = { genv = Global.env (); ltacvars = Id.Set.empty } in
+ let _, tac = Genintern.generic_intern env tac in
+ if not (refining ()) then
+ error "Unknown command of the non proof-editing mode.";
+ set_end_tac tac
+ (* TO DO verifier s'il faut pas mettre exist s | TacId s ici*)
+
+let vernac_set_used_variables e =
+ let env = Global.env () in
+ let tys =
+ List.map snd (Proof.initial_goals (Proof_global.give_me_the_proof ())) in
+ let l = Proof_using.process_expr env e tys in
+ let vars = Environ.named_context env in
+ List.iter (fun id ->
+ if not (List.exists (NamedDecl.get_id %> Id.equal id) vars) then
+ user_err ~hdr:"vernac_set_used_variables"
+ (str "Unknown variable: " ++ pr_id id))
+ l;
+ let _, to_clear = set_used_variables l in
+ let to_clear = List.map snd to_clear in
+ Proof_global.with_current_proof begin fun _ p ->
+ if List.is_empty to_clear then (p, ())
+ else
+ let tac = Tactics.clear to_clear in
+ fst (solve SelectAll None tac p), ()
+ end
+
+(*****************************)
+(* Auxiliary file management *)
+
+let expand filename =
+ Envars.expand_path_macros ~warn:(fun x -> Feedback.msg_warning (str x)) filename
+
+let vernac_add_loadpath implicit pdir ldiropt =
+ let pdir = expand pdir in
+ let alias = Option.default Nameops.default_root_prefix ldiropt in
+ Mltop.add_rec_path Mltop.AddTopML ~unix_path:pdir ~coq_root:alias ~implicit
+
+let vernac_remove_loadpath path =
+ Loadpath.remove_load_path (expand path)
+
+ (* Coq syntax for ML or system commands *)
+
+let vernac_add_ml_path isrec path =
+ (if isrec then Mltop.add_rec_ml_dir else Mltop.add_ml_dir) (expand path)
+
+let vernac_declare_ml_module locality l =
+ let local = make_locality locality in
+ Mltop.declare_ml_modules local (List.map expand l)
+
+let vernac_chdir = function
+ | None -> Feedback.msg_notice (str (Sys.getcwd()))
+ | Some path ->
+ begin
+ try Sys.chdir (expand path)
+ with Sys_error err ->
+ (* Cd is typically used to control the output directory of
+ extraction. A failed Cd could lead to overwriting .ml files
+ so we make it an error. *)
+ CErrors.error ("Cd failed: " ^ err)
+ end;
+ if_verbose Feedback.msg_info (str (Sys.getcwd()))
+
+
+(********************)
+(* State management *)
+
+let vernac_write_state file =
+ Pfedit.delete_all_proofs ();
+ let file = CUnix.make_suffix file ".coq" in
+ States.extern_state file
+
+let vernac_restore_state file =
+ Pfedit.delete_all_proofs ();
+ let file = Loadpath.locate_file (CUnix.make_suffix file ".coq") in
+ States.intern_state file
+
+(************)
+(* Commands *)
+
+let vernac_create_hintdb locality id b =
+ let local = make_module_locality locality in
+ Hints.create_hint_db local id full_transparent_state b
+
+let vernac_remove_hints locality dbs ids =
+ let local = make_module_locality locality in
+ Hints.remove_hints local dbs (List.map Smartlocate.global_with_alias ids)
+
+let vernac_hints locality poly local lb h =
+ let local = enforce_module_locality locality local in
+ Hints.add_hints local lb (Hints.interp_hints poly h)
+
+let vernac_syntactic_definition locality lid x local y =
+ Dumpglob.dump_definition lid false "syndef";
+ let local = enforce_module_locality locality local in
+ Metasyntax.add_syntactic_definition (snd lid) x local y
+
+let vernac_declare_implicits locality r l =
+ let local = make_section_locality locality in
+ match l with
+ | [] ->
+ Impargs.declare_implicits local (smart_global r)
+ | _::_ as imps ->
+ Impargs.declare_manual_implicits local (smart_global r) ~enriching:false
+ (List.map (List.map (fun (ex,b,f) -> ex, (b,true,f))) imps)
+
+let warn_arguments_assert =
+ CWarnings.create ~name:"arguments-assert" ~category:"vernacular"
+ (fun sr ->
+ strbrk "This command is just asserting the names of arguments of " ++
+ pr_global sr ++ strbrk". If this is what you want add " ++
+ strbrk "': assert' to silence the warning. If you want " ++
+ strbrk "to clear implicit arguments add ': clear implicits'. " ++
+ strbrk "If you want to clear notation scopes add ': clear scopes'")
+
+(* [nargs_for_red] is the number of arguments required to trigger reduction,
+ [args] is the main list of arguments statuses,
+ [more_implicits] is a list of extra lists of implicit statuses *)
+let vernac_arguments locality reference args more_implicits nargs_for_red flags =
+ let assert_flag = List.mem `Assert flags in
+ let rename_flag = List.mem `Rename flags in
+ let clear_scopes_flag = List.mem `ClearScopes flags in
+ let extra_scopes_flag = List.mem `ExtraScopes flags in
+ let clear_implicits_flag = List.mem `ClearImplicits flags in
+ let default_implicits_flag = List.mem `DefaultImplicits flags in
+ let never_unfold_flag = List.mem `ReductionNeverUnfold flags in
+
+ let err_incompat x y =
+ error ("Options \""^x^"\" and \""^y^"\" are incompatible.") in
+
+ if assert_flag && rename_flag then
+ err_incompat "assert" "rename";
+ if Option.has_some nargs_for_red && never_unfold_flag then
+ err_incompat "simpl never" "/";
+ if never_unfold_flag && List.mem `ReductionDontExposeCase flags then
+ err_incompat "simpl never" "simpl nomatch";
+ if clear_scopes_flag && extra_scopes_flag then
+ err_incompat "clear scopes" "extra scopes";
+ if clear_implicits_flag && default_implicits_flag then
+ err_incompat "clear implicits" "default implicits";
+
+ let sr = smart_global reference in
+ let inf_names =
+ let ty = Global.type_of_global_unsafe sr in
+ Impargs.compute_implicits_names (Global.env ()) ty
+ in
+ let prev_names =
+ try Arguments_renaming.arguments_names sr with Not_found -> inf_names
+ in
+ let num_args = List.length inf_names in
+ assert (Int.equal num_args (List.length prev_names));
+
+ let names_of args = List.map (fun a -> a.name) args in
+
+ (* Checks *)
+
+ let err_extra_args names =
+ user_err ~hdr:"vernac_declare_arguments"
+ (strbrk "Extra arguments: " ++
+ prlist_with_sep pr_comma pr_name names ++ str ".")
+ in
+ let err_missing_args names =
+ user_err ~hdr:"vernac_declare_arguments"
+ (strbrk "The following arguments are not declared: " ++
+ prlist_with_sep pr_comma pr_name names ++ str ".")
+ in
+
+ let rec check_extra_args extra_args =
+ match extra_args with
+ | [] -> ()
+ | { notation_scope = None } :: _ -> err_extra_args (names_of extra_args)
+ | { name = Anonymous; notation_scope = Some _ } :: args ->
+ check_extra_args args
+ | _ ->
+ error "Extra notation scopes can be set on anonymous and explicit arguments only."
+ in
+
+ let args, scopes =
+ let scopes = List.map (fun { notation_scope = s } -> s) args in
+ if List.length args > num_args then
+ let args, extra_args = List.chop num_args args in
+ if extra_scopes_flag then
+ (check_extra_args extra_args; (args, scopes))
+ else err_extra_args (names_of extra_args)
+ else args, scopes
+ in
+
+ if Option.cata (fun n -> n > num_args) false nargs_for_red then
+ error "The \"/\" modifier should be put before any extra scope.";
+
+ let scopes_specified = List.exists Option.has_some scopes in
+
+ if scopes_specified && clear_scopes_flag then
+ error "The \"clear scopes\" flag is incompatible with scope annotations.";
+
+ let names = List.map (fun { name } -> name) args in
+ let names = names :: List.map (List.map fst) more_implicits in
+
+ let rename_flag_required = ref false in
+ let example_renaming = ref None in
+ let save_example_renaming renaming =
+ rename_flag_required := !rename_flag_required
+ || not (Name.equal (fst renaming) Anonymous);
+ if Option.is_empty !example_renaming then
+ example_renaming := Some renaming
+ in
+
+ let rec names_union names1 names2 =
+ match names1, names2 with
+ | [], [] -> []
+ | _ :: _, [] -> names1
+ | [], _ :: _ -> names2
+ | (Name _ as name) :: names1, Anonymous :: names2
+ | Anonymous :: names1, (Name _ as name) :: names2 ->
+ name :: names_union names1 names2
+ | name1 :: names1, name2 :: names2 ->
+ if Name.equal name1 name2 then
+ name1 :: names_union names1 names2
+ else error "Argument lists should agree on the names they provide."
+ in
+
+ let names = List.fold_left names_union [] names in
+
+ let rec rename prev_names names =
+ match prev_names, names with
+ | [], [] -> []
+ | [], _ :: _ -> err_extra_args names
+ | _ :: _, [] when assert_flag ->
+ (* Error messages are expressed in terms of original names, not
+ renamed ones. *)
+ err_missing_args (List.lastn (List.length prev_names) inf_names)
+ | _ :: _, [] -> prev_names
+ | prev :: prev_names, Anonymous :: names ->
+ prev :: rename prev_names names
+ | prev :: prev_names, (Name id as name) :: names ->
+ if not (Name.equal prev name) then save_example_renaming (prev,name);
+ name :: rename prev_names names
+ in
+
+ let names = rename prev_names names in
+ let renaming_specified = Option.has_some !example_renaming in
+
+ if !rename_flag_required && not rename_flag then
+ user_err ~hdr:"vernac_declare_arguments"
+ (strbrk "To rename arguments the \"rename\" flag must be specified."
+ ++ spc () ++
+ match !example_renaming with
+ | None -> mt ()
+ | Some (o,n) ->
+ str "Argument " ++ pr_name o ++
+ str " renamed to " ++ pr_name n ++ str ".");
+
+ let duplicate_names =
+ List.duplicates Name.equal (List.filter ((!=) Anonymous) names)
+ in
+ if not (List.is_empty duplicate_names) then begin
+ let duplicates = prlist_with_sep pr_comma pr_name duplicate_names in
+ user_err (strbrk "Some argument names are duplicated: " ++ duplicates)
+ end;
+
+ (* Parts of this code are overly complicated because the implicit arguments
+ API is completely crazy: positions (ExplByPos) are elaborated to
+ names. This is broken by design, since not all arguments have names. So
+ even though we eventually want to map only positions to implicit statuses,
+ we have to check whether the corresponding arguments have names, not to
+ trigger an error in the impargs code. Even better, the names we have to
+ check are not the current ones (after previous renamings), but the original
+ ones (inferred from the type). *)
+
+ let implicits =
+ List.map (fun { name; implicit_status = i } -> (name,i)) args
+ in
+ let implicits = implicits :: more_implicits in
+
+ let open Vernacexpr in
+ let rec build_implicits inf_names implicits =
+ match inf_names, implicits with
+ | _, [] -> []
+ | _ :: inf_names, (_, NotImplicit) :: implicits ->
+ build_implicits inf_names implicits
+
+ (* With the current impargs API, it is impossible to make an originally
+ anonymous argument implicit *)
+ | Anonymous :: _, (name, _) :: _ ->
+ user_err ~hdr:"vernac_declare_arguments"
+ (strbrk"Argument "++ pr_name name ++
+ strbrk " cannot be declared implicit.")
+
+ | Name id :: inf_names, (name, impl) :: implicits ->
+ let max = impl = MaximallyImplicit in
+ (ExplByName id,max,false) :: build_implicits inf_names implicits
+
+ | _ -> assert false (* already checked in [names_union] *)
+ in
+
+ let implicits = List.map (build_implicits inf_names) implicits in
+ let implicits_specified = match implicits with [[]] -> false | _ -> true in
+
+ if implicits_specified && clear_implicits_flag then
+ error "The \"clear implicits\" flag is incompatible with implicit annotations";
+
+ if implicits_specified && default_implicits_flag then
+ error "The \"default implicits\" flag is incompatible with implicit annotations";
+
+ let rargs =
+ Util.List.map_filter (function (n, true) -> Some n | _ -> None)
+ (Util.List.map_i (fun i { recarg_like = b } -> i, b) 0 args)
+ in
+
+ let rec narrow = function
+ | #Reductionops.ReductionBehaviour.flag as x :: tl -> x :: narrow tl
+ | [] -> [] | _ :: tl -> narrow tl
+ in
+ let red_flags = narrow flags in
+ let red_modifiers_specified =
+ not (List.is_empty rargs) || Option.has_some nargs_for_red
+ || not (List.is_empty red_flags)
+ in
+
+ if not (List.is_empty rargs) && never_unfold_flag then
+ err_incompat "simpl never" "!";
+
+
+ (* Actions *)
+
+ if renaming_specified then begin
+ let local = make_section_locality locality in
+ Arguments_renaming.rename_arguments local sr names
+ end;
+
+ if scopes_specified || clear_scopes_flag then begin
+ let scopes = List.map (Option.map (fun (o,k) ->
+ try ignore (Notation.find_scope k); k
+ with UserError _ ->
+ Notation.find_delimiters_scope o k)) scopes
+ in
+ vernac_arguments_scope locality reference scopes
+ end;
+
+ if implicits_specified || clear_implicits_flag then
+ vernac_declare_implicits locality reference implicits;
+
+ if default_implicits_flag then
+ vernac_declare_implicits locality reference [];
+
+ if red_modifiers_specified then begin
+ match sr with
+ | ConstRef _ as c ->
+ Reductionops.ReductionBehaviour.set
+ (make_section_locality locality) c
+ (rargs, Option.default ~-1 nargs_for_red, red_flags)
+ | _ -> user_err
+ (strbrk "Modifiers of the behavior of the simpl tactic "++
+ strbrk "are relevant for constants only.")
+ end;
+
+ if not (renaming_specified ||
+ implicits_specified ||
+ scopes_specified ||
+ red_modifiers_specified) && (List.is_empty flags) then
+ warn_arguments_assert sr
+
+let default_env () = {
+ Notation_term.ninterp_var_type = Id.Map.empty;
+ ninterp_rec_vars = Id.Map.empty;
+}
+
+let vernac_reserve bl =
+ let sb_decl = (fun (idl,c) ->
+ let env = Global.env() in
+ let sigma = Evd.from_env env in
+ let t,ctx = Constrintern.interp_type env sigma c in
+ let t = Detyping.detype false [] env (Evd.from_ctx ctx) t in
+ let t,_ = Notation_ops.notation_constr_of_glob_constr (default_env ()) t in
+ Reserve.declare_reserved_type idl t)
+ in List.iter sb_decl bl
+
+let vernac_generalizable locality =
+ let local = make_non_locality locality in
+ Implicit_quantifiers.declare_generalizable local
+
+let _ =
+ declare_bool_option
+ { optsync = false;
+ optdepr = false;
+ optname = "silent";
+ optkey = ["Silent"];
+ optread = is_silent;
+ optwrite = make_silent }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "implicit arguments";
+ optkey = ["Implicit";"Arguments"];
+ optread = Impargs.is_implicit_args;
+ optwrite = Impargs.make_implicit_args }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "strict implicit arguments";
+ optkey = ["Strict";"Implicit"];
+ optread = Impargs.is_strict_implicit_args;
+ optwrite = Impargs.make_strict_implicit_args }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "strong strict implicit arguments";
+ optkey = ["Strongly";"Strict";"Implicit"];
+ optread = Impargs.is_strongly_strict_implicit_args;
+ optwrite = Impargs.make_strongly_strict_implicit_args }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "contextual implicit arguments";
+ optkey = ["Contextual";"Implicit"];
+ optread = Impargs.is_contextual_implicit_args;
+ optwrite = Impargs.make_contextual_implicit_args }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "implicit status of reversible patterns";
+ optkey = ["Reversible";"Pattern";"Implicit"];
+ optread = Impargs.is_reversible_pattern_implicit_args;
+ optwrite = Impargs.make_reversible_pattern_implicit_args }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "maximal insertion of implicit";
+ optkey = ["Maximal";"Implicit";"Insertion"];
+ optread = Impargs.is_maximal_implicit_args;
+ optwrite = Impargs.make_maximal_implicit_args }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "automatic introduction of variables";
+ optkey = ["Automatic";"Introduction"];
+ optread = Flags.is_auto_intros;
+ optwrite = make_auto_intros }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "coercion printing";
+ optkey = ["Printing";"Coercions"];
+ optread = (fun () -> !Constrextern.print_coercions);
+ optwrite = (fun b -> Constrextern.print_coercions := b) }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "printing of existential variable instances";
+ optkey = ["Printing";"Existential";"Instances"];
+ optread = (fun () -> !Detyping.print_evar_arguments);
+ optwrite = (:=) Detyping.print_evar_arguments }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "implicit arguments printing";
+ optkey = ["Printing";"Implicit"];
+ optread = (fun () -> !Constrextern.print_implicits);
+ optwrite = (fun b -> Constrextern.print_implicits := b) }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "implicit arguments defensive printing";
+ optkey = ["Printing";"Implicit";"Defensive"];
+ optread = (fun () -> !Constrextern.print_implicits_defensive);
+ optwrite = (fun b -> Constrextern.print_implicits_defensive := b) }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "projection printing using dot notation";
+ optkey = ["Printing";"Projections"];
+ optread = (fun () -> !Constrextern.print_projections);
+ optwrite = (fun b -> Constrextern.print_projections := b) }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "notations printing";
+ optkey = ["Printing";"Notations"];
+ optread = (fun () -> not !Constrextern.print_no_symbol);
+ optwrite = (fun b -> Constrextern.print_no_symbol := not b) }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "raw printing";
+ optkey = ["Printing";"All"];
+ optread = (fun () -> !Flags.raw_print);
+ optwrite = (fun b -> Flags.raw_print := b) }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "record printing";
+ optkey = ["Printing";"Records"];
+ optread = (fun () -> !Flags.record_print);
+ optwrite = (fun b -> Flags.record_print := b) }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "use of the program extension";
+ optkey = ["Program";"Mode"];
+ optread = (fun () -> !Flags.program_mode);
+ optwrite = (fun b -> Flags.program_mode:=b) }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "universe polymorphism";
+ optkey = ["Universe"; "Polymorphism"];
+ optread = Flags.is_universe_polymorphism;
+ optwrite = Flags.make_universe_polymorphism }
+
+let _ =
+ declare_int_option
+ { optsync = true;
+ optdepr = false;
+ optname = "the level of inlining during functor application";
+ optkey = ["Inline";"Level"];
+ optread = (fun () -> Some (Flags.get_inline_level ()));
+ optwrite = (fun o ->
+ let lev = Option.default Flags.default_inline_level o in
+ Flags.set_inline_level lev) }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "kernel term sharing";
+ optkey = ["Kernel"; "Term"; "Sharing"];
+ optread = (fun () -> !CClosure.share);
+ optwrite = (fun b -> CClosure.share := b) }
+
+(* No more undo limit in the new proof engine.
+ The command still exists for compatibility (e.g. with ProofGeneral) *)
+
+let _ =
+ declare_int_option
+ { optsync = false;
+ optdepr = true;
+ optname = "the undo limit (OBSOLETE)";
+ optkey = ["Undo"];
+ optread = (fun _ -> None);
+ optwrite = (fun _ -> ()) }
+
+let _ =
+ declare_int_option
+ { optsync = false;
+ optdepr = false;
+ optname = "the hypotheses limit";
+ optkey = ["Hyps";"Limit"];
+ optread = Flags.print_hyps_limit;
+ optwrite = Flags.set_print_hyps_limit }
+
+let _ =
+ declare_int_option
+ { optsync = true;
+ optdepr = false;
+ optname = "the printing depth";
+ optkey = ["Printing";"Depth"];
+ optread = Pp_control.get_depth_boxes;
+ optwrite = Pp_control.set_depth_boxes }
+
+let _ =
+ declare_int_option
+ { optsync = true;
+ optdepr = false;
+ optname = "the printing width";
+ optkey = ["Printing";"Width"];
+ optread = Pp_control.get_margin;
+ optwrite = Pp_control.set_margin }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "printing of universes";
+ optkey = ["Printing";"Universes"];
+ optread = (fun () -> !Constrextern.print_universes);
+ optwrite = (fun b -> Constrextern.print_universes:=b) }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "dumping bytecode after compilation";
+ optkey = ["Dump";"Bytecode"];
+ optread = Flags.get_dump_bytecode;
+ optwrite = Flags.set_dump_bytecode }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "explicitly parsing implicit arguments";
+ optkey = ["Parsing";"Explicit"];
+ optread = (fun () -> !Constrintern.parsing_explicit);
+ optwrite = (fun b -> Constrintern.parsing_explicit := b) }
+
+let _ =
+ declare_string_option ~preprocess:CWarnings.normalize_flags_string
+ { optsync = true;
+ optdepr = false;
+ optname = "warnings display";
+ optkey = ["Warnings"];
+ optread = CWarnings.get_flags;
+ optwrite = CWarnings.set_flags }
+
+let vernac_set_strategy locality l =
+ let local = make_locality locality in
+ let glob_ref r =
+ match smart_global r with
+ | ConstRef sp -> EvalConstRef sp
+ | VarRef id -> EvalVarRef id
+ | _ -> error
+ "cannot set an inductive type or a constructor as transparent" in
+ let l = List.map (fun (lev,ql) -> (lev,List.map glob_ref ql)) l in
+ Redexpr.set_strategy local l
+
+let vernac_set_opacity locality (v,l) =
+ let local = make_non_locality locality in
+ let glob_ref r =
+ match smart_global r with
+ | ConstRef sp -> EvalConstRef sp
+ | VarRef id -> EvalVarRef id
+ | _ -> error
+ "cannot set an inductive type or a constructor as transparent" in
+ let l = List.map glob_ref l in
+ Redexpr.set_strategy local [v,l]
+
+let vernac_set_option locality key = function
+ | StringValue s -> set_string_option_value_gen locality key s
+ | StringOptValue (Some s) -> set_string_option_value_gen locality key s
+ | StringOptValue None -> unset_option_value_gen locality key
+ | IntValue n -> set_int_option_value_gen locality key n
+ | BoolValue b -> set_bool_option_value_gen locality key b
+
+let vernac_set_append_option locality key s =
+ set_string_option_append_value_gen locality key s
+
+let vernac_unset_option locality key =
+ unset_option_value_gen locality key
+
+let vernac_add_option key lv =
+ let f = function
+ | StringRefValue s -> (get_string_table key)#add s
+ | QualidRefValue locqid -> (get_ref_table key)#add locqid
+ in
+ try List.iter f lv with Not_found -> error_undeclared_key key
+
+let vernac_remove_option key lv =
+ let f = function
+ | StringRefValue s -> (get_string_table key)#remove s
+ | QualidRefValue locqid -> (get_ref_table key)#remove locqid
+ in
+ try List.iter f lv with Not_found -> error_undeclared_key key
+
+let vernac_mem_option key lv =
+ let f = function
+ | StringRefValue s -> (get_string_table key)#mem s
+ | QualidRefValue locqid -> (get_ref_table key)#mem locqid
+ in
+ try List.iter f lv with Not_found -> error_undeclared_key key
+
+let vernac_print_option key =
+ try (get_ref_table key)#print
+ with Not_found ->
+ try (get_string_table key)#print
+ with Not_found ->
+ try print_option_value key
+ with Not_found -> error_undeclared_key key
+
+let get_current_context_of_args = function
+ | Some n -> get_goal_context n
+ | None -> get_current_context ()
+
+let vernac_check_may_eval redexp glopt rc =
+ let (sigma, env) = get_current_context_of_args glopt in
+ let sigma', c = interp_open_constr env sigma rc in
+ let sigma' = Evarconv.solve_unif_constraints_with_heuristics env sigma' in
+ Evarconv.check_problems_are_solved env sigma';
+ let sigma',nf = Evarutil.nf_evars_and_universes sigma' in
+ let pl, uctx = Evd.universe_context sigma' in
+ let env = Environ.push_context uctx (Evarutil.nf_env_evar sigma' env) in
+ let c = nf c in
+ let j =
+ if Evarutil.has_undefined_evars sigma' c then
+ Evarutil.j_nf_evar sigma' (Retyping.get_judgment_of env sigma' c)
+ else
+ (* OK to call kernel which does not support evars *)
+ Arguments_renaming.rename_typing env c in
+ match redexp with
+ | None ->
+ let l = Evar.Set.union (Evd.evars_of_term j.Environ.uj_val) (Evd.evars_of_term j.Environ.uj_type) in
+ let j = { j with Environ.uj_type = Reductionops.nf_betaiota sigma' j.Environ.uj_type } in
+ Feedback.msg_notice (print_judgment env sigma' j ++
+ pr_ne_evar_set (fnl () ++ str "where" ++ fnl ()) (mt ()) sigma' l ++
+ Printer.pr_universe_ctx sigma uctx)
+ | Some r ->
+ let (sigma',r_interp) = Hook.get f_interp_redexp env sigma' r in
+ let redfun env evm c =
+ let (redfun, _) = reduction_of_red_expr env r_interp in
+ let evm = Sigma.Unsafe.of_evar_map evm in
+ let Sigma (c, _, _) = redfun.Reductionops.e_redfun env evm c in
+ c
+ in
+ Feedback.msg_notice (print_eval redfun env sigma' rc j)
+
+let vernac_declare_reduction locality s r =
+ let local = make_locality locality in
+ declare_red_expr local s (snd (Hook.get f_interp_redexp (Global.env()) Evd.empty r))
+
+ (* The same but avoiding the current goal context if any *)
+let vernac_global_check c =
+ let env = Global.env() in
+ let sigma = Evd.from_env env in
+ let c,ctx = interp_constr env sigma c in
+ let senv = Global.safe_env() in
+ let cstrs = snd (UState.context_set ctx) in
+ let senv = Safe_typing.add_constraints cstrs senv in
+ let j = Safe_typing.typing senv c in
+ let env = Safe_typing.env_of_safe_env senv in
+ Feedback.msg_notice (print_safe_judgment env sigma j)
+
+
+let get_nth_goal n =
+ let pf = get_pftreestate() in
+ let {Evd.it=gls ; sigma=sigma; } = Proof.V82.subgoals pf in
+ let gl = {Evd.it=List.nth gls (n-1) ; sigma = sigma; } in
+ gl
+
+exception NoHyp
+(* Printing "About" information of a hypothesis of the current goal.
+ We only print the type and a small statement to this comes from the
+ goal. Precondition: there must be at least one current goal. *)
+let print_about_hyp_globs ref_or_by_not glnumopt =
+ let open Context.Named.Declaration in
+ try
+ let gl,id =
+ match glnumopt,ref_or_by_not with
+ | None,AN (Ident (_loc,id)) -> (* goal number not given, catch any failure *)
+ (try get_nth_goal 1,id with _ -> raise NoHyp)
+ | Some n,AN (Ident (_loc,id)) -> (* goal number given, catch if wong *)
+ (try get_nth_goal n,id
+ with
+ Failure _ -> user_err ~hdr:"print_about_hyp_globs"
+ (str "No such goal: " ++ int n ++ str "."))
+ | _ , _ -> raise NoHyp in
+ let hyps = pf_hyps gl in
+ let decl = Context.Named.lookup id hyps in
+ let natureofid = match decl with
+ | LocalAssum _ -> "Hypothesis"
+ | LocalDef (_,bdy,_) ->"Constant (let in)" in
+ v 0 (pr_id id ++ str":" ++ pr_constr (NamedDecl.get_type decl) ++ fnl() ++ fnl()
+ ++ str natureofid ++ str " of the goal context.")
+ with (* fallback to globals *)
+ | NoHyp | Not_found -> print_about ref_or_by_not
+
+
+let vernac_print = let open Feedback in function
+ | PrintTables -> msg_notice (print_tables ())
+ | PrintFullContext-> msg_notice (print_full_context_typ ())
+ | PrintSectionContext qid -> msg_notice (print_sec_context_typ qid)
+ | PrintInspect n -> msg_notice (inspect n)
+ | PrintGrammar ent -> msg_notice (Metasyntax.pr_grammar ent)
+ | PrintLoadPath dir -> (* For compatibility ? *) msg_notice (print_loadpath dir)
+ | PrintModules -> msg_notice (print_modules ())
+ | PrintModule qid -> print_module qid
+ | PrintModuleType qid -> print_modtype qid
+ | PrintNamespace ns -> print_namespace ns
+ | PrintMLLoadPath -> msg_notice (Mltop.print_ml_path ())
+ | PrintMLModules -> msg_notice (Mltop.print_ml_modules ())
+ | PrintDebugGC -> msg_notice (Mltop.print_gc ())
+ | PrintName qid -> dump_global qid; msg_notice (print_name qid)
+ | PrintGraph -> msg_notice (Prettyp.print_graph())
+ | PrintClasses -> msg_notice (Prettyp.print_classes())
+ | PrintTypeClasses -> msg_notice (Prettyp.print_typeclasses())
+ | PrintInstances c -> msg_notice (Prettyp.print_instances (smart_global c))
+ | PrintCoercions -> msg_notice (Prettyp.print_coercions())
+ | PrintCoercionPaths (cls,clt) ->
+ msg_notice (Prettyp.print_path_between (cl_of_qualid cls) (cl_of_qualid clt))
+ | PrintCanonicalConversions -> msg_notice (Prettyp.print_canonical_projections ())
+ | PrintUniverses (b, dst) ->
+ let univ = Global.universes () in
+ let univ = if b then UGraph.sort_universes univ else univ in
+ let pr_remaining =
+ if Global.is_joined_environment () then mt ()
+ else str"There may remain asynchronous universe constraints"
+ in
+ begin match dst with
+ | None -> msg_notice (UGraph.pr_universes Universes.pr_with_global_universes univ ++ pr_remaining)
+ | Some s -> dump_universes_gen univ s
+ end
+ | PrintHint r -> msg_notice (Hints.pr_hint_ref (smart_global r))
+ | PrintHintGoal -> msg_notice (Hints.pr_applicable_hint ())
+ | PrintHintDbName s -> msg_notice (Hints.pr_hint_db_by_name s)
+ | PrintHintDb -> msg_notice (Hints.pr_searchtable ())
+ | PrintScopes ->
+ msg_notice (Notation.pr_scopes (Constrextern.without_symbols pr_lglob_constr))
+ | PrintScope s ->
+ msg_notice (Notation.pr_scope (Constrextern.without_symbols pr_lglob_constr) s)
+ | PrintVisibility s ->
+ msg_notice (Notation.pr_visibility (Constrextern.without_symbols pr_lglob_constr) s)
+ | PrintAbout (ref_or_by_not,glnumopt) ->
+ msg_notice (print_about_hyp_globs ref_or_by_not glnumopt)
+ | PrintImplicit qid ->
+ dump_global qid; msg_notice (print_impargs qid)
+ | PrintAssumptions (o,t,r) ->
+ (* Prints all the axioms and section variables used by a term *)
+ let gr = smart_global r in
+ let cstr = printable_constr_of_global gr in
+ let st = Conv_oracle.get_transp_state (Environ.oracle (Global.env())) in
+ let nassums =
+ Assumptions.assumptions st ~add_opaque:o ~add_transparent:t gr cstr in
+ msg_notice (Printer.pr_assumptionset (Global.env ()) nassums)
+ | PrintStrategy r -> print_strategy r
+
+let global_module r =
+ let (loc,qid) = qualid_of_reference r in
+ try Nametab.full_name_module qid
+ with Not_found ->
+ user_err ~loc ~hdr:"global_module"
+ (str "Module/section " ++ pr_qualid qid ++ str " not found.")
+
+let interp_search_restriction = function
+ | SearchOutside l -> (List.map global_module l, true)
+ | SearchInside l -> (List.map global_module l, false)
+
+open Search
+
+let interp_search_about_item env =
+ function
+ | SearchSubPattern pat ->
+ let _,pat = intern_constr_pattern env pat in
+ GlobSearchSubPattern pat
+ | SearchString (s,None) when Id.is_valid s ->
+ GlobSearchString s
+ | SearchString (s,sc) ->
+ try
+ let ref =
+ Notation.interp_notation_as_global_reference Loc.ghost
+ (fun _ -> true) s sc in
+ GlobSearchSubPattern (Pattern.PRef ref)
+ with UserError _ ->
+ user_err ~hdr:"interp_search_about_item"
+ (str "Unable to interp \"" ++ str s ++ str "\" either as a reference or as an identifier component")
+
+(* 05f22a5d6d5b8e3e80f1a37321708ce401834430 introduced the
+ `search_output_name_only` option to avoid excessive printing when
+ searching.
+
+ The motivation was to make search usable for IDE completion,
+ however, it is still too slow due to the non-indexed nature of the
+ underlying search mechanism.
+
+ In the future we should deprecate the option and provide a fast,
+ indexed name-searching interface.
+*)
+let search_output_name_only = ref false
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "output-name-only search";
+ optkey = ["Search";"Output";"Name";"Only"];
+ optread = (fun () -> !search_output_name_only);
+ optwrite = (:=) search_output_name_only }
+
+let vernac_search s gopt r =
+ let r = interp_search_restriction r in
+ let env,gopt =
+ match gopt with | None ->
+ (* 1st goal by default if it exists, otherwise no goal at all *)
+ (try snd (Pfedit.get_goal_context 1) , Some 1
+ with _ -> Global.env (),None)
+ (* if goal selector is given and wrong, then let exceptions be raised. *)
+ | Some g -> snd (Pfedit.get_goal_context g) , Some g
+ in
+ let get_pattern c = snd (intern_constr_pattern env c) in
+ let pr_search ref env c =
+ let pr = pr_global ref in
+ let pp = if !search_output_name_only
+ then pr
+ else begin
+ let pc = pr_lconstr_env env Evd.empty c in
+ hov 2 (pr ++ str":" ++ spc () ++ pc)
+ end
+ in Feedback.msg_notice pp
+ in
+ match s with
+ | SearchPattern c ->
+ (Search.search_pattern gopt (get_pattern c) r |> Search.prioritize_search) pr_search
+ | SearchRewrite c ->
+ (Search.search_rewrite gopt (get_pattern c) r |> Search.prioritize_search) pr_search
+ | SearchHead c ->
+ (Search.search_by_head gopt (get_pattern c) r |> Search.prioritize_search) pr_search
+ | SearchAbout sl ->
+ (Search.search_about gopt (List.map (on_snd (interp_search_about_item env)) sl) r |> Search.prioritize_search) pr_search
+
+let vernac_locate = let open Feedback in function
+ | LocateAny (AN qid) -> msg_notice (print_located_qualid qid)
+ | LocateTerm (AN qid) -> msg_notice (print_located_term qid)
+ | LocateAny (ByNotation (_, ntn, sc)) (** TODO : handle Ltac notations *)
+ | LocateTerm (ByNotation (_, ntn, sc)) ->
+ msg_notice
+ (Notation.locate_notation
+ (Constrextern.without_symbols pr_lglob_constr) ntn sc)
+ | LocateLibrary qid -> print_located_library qid
+ | LocateModule qid -> msg_notice (print_located_module qid)
+ | LocateTactic qid -> msg_notice (print_located_tactic qid)
+ | LocateFile f -> msg_notice (locate_file f)
+
+let vernac_register id r =
+ if Pfedit.refining () then
+ error "Cannot register a primitive while in proof editing mode.";
+ let t = (Constrintern.global_reference (snd id)) in
+ if not (isConst t) then
+ error "Register inline: a constant is expected";
+ let kn = destConst t in
+ match r with
+ | RegisterInline -> Global.register_inline (Univ.out_punivs kn)
+
+(********************)
+(* Proof management *)
+
+let vernac_focus gln =
+ Proof_global.simple_with_current_proof (fun _ p ->
+ match gln with
+ | None -> Proof.focus focus_command_cond () 1 p
+ | Some 0 ->
+ CErrors.error "Invalid goal number: 0. Goal numbering starts with 1."
+ | Some n ->
+ Proof.focus focus_command_cond () n p)
+
+ (* Unfocuses one step in the focus stack. *)
+let vernac_unfocus () =
+ Proof_global.simple_with_current_proof
+ (fun _ p -> Proof.unfocus command_focus p ())
+
+(* Checks that a proof is fully unfocused. Raises an error if not. *)
+let vernac_unfocused () =
+ let p = Proof_global.give_me_the_proof () in
+ if Proof.unfocused p then
+ Feedback.msg_notice (str"The proof is indeed fully unfocused.")
+ else
+ error "The proof is not fully unfocused."
+
+
+(* BeginSubproof / EndSubproof.
+ BeginSubproof (vernac_subproof) focuses on the first goal, or the goal
+ given as argument.
+ EndSubproof (vernac_end_subproof) unfocuses from a BeginSubproof, provided
+ that the proof of the goal has been completed.
+*)
+let subproof_kind = Proof.new_focus_kind ()
+let subproof_cond = Proof.done_cond subproof_kind
+
+let vernac_subproof gln =
+ Proof_global.simple_with_current_proof (fun _ p ->
+ match gln with
+ | None -> Proof.focus subproof_cond () 1 p
+ | Some n -> Proof.focus subproof_cond () n p)
+
+let vernac_end_subproof () =
+ Proof_global.simple_with_current_proof (fun _ p ->
+ Proof.unfocus subproof_kind p ())
+
+let vernac_bullet (bullet:Proof_global.Bullet.t) =
+ Proof_global.simple_with_current_proof (fun _ p ->
+ Proof_global.Bullet.put p bullet)
+
+let vernac_show = let open Feedback in function
+ | ShowGoal goalref ->
+ let info = match goalref with
+ | OpenSubgoals -> pr_open_subgoals ()
+ | NthGoal n -> pr_nth_open_subgoal n
+ | GoalId id -> pr_goal_by_id id
+ | GoalUid id -> pr_goal_by_uid id
+ in
+ msg_notice info
+ | ShowGoalImplicitly None ->
+ Constrextern.with_implicits msg_notice (pr_open_subgoals ())
+ | ShowGoalImplicitly (Some n) ->
+ Constrextern.with_implicits msg_notice (pr_nth_open_subgoal n)
+ | ShowProof -> show_proof ()
+ | ShowNode -> show_node ()
+ | ShowScript -> (* Stm.show_script () *) ()
+ | ShowExistentials -> show_top_evars ()
+ | ShowUniverses -> show_universes ()
+ | ShowTree -> show_prooftree ()
+ | ShowProofNames ->
+ msg_notice (pr_sequence pr_id (Pfedit.get_all_proof_names()))
+ | ShowIntros all -> show_intro all
+ | ShowMatch id -> show_match id
+ | ShowThesis -> show_thesis ()
+
+
+let vernac_check_guard () =
+ let pts = get_pftreestate () in
+ let pfterm = List.hd (Proof.partial_proof pts) in
+ let message =
+ try
+ let { Evd.it=gl ; sigma=sigma } = Proof.V82.top_goal pts in
+ Inductiveops.control_only_guard (Goal.V82.env sigma gl)
+ pfterm;
+ (str "The condition holds up to here")
+ with UserError(_,s) ->
+ (str ("Condition violated: ") ++s)
+ in
+ Feedback.msg_notice message
+
+exception End_of_input
+
+(* XXX: This won't properly set the proof mode, as of today, it is
+ controlled by the STM. Thus, we would need access information from
+ the classifier. The proper fix is to move it to the STM, however,
+ the way the proof mode is set there makes the task non trivial
+ without a considerable amount of refactoring.
+ *)
+let vernac_load interp fname =
+ let interp x =
+ let proof_mode = Proof_global.get_default_proof_mode_name () in
+ Proof_global.activate_proof_mode proof_mode;
+ interp x in
+ let parse_sentence = Flags.with_option Flags.we_are_parsing
+ (fun po ->
+ match Pcoq.Gram.entry_parse Pcoq.main_entry po with
+ | Some x -> x
+ | None -> raise End_of_input) in
+ let fname =
+ Envars.expand_path_macros ~warn:(fun x -> Feedback.msg_warning (str x)) fname in
+ let fname = CUnix.make_suffix fname ".v" in
+ let input =
+ let longfname = Loadpath.locate_file fname in
+ let in_chan = open_utf8_file_in longfname in
+ Pcoq.Gram.parsable ~file:longfname (Stream.of_channel in_chan) in
+ try while true do interp (snd (parse_sentence input)) done
+ with End_of_input -> ()
+
+(* "locality" is the prefix "Local" attribute, while the "local" component
+ * is the outdated/deprecated "Local" attribute of some vernacular commands
+ * still parsed as the obsolete_locality grammar entry for retrocompatibility.
+ * loc is the Loc.t of the vernacular command being interpreted. *)
+let interp ?proof ~loc locality poly c =
+ prerr_endline (fun () -> "interpreting: " ^ Pp.string_of_ppcmds (Ppvernac.pr_vernac c));
+ match c with
+ (* The below vernac are candidates for removal from the main type
+ and to be put into a new doc_command datatype: *)
+
+ | VernacLoad _ -> assert false
+
+ (* Done later in this file *)
+ | VernacFail _ -> assert false
+ | VernacTime _ -> assert false
+ | VernacRedirect _ -> assert false
+ | VernacTimeout _ -> assert false
+ | VernacStm _ -> assert false
+
+ (* The STM should handle that, but LOAD bypasses the STM... *)
+ | VernacAbortAll -> CErrors.user_err (str "AbortAll cannot be used through the Load command")
+ | VernacRestart -> CErrors.user_err (str "Restart cannot be used through the Load command")
+ | VernacUndo _ -> CErrors.user_err (str "Undo cannot be used through the Load command")
+ | VernacUndoTo _ -> CErrors.user_err (str "UndoTo cannot be used through the Load command")
+ | VernacBacktrack _ -> CErrors.user_err (str "Backtrack cannot be used through the Load command")
+
+ (* Toplevel control *)
+ | VernacToplevelControl e -> raise e
+
+ (* Resetting *)
+ | VernacResetName _ -> anomaly (str "VernacResetName not handled by Stm")
+ | VernacResetInitial -> anomaly (str "VernacResetInitial not handled by Stm")
+ | VernacBack _ -> anomaly (str "VernacBack not handled by Stm")
+ | VernacBackTo _ -> anomaly (str "VernacBackTo not handled by Stm")
+
+ (* Horrible Hack that should die. *)
+ | VernacError e -> raise e
+
+ (* This one is possible to handle here *)
+ | VernacAbort id -> CErrors.user_err (str "Abort cannot be used through the Load command")
+
+ (* Handled elsewhere *)
+ | VernacProgram _
+ | VernacPolymorphic _
+ | VernacLocal _ -> assert false
+
+ (* Syntax *)
+ | VernacSyntaxExtension (local,sl) ->
+ vernac_syntax_extension locality local sl
+ | VernacDelimiters (sc,lr) -> vernac_delimiters sc lr
+ | VernacBindScope (sc,rl) -> vernac_bind_scope sc rl
+ | VernacOpenCloseScope (local, s) -> vernac_open_close_scope locality local s
+ | VernacArgumentsScope (qid,scl) -> vernac_arguments_scope locality qid scl
+ | VernacInfix (local,mv,qid,sc) -> vernac_infix locality local mv qid sc
+ | VernacNotation (local,c,infpl,sc) ->
+ vernac_notation locality local c infpl sc
+ | VernacNotationAddFormat(n,k,v) ->
+ Metasyntax.add_notation_extra_printing_rule n k v
+
+ (* Gallina *)
+ | VernacDefinition (k,lid,d) -> vernac_definition locality poly k lid d
+ | VernacStartTheoremProof (k,l,top) -> vernac_start_proof locality poly k l top
+ | VernacEndProof e -> vernac_end_proof ?proof e
+ | VernacExactProof c -> vernac_exact_proof c
+ | VernacAssumption (stre,nl,l) -> vernac_assumption locality poly stre l nl
+ | VernacInductive (priv,finite,l) -> vernac_inductive poly priv finite l
+ | VernacFixpoint (local, l) -> vernac_fixpoint locality poly local l
+ | VernacCoFixpoint (local, l) -> vernac_cofixpoint locality poly local l
+ | VernacScheme l -> vernac_scheme l
+ | VernacCombinedScheme (id, l) -> vernac_combined_scheme id l
+ | VernacUniverse l -> vernac_universe loc poly l
+ | VernacConstraint l -> vernac_constraint loc poly l
+
+ (* Modules *)
+ | VernacDeclareModule (export,lid,bl,mtyo) ->
+ vernac_declare_module export lid bl mtyo
+ | VernacDefineModule (export,lid,bl,mtys,mexprl) ->
+ vernac_define_module export lid bl mtys mexprl
+ | VernacDeclareModuleType (lid,bl,mtys,mtyo) ->
+ vernac_declare_module_type lid bl mtys mtyo
+ | VernacInclude in_asts ->
+ vernac_include in_asts
+ (* Gallina extensions *)
+ | VernacBeginSection lid -> vernac_begin_section lid
+
+ | VernacEndSegment lid -> vernac_end_segment lid
+
+ | VernacNameSectionHypSet (lid, set) -> vernac_name_sec_hyp lid set
+
+ | VernacRequire (from, export, qidl) -> vernac_require from export qidl
+ | VernacImport (export,qidl) -> vernac_import export qidl
+ | VernacCanonical qid -> vernac_canonical qid
+ | VernacCoercion (local,r,s,t) -> vernac_coercion locality poly local r s t
+ | VernacIdentityCoercion (local,(_,id),s,t) ->
+ vernac_identity_coercion locality poly local id s t
+
+ (* Type classes *)
+ | VernacInstance (abst, sup, inst, props, info) ->
+ vernac_instance abst locality poly sup inst props info
+ | VernacContext sup -> vernac_context poly sup
+ | VernacDeclareInstances insts -> vernac_declare_instances locality insts
+ | VernacDeclareClass id -> vernac_declare_class id
+
+ (* Solving *)
+ | VernacSolveExistential (n,c) -> vernac_solve_existential n c
+
+ (* Auxiliary file and library management *)
+ | VernacAddLoadPath (isrec,s,alias) -> vernac_add_loadpath isrec s alias
+ | VernacRemoveLoadPath s -> vernac_remove_loadpath s
+ | VernacAddMLPath (isrec,s) -> vernac_add_ml_path isrec s
+ | VernacDeclareMLModule l -> vernac_declare_ml_module locality l
+ | VernacChdir s -> vernac_chdir s
+
+ (* State management *)
+ | VernacWriteState s -> vernac_write_state s
+ | VernacRestoreState s -> vernac_restore_state s
+
+ (* Commands *)
+ | VernacCreateHintDb (dbname,b) -> vernac_create_hintdb locality dbname b
+ | VernacRemoveHints (dbnames,ids) -> vernac_remove_hints locality dbnames ids
+ | VernacHints (local,dbnames,hints) ->
+ vernac_hints locality poly local dbnames hints
+ | VernacSyntacticDefinition (id,c,local,b) ->
+ vernac_syntactic_definition locality id c local b
+ | VernacDeclareImplicits (qid,l) ->
+ vernac_declare_implicits locality qid l
+ | VernacArguments (qid, args, more_implicits, nargs, flags) ->
+ vernac_arguments locality qid args more_implicits nargs flags
+ | VernacReserve bl -> vernac_reserve bl
+ | VernacGeneralizable gen -> vernac_generalizable locality gen
+ | VernacSetOpacity qidl -> vernac_set_opacity locality qidl
+ | VernacSetStrategy l -> vernac_set_strategy locality l
+ | VernacSetOption (key,v) -> vernac_set_option locality key v
+ | VernacSetAppendOption (key,v) -> vernac_set_append_option locality key v
+ | VernacUnsetOption key -> vernac_unset_option locality key
+ | VernacRemoveOption (key,v) -> vernac_remove_option key v
+ | VernacAddOption (key,v) -> vernac_add_option key v
+ | VernacMemOption (key,v) -> vernac_mem_option key v
+ | VernacPrintOption key -> vernac_print_option key
+ | VernacCheckMayEval (r,g,c) -> vernac_check_may_eval r g c
+ | VernacDeclareReduction (s,r) -> vernac_declare_reduction locality s r
+ | VernacGlobalCheck c -> vernac_global_check c
+ | VernacPrint p -> vernac_print p
+ | VernacSearch (s,g,r) -> vernac_search s g r
+ | VernacLocate l -> vernac_locate l
+ | VernacRegister (id, r) -> vernac_register id r
+ | VernacComments l -> if_verbose Feedback.msg_info (str "Comments ok\n")
+
+ (* Proof management *)
+ | VernacGoal t -> vernac_start_proof locality poly Theorem [None,([],t,None)] false
+ | VernacFocus n -> vernac_focus n
+ | VernacUnfocus -> vernac_unfocus ()
+ | VernacUnfocused -> vernac_unfocused ()
+ | VernacBullet b -> vernac_bullet b
+ | VernacSubproof n -> vernac_subproof n
+ | VernacEndSubproof -> vernac_end_subproof ()
+ | VernacShow s -> vernac_show s
+ | VernacCheckGuard -> vernac_check_guard ()
+ | VernacProof (None, None) ->
+ Aux_file.record_in_aux_at loc "VernacProof" "tac:no using:no"
+ | VernacProof (Some tac, None) ->
+ Aux_file.record_in_aux_at loc "VernacProof" "tac:yes using:no";
+ vernac_set_end_tac tac
+ | VernacProof (None, Some l) ->
+ Aux_file.record_in_aux_at loc "VernacProof" "tac:no using:yes";
+ vernac_set_used_variables l
+ | VernacProof (Some tac, Some l) ->
+ Aux_file.record_in_aux_at loc "VernacProof" "tac:yes using:yes";
+ vernac_set_end_tac tac; vernac_set_used_variables l
+ | VernacProofMode mn -> Proof_global.set_proof_mode mn
+
+ (* Extensions *)
+ | VernacExtend (opn,args) -> Vernacinterp.call ?locality (opn,args)
+
+(* Vernaculars that take a locality flag *)
+let check_vernac_supports_locality c l =
+ match l, c with
+ | None, _ -> ()
+ | Some _, (
+ VernacOpenCloseScope _
+ | VernacSyntaxExtension _ | VernacInfix _ | VernacNotation _
+ | VernacDefinition _ | VernacFixpoint _ | VernacCoFixpoint _
+ | VernacAssumption _ | VernacStartTheoremProof _
+ | VernacCoercion _ | VernacIdentityCoercion _
+ | VernacInstance _ | VernacDeclareInstances _
+ | VernacDeclareMLModule _
+ | VernacCreateHintDb _ | VernacRemoveHints _ | VernacHints _
+ | VernacSyntacticDefinition _
+ | VernacArgumentsScope _ | VernacDeclareImplicits _ | VernacArguments _
+ | VernacGeneralizable _
+ | VernacSetOpacity _ | VernacSetStrategy _
+ | VernacSetOption _ | VernacSetAppendOption _ | VernacUnsetOption _
+ | VernacDeclareReduction _
+ | VernacExtend _
+ | VernacInductive _) -> ()
+ | Some _, _ -> CErrors.error "This command does not support Locality"
+
+(* Vernaculars that take a polymorphism flag *)
+let check_vernac_supports_polymorphism c p =
+ match p, c with
+ | None, _ -> ()
+ | Some _, (
+ VernacDefinition _ | VernacFixpoint _ | VernacCoFixpoint _
+ | VernacAssumption _ | VernacInductive _
+ | VernacStartTheoremProof _
+ | VernacCoercion _ | VernacIdentityCoercion _
+ | VernacInstance _ | VernacDeclareInstances _
+ | VernacHints _ | VernacContext _
+ | VernacExtend _ | VernacUniverse _ | VernacConstraint _) -> ()
+ | Some _, _ -> CErrors.error "This command does not support Polymorphism"
+
+let enforce_polymorphism = function
+ | None -> Flags.is_universe_polymorphism ()
+ | Some b -> Flags.make_polymorphic_flag b; b
+
+(** A global default timeout, controlled by option "Set Default Timeout n".
+ Use "Unset Default Timeout" to deactivate it (or set it to 0). *)
+
+let default_timeout = ref None
+
+let _ =
+ Goptions.declare_int_option
+ { Goptions.optsync = true;
+ Goptions.optdepr = false;
+ Goptions.optname = "the default timeout";
+ Goptions.optkey = ["Default";"Timeout"];
+ Goptions.optread = (fun () -> !default_timeout);
+ Goptions.optwrite = ((:=) default_timeout) }
+
+(** When interpreting a command, the current timeout is initially
+ the default one, but may be modified locally by a Timeout command. *)
+
+let current_timeout = ref None
+
+let vernac_timeout f =
+ match !current_timeout, !default_timeout with
+ | Some n, _ | None, Some n ->
+ let f () = f (); current_timeout := None in
+ Control.timeout n f Timeout
+ | None, None -> f ()
+
+let restore_timeout () = current_timeout := None
+
+let locate_if_not_already loc (e, info) =
+ match Loc.get_loc info with
+ | None -> (e, Loc.add_loc info loc)
+ | Some l -> if Loc.is_ghost l then (e, Loc.add_loc info loc) else (e, info)
+
+exception HasNotFailed
+exception HasFailed of std_ppcmds
+
+let with_fail b f =
+ if not b then f ()
+ else begin try
+ (* If the command actually works, ignore its effects on the state.
+ * Note that error has to be printed in the right state, hence
+ * within the purified function *)
+ Future.purify
+ (fun v ->
+ try f v; raise HasNotFailed
+ with
+ | HasNotFailed as e -> raise e
+ | e ->
+ let e = CErrors.push e in
+ raise (HasFailed (CErrors.iprint
+ (ExplainErr.process_vernac_interp_error ~allow_uncaught:false ~with_header:false e))))
+ ()
+ with e when CErrors.noncritical e ->
+ let (e, _) = CErrors.push e in
+ match e with
+ | HasNotFailed ->
+ user_err ~hdr:"Fail" (str "The command has not failed!")
+ | HasFailed msg ->
+ if is_verbose () || !test_mode || !ide_slave then Feedback.msg_info
+ (str "The command has indeed failed with message:" ++ fnl () ++ msg)
+ | _ -> assert false
+ end
+
+let interp ?(verbosely=true) ?proof (loc,c) =
+ let orig_program_mode = Flags.is_program_mode () in
+ let rec aux ?locality ?polymorphism isprogcmd = function
+ | VernacProgram c when not isprogcmd -> aux ?locality ?polymorphism true c
+ | VernacProgram _ -> CErrors.error "Program mode specified twice"
+ | VernacLocal (b, c) when Option.is_empty locality ->
+ aux ~locality:b ?polymorphism isprogcmd c
+ | VernacPolymorphic (b, c) when polymorphism = None ->
+ aux ?locality ~polymorphism:b isprogcmd c
+ | VernacPolymorphic (b, c) -> CErrors.error "Polymorphism specified twice"
+ | VernacLocal _ -> CErrors.error "Locality specified twice"
+ | VernacStm (Command c) -> aux ?locality ?polymorphism isprogcmd c
+ | VernacStm (PGLast c) -> aux ?locality ?polymorphism isprogcmd c
+ | VernacStm _ -> assert false (* Done by Stm *)
+ | VernacFail v ->
+ with_fail true (fun () -> aux ?locality ?polymorphism isprogcmd v)
+ | VernacTimeout (n,v) ->
+ current_timeout := Some n;
+ aux ?locality ?polymorphism isprogcmd v
+ | VernacRedirect (s, (_,v)) ->
+ Feedback.with_output_to_file s (aux false) v
+ | VernacTime (_,v) ->
+ System.with_time !Flags.time
+ (aux ?locality ?polymorphism isprogcmd) v;
+ | VernacLoad (_,fname) -> vernac_load (aux false) fname
+ | c ->
+ check_vernac_supports_locality c locality;
+ check_vernac_supports_polymorphism c polymorphism;
+ let poly = enforce_polymorphism polymorphism in
+ Obligations.set_program_mode isprogcmd;
+ try
+ vernac_timeout begin fun () ->
+ if verbosely
+ then Flags.verbosely (interp ?proof ~loc locality poly) c
+ else Flags.silently (interp ?proof ~loc locality poly) c;
+ if orig_program_mode || not !Flags.program_mode || isprogcmd then
+ Flags.program_mode := orig_program_mode;
+ ignore (Flags.use_polymorphic_flag ())
+ end
+ with
+ | reraise when
+ (match reraise with
+ | Timeout -> true
+ | e -> CErrors.noncritical e)
+ ->
+ let e = CErrors.push reraise in
+ let e = locate_if_not_already loc e in
+ let () = restore_timeout () in
+ Flags.program_mode := orig_program_mode;
+ ignore (Flags.use_polymorphic_flag ());
+ iraise e
+ in
+ if verbosely then Flags.verbosely (aux false) c
+ else aux false c
diff --git a/vernac/vernacentries.mli b/vernac/vernacentries.mli
new file mode 100644
index 000000000..7cdc8dd06
--- /dev/null
+++ b/vernac/vernacentries.mli
@@ -0,0 +1,66 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Misctypes
+
+val dump_global : Libnames.reference or_by_notation -> unit
+
+(** Vernacular entries *)
+
+val show_prooftree : unit -> unit
+
+val show_node : unit -> unit
+
+val vernac_require :
+ Libnames.reference option -> bool option -> Libnames.reference list -> unit
+
+(** This function can be used by any command that want to observe terms
+ in the context of the current goal *)
+val get_current_context_of_args : int option -> Evd.evar_map * Environ.env
+
+(** The main interpretation function of vernacular expressions *)
+val interp :
+ ?verbosely:bool ->
+ ?proof:Proof_global.closed_proof ->
+ Loc.t * Vernacexpr.vernac_expr -> unit
+
+(** Print subgoals when the verbose flag is on.
+ Meant to be used inside vernac commands from plugins. *)
+
+val print_subgoals : unit -> unit
+val try_print_subgoals : unit -> unit
+
+(** The printing of goals via [print_subgoals] or during
+ [interp] can be controlled by the following flag.
+ Used for instance by coqide, since it has its own
+ goal-fetching mechanism. *)
+
+val enable_goal_printing : bool ref
+
+(** Should Qed try to display the proof script ?
+ True by default, but false in ProofGeneral and coqIDE *)
+
+val qed_display_script : bool ref
+
+(** Prepare a "match" template for a given inductive type.
+ For each branch of the match, we list the constructor name
+ followed by enough pattern variables.
+ [Not_found] is raised if the given string isn't the qualid of
+ a known inductive type. *)
+
+val make_cases : string -> string list list
+
+val vernac_end_proof :
+ ?proof:Proof_global.closed_proof -> Vernacexpr.proof_end -> unit
+
+val with_fail : bool -> (unit -> unit) -> unit
+
+val command_focus : unit Proof.focus_kind
+
+val interp_redexp_hook : (Environ.env -> Evd.evar_map -> Genredexpr.raw_red_expr ->
+ Evd.evar_map * Redexpr.red_expr) Hook.t
diff --git a/vernac/vernacinterp.ml b/vernac/vernacinterp.ml
new file mode 100644
index 000000000..f26ef460d
--- /dev/null
+++ b/vernac/vernacinterp.ml
@@ -0,0 +1,77 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Util
+open Pp
+open CErrors
+
+type deprecation = bool
+type vernac_command = Genarg.raw_generic_argument list -> unit -> unit
+
+(* Table of vernac entries *)
+let vernac_tab =
+ (Hashtbl.create 51 :
+ (Vernacexpr.extend_name, deprecation * vernac_command) Hashtbl.t)
+
+let vinterp_add depr s f =
+ try
+ Hashtbl.add vernac_tab s (depr, f)
+ with Failure _ ->
+ user_err ~hdr:"vinterp_add"
+ (str"Cannot add the vernac command " ++ str (fst s) ++ str" twice.")
+
+let overwriting_vinterp_add s f =
+ begin
+ try
+ let _ = Hashtbl.find vernac_tab s in Hashtbl.remove vernac_tab s
+ with Not_found -> ()
+ end;
+ Hashtbl.add vernac_tab s (false, f)
+
+let vinterp_map s =
+ try
+ Hashtbl.find vernac_tab s
+ with Failure _ | Not_found ->
+ user_err ~hdr:"Vernac Interpreter"
+ (str"Cannot find vernac command " ++ str (fst s) ++ str".")
+
+let vinterp_init () = Hashtbl.clear vernac_tab
+
+let warn_deprecated_command =
+ let open CWarnings in
+ create ~name:"deprecated-command" ~category:"deprecated"
+ (fun pr -> str "Deprecated vernacular command: " ++ pr)
+
+(* Interpretation of a vernac command *)
+
+let call ?locality (opn,converted_args) =
+ let loc = ref "Looking up command" in
+ try
+ let depr, callback = vinterp_map opn in
+ let () = if depr then
+ let rules = Egramml.get_extend_vernac_rule opn in
+ let pr_gram = function
+ | Egramml.GramTerminal s -> str s
+ | Egramml.GramNonTerminal _ -> str "_"
+ in
+ let pr = pr_sequence pr_gram rules in
+ warn_deprecated_command pr;
+ in
+ loc:= "Checking arguments";
+ let hunk = callback converted_args in
+ loc:= "Executing command";
+ Locality.LocalityFixme.set locality;
+ hunk();
+ Locality.LocalityFixme.assert_consumed()
+ with
+ | Drop -> raise Drop
+ | reraise ->
+ let reraise = CErrors.push reraise in
+ if !Flags.debug then
+ Feedback.msg_debug (str"Vernac Interpreter " ++ str !loc);
+ iraise reraise
diff --git a/vernac/vernacinterp.mli b/vernac/vernacinterp.mli
new file mode 100644
index 000000000..5149b5416
--- /dev/null
+++ b/vernac/vernacinterp.mli
@@ -0,0 +1,20 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** Interpretation of extended vernac phrases. *)
+
+type deprecation = bool
+type vernac_command = Genarg.raw_generic_argument list -> unit -> unit
+
+val vinterp_add : deprecation -> Vernacexpr.extend_name ->
+ vernac_command -> unit
+val overwriting_vinterp_add :
+ Vernacexpr.extend_name -> vernac_command -> unit
+
+val vinterp_init : unit -> unit
+val call : ?locality:bool -> Vernacexpr.extend_name * Genarg.raw_generic_argument list -> unit