aboutsummaryrefslogtreecommitdiffhomepage
path: root/vernac
diff options
context:
space:
mode:
Diffstat (limited to 'vernac')
-rw-r--r--vernac/assumptions.mli2
-rw-r--r--vernac/auto_ind_decl.ml64
-rw-r--r--vernac/auto_ind_decl.mli2
-rw-r--r--vernac/class.ml5
-rw-r--r--vernac/class.mli9
-rw-r--r--vernac/classes.ml51
-rw-r--r--vernac/classes.mli12
-rw-r--r--vernac/comAssumption.ml9
-rw-r--r--vernac/comAssumption.mli9
-rw-r--r--vernac/comDefinition.ml14
-rw-r--r--vernac/comDefinition.mli2
-rw-r--r--vernac/comFixpoint.ml19
-rw-r--r--vernac/comFixpoint.mli10
-rw-r--r--vernac/comInductive.ml98
-rw-r--r--vernac/comInductive.mli4
-rw-r--r--vernac/comProgramFixpoint.ml16
-rw-r--r--vernac/declareDef.mli10
-rw-r--r--vernac/egramcoq.ml483
-rw-r--r--vernac/egramcoq.mli19
-rw-r--r--vernac/egramml.ml84
-rw-r--r--vernac/egramml.mli33
-rw-r--r--vernac/explainErr.ml4
-rw-r--r--vernac/g_proofs.ml4131
-rw-r--r--vernac/g_vernac.ml41156
-rw-r--r--vernac/himsg.ml67
-rw-r--r--vernac/himsg.mli2
-rw-r--r--vernac/indschemes.ml2
-rw-r--r--vernac/indschemes.mli6
-rw-r--r--vernac/lemmas.ml12
-rw-r--r--vernac/lemmas.mli10
-rw-r--r--vernac/metasyntax.ml43
-rw-r--r--vernac/metasyntax.mli1
-rw-r--r--vernac/misctypes.ml75
-rw-r--r--vernac/mltop.ml9
-rw-r--r--vernac/mltop.mli3
-rw-r--r--vernac/obligations.ml72
-rw-r--r--vernac/obligations.mli11
-rw-r--r--vernac/ppvernac.ml1222
-rw-r--r--vernac/ppvernac.mli26
-rw-r--r--vernac/pvernac.ml56
-rw-r--r--vernac/pvernac.mli36
-rw-r--r--vernac/record.ml22
-rw-r--r--vernac/record.mli7
-rw-r--r--vernac/search.ml25
-rw-r--r--vernac/search.mli5
-rw-r--r--vernac/topfmt.ml35
-rw-r--r--vernac/topfmt.mli11
-rw-r--r--vernac/vernac.mllib17
-rw-r--r--vernac/vernacentries.ml71
-rw-r--r--vernac/vernacentries.mli4
-rw-r--r--vernac/vernacexpr.ml522
51 files changed, 4284 insertions, 334 deletions
diff --git a/vernac/assumptions.mli b/vernac/assumptions.mli
index 7e13f8f28..0e2b0c80e 100644
--- a/vernac/assumptions.mli
+++ b/vernac/assumptions.mli
@@ -30,4 +30,4 @@ val traverse :
{!traverse} also applies. *)
val assumptions :
?add_opaque:bool -> ?add_transparent:bool -> transparent_state ->
- global_reference -> constr -> types ContextObjectMap.t
+ GlobRef.t -> constr -> types ContextObjectMap.t
diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml
index 1a6b4dcdb..ee578669c 100644
--- a/vernac/auto_ind_decl.ml
+++ b/vernac/auto_ind_decl.ml
@@ -24,7 +24,8 @@ open Globnames
open Inductiveops
open Tactics
open Ind_tables
-open Misctypes
+open Namegen
+open Tactypes
open Proofview.Notations
module RelDecl = Context.Rel.Declaration
@@ -54,20 +55,20 @@ exception EqUnknown of string
exception UndefinedCst of string
exception InductiveWithProduct
exception InductiveWithSort
-exception ParameterWithoutEquality of global_reference
+exception ParameterWithoutEquality of GlobRef.t
exception NonSingletonProp of inductive
exception DecidabilityMutualNotSupported
exception NoDecidabilityCoInductive
-let constr_of_global g = lazy (Universes.constr_of_global g)
+let constr_of_global g = lazy (UnivGen.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 _ -> Universes.constr_of_global (Coqlib.build_bool_type()).Coqlib.andb_prop
+let andb_prop = fun _ -> UnivGen.constr_of_global (Coqlib.build_bool_type()).Coqlib.andb_prop
let andb_true_intro = fun _ ->
- Universes.constr_of_global
+ UnivGen.constr_of_global
(Coqlib.build_bool_type()).Coqlib.andb_true_intro
let tt = constr_of_global Coqlib.glob_true
@@ -76,9 +77,9 @@ let ff = constr_of_global Coqlib.glob_false
let eq = constr_of_global Coqlib.glob_eq
-let sumbool () = Universes.constr_of_global (Coqlib.build_coq_sumbool ())
+let sumbool () = UnivGen.constr_of_global (Coqlib.build_coq_sumbool ())
-let andb = fun _ -> Universes.constr_of_global (Coqlib.build_bool_type()).Coqlib.andb
+let andb = fun _ -> UnivGen.constr_of_global (Coqlib.build_bool_type()).Coqlib.andb
let induct_on c = induction false None c None None
@@ -186,10 +187,10 @@ let build_beq_scheme mode kn =
*)
let compute_A_equality rel_list nlist eqA ndx t =
let lifti = ndx in
- let sigma = Evd.empty (** FIXME *) in
let rec aux c =
- let (c,a) = Reductionops.whd_betaiota_stack Evd.empty c in
- match EConstr.kind sigma c with
+ let (c,a) = Reductionops.whd_betaiota_stack Evd.empty EConstr.(of_constr c) in
+ let (c,a) = EConstr.Unsafe.(to_constr c, List.map to_constr a) in
+ match Constr.kind c with
| Rel x -> mkRel (x-nlist+ndx), Safe_typing.empty_private_constants
| Var x ->
let eid = Id.of_string ("eq_"^(Id.to_string x)) in
@@ -198,7 +199,7 @@ let build_beq_scheme mode kn =
with Not_found -> raise (ParameterWithoutEquality (VarRef x))
in
mkVar eid, Safe_typing.empty_private_constants
- | Cast (x,_,_) -> aux (EConstr.applist (x,a))
+ | Cast (x,_,_) -> aux (Term.applist (x,a))
| App _ -> assert false
| Ind ((kn',i as ind'),u) (*FIXME: universes *) ->
if MutInd.equal kn kn' then mkRel(eqA-nlist-i+nb_ind-1), Safe_typing.empty_private_constants
@@ -213,8 +214,8 @@ let build_beq_scheme mode kn =
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 (EConstr.Unsafe.to_constr x)) a)) eqa in
+ 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))
@@ -224,10 +225,9 @@ let build_beq_scheme mode kn =
| Lambda _-> raise (EqUnknown "abstraction")
| LetIn _ -> raise (EqUnknown "let-in")
| Const (kn, u) ->
- let u = EConstr.EInstance.kind sigma u in
(match Environ.constant_opt_value_in env (kn, u) with
| None -> raise (ParameterWithoutEquality (ConstRef kn))
- | Some c -> aux (EConstr.applist (EConstr.of_constr c,a)))
+ | Some c -> aux (Term.applist (c,a)))
| Proj _ -> raise (EqUnknown "projection")
| Construct _ -> raise (EqUnknown "constructor")
| Case _ -> raise (EqUnknown "match")
@@ -271,7 +271,7 @@ let build_beq_scheme mode kn =
nparrec
(nparrec+3+2*nb_cstr_args)
(nb_cstr_args+ndx+1)
- (EConstr.of_constr cc)
+ cc
in
eff := Safe_typing.concat_private eff' !eff;
Array.set eqs ndx
@@ -319,9 +319,17 @@ let build_beq_scheme mode kn =
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 = CoFinite then
+ let fix = match mib.mind_finite with
+ | CoFinite ->
raise NoDecidabilityCoInductive;
- let fix = mkFix (((Array.make nb_ind 0),i),(names,types,cores)) in
+ | Finite ->
+ mkFix (((Array.make nb_ind 0),i),(names,types,cores))
+ | BiFinite ->
+ (** If the inductive type is not recursive, the fixpoint is not
+ used, so let's replace it with garbage *)
+ let subst = List.init nb_ind (fun _ -> mkProp) in
+ Vars.substl subst cores.(i)
+ in
create_input fix),
UState.make (Global.universes ())),
!eff
@@ -400,9 +408,9 @@ let do_replace_lb mode lb_scheme_key aavoid narg p q =
lb_type_of_p >>= fun (lb_type_of_p,eff) ->
Proofview.tclEVARMAP >>= fun sigma ->
let lb_args = Array.append (Array.append
- (Array.map (fun x -> x) v)
- (Array.map (fun x -> do_arg sigma x 1) v))
- (Array.map (fun x -> do_arg sigma x 2) v)
+ v
+ (Array.Smart.map (fun x -> do_arg sigma x 1) v))
+ (Array.Smart.map (fun x -> do_arg sigma x 2) v)
in let app = if Array.is_empty lb_args
then lb_type_of_p else mkApp (lb_type_of_p,lb_args)
in
@@ -471,9 +479,9 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt =
user_err err_msg
in let bl_args =
Array.append (Array.append
- (Array.map (fun x -> x) v)
- (Array.map (fun x -> do_arg sigma x 1) v))
- (Array.map (fun x -> do_arg sigma x 2) v )
+ v
+ (Array.Smart.map (fun x -> do_arg sigma x 1) v))
+ (Array.Smart.map (fun x -> do_arg sigma x 2) v )
in
let app = if Array.is_empty bl_args
then bl_t1 else mkApp (bl_t1,bl_args)
@@ -635,7 +643,7 @@ repeat ( apply andb_prop in z;let z1:= fresh "Z" in destruct z as [z1 z]).
| App (c,ca) -> (
match EConstr.kind sigma c with
| Ind (indeq, u) ->
- if eq_gr (IndRef indeq) Coqlib.glob_eq
+ if GlobRef.equal (IndRef indeq) Coqlib.glob_eq
then
Tacticals.New.tclTHEN
(do_replace_bl mode bl_scheme_key ind
@@ -863,7 +871,7 @@ let compute_dec_goal ind lnamesparrec nparrec =
create_input (
mkNamedProd n (mkFullInd ind (2*nparrec)) (
mkNamedProd m (mkFullInd ind (2*nparrec+1)) (
- mkApp(sumbool(),[|eqnm;mkApp (Universes.constr_of_global @@ Coqlib.build_coq_not(),[|eqnm|])|])
+ mkApp(sumbool(),[|eqnm;mkApp (UnivGen.constr_of_global @@ Coqlib.build_coq_not(),[|eqnm|])|])
)
)
)
@@ -923,7 +931,7 @@ let compute_dec_tact ind lnamesparrec nparrec =
(* left *)
Tacticals.New.tclTHENLIST [
simplest_left;
- apply (EConstr.of_constr (mkApp(blI,Array.map(fun x->mkVar x) xargs)));
+ apply (EConstr.of_constr (mkApp(blI,Array.map mkVar xargs)));
Auto.default_auto
]
;
@@ -939,7 +947,7 @@ let compute_dec_tact ind lnamesparrec nparrec =
assert_by (Name freshH3)
(EConstr.of_constr (mkApp(eq,[|bb;mkApp(eqI,[|mkVar freshm;mkVar freshm|]);tt|])))
(Tacticals.New.tclTHENLIST [
- apply (EConstr.of_constr (mkApp(lbI,Array.map (fun x->mkVar x) xargs)));
+ apply (EConstr.of_constr (mkApp(lbI,Array.map mkVar xargs)));
Auto.default_auto
]);
Equality.general_rewrite_bindings_in true
diff --git a/vernac/auto_ind_decl.mli b/vernac/auto_ind_decl.mli
index 5cc783df7..11f26c7c3 100644
--- a/vernac/auto_ind_decl.mli
+++ b/vernac/auto_ind_decl.mli
@@ -23,7 +23,7 @@ exception EqUnknown of string
exception UndefinedCst of string
exception InductiveWithProduct
exception InductiveWithSort
-exception ParameterWithoutEquality of Globnames.global_reference
+exception ParameterWithoutEquality of GlobRef.t
exception NonSingletonProp of inductive
exception DecidabilityMutualNotSupported
exception NoDecidabilityCoInductive
diff --git a/vernac/class.ml b/vernac/class.ml
index 59d933108..133726702 100644
--- a/vernac/class.ml
+++ b/vernac/class.ml
@@ -37,7 +37,7 @@ type coercion_error_kind =
| ForbiddenSourceClass of cl_typ
| NoTarget
| WrongTarget of cl_typ * cl_typ
- | NotAClass of global_reference
+ | NotAClass of GlobRef.t
exception CoercionError of coercion_error_kind
@@ -67,7 +67,7 @@ let explain_coercion_error g = function
let check_reference_arity ref =
let env = Global.env () in
let c, _ = Global.type_of_global_in_context env ref in
- if not (Reductionops.is_arity env Evd.empty (EConstr.of_constr c)) (** FIXME *) then
+ if not (Reductionops.is_arity env (Evd.from_env env) (EConstr.of_constr c)) (** FIXME *) then
raise (CoercionError (NotAClass ref))
let check_arity = function
@@ -181,6 +181,7 @@ let build_id_coercion idf_opt source poly =
let sigma, vs = match source with
| CL_CONST sp -> Evd.fresh_global env sigma (ConstRef sp)
| _ -> error_not_transparent source in
+ let vs = EConstr.Unsafe.to_constr vs in
let c = match constant_opt_value_in env (destConst vs) with
| Some c -> c
| None -> error_not_transparent source in
diff --git a/vernac/class.mli b/vernac/class.mli
index 33d31fe1f..f7e837f3b 100644
--- a/vernac/class.mli
+++ b/vernac/class.mli
@@ -10,19 +10,18 @@
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 ->
+val try_add_new_coercion_with_target : GlobRef.t -> 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 ->
+val try_add_new_coercion : GlobRef.t -> local:bool ->
Decl_kinds.polymorphic -> unit
(** [try_add_new_coercion_subclass cst s] expects that [cst] denotes a
@@ -34,7 +33,7 @@ val try_add_new_coercion_subclass : cl_typ -> local:bool ->
(** [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 ->
+val try_add_new_coercion_with_source : GlobRef.t -> local:bool ->
Decl_kinds.polymorphic -> source:cl_typ -> unit
(** [try_add_new_identity_coercion id s src tg] enriches the
@@ -47,4 +46,4 @@ 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
+val class_of_global : GlobRef.t -> cl_typ
diff --git a/vernac/classes.ml b/vernac/classes.ml
index 76d427add..8cf3895fb 100644
--- a/vernac/classes.ml
+++ b/vernac/classes.ml
@@ -41,7 +41,7 @@ let _ = Goptions.declare_bool_option {
let typeclasses_db = "typeclass_instances"
let set_typeclass_transparency c local b =
- Hints.add_hints local [typeclasses_db]
+ Hints.add_hints ~local [typeclasses_db]
(Hints.HintsTransparencyEntry ([c], b))
let _ =
@@ -50,24 +50,25 @@ let _ =
let inst' = match inst with IsConstr c -> Hints.IsConstr (EConstr.of_constr 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()) Evd.(from_env Global.(env())))
- info.hint_pattern } in
Flags.silently (fun () ->
- Hints.add_hints local [typeclasses_db]
+ 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))
+let intern_info {hint_priority;hint_pattern} =
+ let env = Global.env() in
+ let sigma = Evd.from_env env in
+ let hint_pattern = Option.map (Constrintern.intern_constr_pattern env sigma) hint_pattern in
+ {hint_priority;hint_pattern}
+
(** 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 info = intern_info info in
let instance, _ = Global.type_of_global_in_context (Global.env ()) c in
let _, r = Term.decompose_prod_assum instance in
match class_of_constr Evd.empty (EConstr.of_constr r) with
@@ -76,8 +77,8 @@ let existing_instance glob g info =
~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
+let mismatched_params env n m = Implicit_quantifiers.mismatched_ctx_inst_err env Parameters n m
+let mismatched_props env n m = Implicit_quantifiers.mismatched_ctx_inst_err env Properties n m
(* Declare everything in the parameters as implicit, and the class instance as well *)
@@ -108,6 +109,7 @@ open Pp
let instance_hook k info global imps ?hook cst =
Impargs.maybe_declare_manual_implicits false cst ~enriching:false imps;
+ let info = intern_info info in
Typeclasses.declare_instance (Some info) (not global) cst;
(match hook with Some h -> h cst | None -> ())
@@ -135,7 +137,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance)
?(tac:unit Proofview.tactic option) ?hook pri =
let env = Global.env() in
let ({CAst.loc;v=instid}, pl) = instid in
- let sigma, decl = Univdecls.interp_univ_decl_opt env pl in
+ let sigma, decl = Constrexpr_ops.interp_univ_decl_opt env pl in
let tclass, ids =
match bk with
| Decl_kinds.Implicit ->
@@ -143,7 +145,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance)
(fun avoid (clname, _) ->
match clname with
| Some cl ->
- let t = CAst.make @@ CHole (None, Misctypes.IntroAnonymous, None) in
+ let t = CAst.make @@ CHole (None, Namegen.IntroAnonymous, None) in
t, avoid
| None -> failwith ("new instance: under-applied typeclass"))
cl
@@ -196,8 +198,8 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance)
in
let (_, ty_constr) = instance_constructor (k,u) subst in
let termtype = it_mkProd_or_LetIn ty_constr (ctx' @ ctx) in
- let sigma,_ = Evarutil.nf_evars_and_universes sigma in
- Pretyping.check_evars env Evd.empty sigma termtype;
+ let sigma = Evd.minimize_universes sigma in
+ Pretyping.check_evars env (Evd.from_env env) sigma termtype;
let univs = Evd.check_univ_decl ~poly sigma decl in
let termtype = to_constr sigma termtype in
let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id
@@ -253,7 +255,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance)
k.cl_projs;
c :: props, rest'
with Not_found ->
- ((CAst.make @@ CHole (None(* Some Evar_kinds.GoalEvar *), Misctypes.IntroAnonymous, None)) :: props), rest
+ ((CAst.make @@ CHole (None(* Some Evar_kinds.GoalEvar *), Namegen.IntroAnonymous, None)) :: props), rest
else props, rest)
([], props) k.cl_props
in
@@ -289,11 +291,11 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance)
let sigma = Typeclasses.resolve_typeclasses ~filter:Typeclasses.all_evars ~fail:false env sigma in
let sigma = Evarutil.nf_evar_map_undefined sigma in
(* Beware of this step, it is required as to minimize universes. *)
- let sigma, _nf = Evarutil.nf_evar_map_universes sigma in
+ let sigma = Evd.minimize_universes sigma in
(* Check that the type is free of evars now. *)
- Pretyping.check_evars env Evd.empty sigma termtype;
+ Pretyping.check_evars env (Evd.from_env env) sigma termtype;
let termtype = to_constr sigma termtype in
- let term = Option.map (to_constr sigma) term in
+ let term = Option.map (to_constr ~abort_on_undefined_evars:false sigma) term in
if not (Evd.has_undefined sigma) && not (Option.is_empty term) then
declare_instance_constant k pri global imps ?hook id decl
poly sigma (Option.get term) termtype
@@ -302,7 +304,8 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance)
if 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];
+ Impargs.declare_manual_implicits false gr ~enriching:false [imps];
+ let pri = intern_info pri in
Typeclasses.declare_instance (Some pri) (not global) (ConstRef cst)
in
let obls, constr, typ =
@@ -365,8 +368,8 @@ let context poly l =
let sigma = Evd.from_env env in
let sigma, (_, ((env', fullctx), impls)) = interp_context_evars env sigma l in
(* Note, we must use the normalized evar from now on! *)
- let sigma,_ = Evarutil.nf_evars_and_universes sigma in
- let ce t = Pretyping.check_evars env Evd.empty sigma t in
+ let sigma = Evd.minimize_universes sigma in
+ let ce t = Pretyping.check_evars env (Evd.from_env env) sigma t in
let () = List.iter (fun decl -> Context.Rel.Declaration.iter_constr ce decl) fullctx in
let ctx =
try named_of_rel_context fullctx
@@ -424,13 +427,13 @@ let context poly l =
let decl = (Discharge, poly, Definitional) in
let nstatus = match b with
| None ->
- pi3 (ComAssumption.declare_assumption false decl (t, univs) Universes.empty_binders [] impl
- Vernacexpr.NoInline (CAst.make id))
+ pi3 (ComAssumption.declare_assumption false decl (t, univs) UnivNames.empty_binders [] impl
+ Declaremods.NoInline (CAst.make id))
| Some b ->
let decl = (Discharge, poly, Definition) in
let entry = Declare.definition_entry ~univs ~types:t b in
let hook = Lemmas.mk_hook (fun _ gr -> gr) in
- let _ = DeclareDef.declare_definition id decl entry Universes.empty_binders [] hook in
+ let _ = DeclareDef.declare_definition id decl entry UnivNames.empty_binders [] hook in
Lib.sections_are_opened () || Lib.is_modtype_strict ()
in
status && nstatus
diff --git a/vernac/classes.mli b/vernac/classes.mli
index 0342c840e..eea2a211d 100644
--- a/vernac/classes.mli
+++ b/vernac/classes.mli
@@ -22,17 +22,17 @@ val mismatched_props : env -> constr_expr list -> Context.Rel.t -> 'a
(** Instance declaration *)
-val existing_instance : bool -> reference -> Vernacexpr.hint_info_expr option -> unit
+val existing_instance : bool -> reference -> Hints.hint_info_expr option -> unit
(** globality, reference, optional priority and pattern information *)
val declare_instance_constant :
typeclass ->
- Vernacexpr.hint_info_expr -> (** priority *)
+ Hints.hint_info_expr -> (** priority *)
bool -> (** globality *)
Impargs.manual_explicitation list -> (** implicits *)
- ?hook:(Globnames.global_reference -> unit) ->
+ ?hook:(GlobRef.t -> unit) ->
Id.t -> (** name *)
- Univdecls.universe_decl ->
+ UState.universe_decl ->
bool -> (* polymorphic *)
Evd.evar_map -> (* Universes *)
Constr.t -> (** body *)
@@ -50,8 +50,8 @@ val new_instance :
(bool * constr_expr) option ->
?generalize:bool ->
?tac:unit Proofview.tactic ->
- ?hook:(Globnames.global_reference -> unit) ->
- Vernacexpr.hint_info_expr ->
+ ?hook:(GlobRef.t -> unit) ->
+ Hints.hint_info_expr ->
Id.t
(** Setting opacity *)
diff --git a/vernac/comAssumption.ml b/vernac/comAssumption.ml
index 6a590758f..a8ac52846 100644
--- a/vernac/comAssumption.ml
+++ b/vernac/comAssumption.ml
@@ -20,7 +20,6 @@ open Constrintern
open Impargs
open Decl_kinds
open Pretyping
-open Vernacexpr
open Entries
(* 2| Variable/Hypothesis/Parameter/Axiom declarations *)
@@ -66,7 +65,7 @@ match local with
| Global | Local | Discharge ->
let do_instance = should_axiom_into_instance local in
let local = DeclareDef.get_locality ident ~kind:"axiom" local in
- let inl = match nl with
+ let inl = let open Declaremods in match nl with
| NoInline -> None
| DefaultInline -> Some (Flags.get_inline_level())
| InlineAt i -> Some i
@@ -137,7 +136,7 @@ let do_assumptions kind nl l =
let open Context.Named.Declaration in
let env = Global.env () in
let udecl, l = process_assumptions_udecls kind l in
- let sigma, udecl = Univdecls.interp_univ_decl_opt env udecl in
+ let sigma, udecl = interp_univ_decl_opt env udecl in
let l =
if pi2 kind (* poly *) then
(* Separate declarations so that A B : Type puts A and B in different levels. *)
@@ -158,7 +157,7 @@ let do_assumptions kind nl l =
((sigma,env,ienv),((is_coe,idl),t,imps)))
(sigma,env,empty_internalization_env) l
in
- let sigma = solve_remaining_evars all_and_fail_flags env sigma Evd.empty in
+ let sigma = solve_remaining_evars all_and_fail_flags env sigma (Evd.from_env env) in
(* The universe constraints come from the whole telescope. *)
let sigma = Evd.minimize_universes sigma in
let nf_evar c = EConstr.to_constr sigma c in
@@ -175,7 +174,7 @@ let do_assumptions kind nl l =
let t = replace_vars subst t in
let refs, status' = declare_assumptions idl is_coe kind (t,uctx) ubinders imps nl in
let subst' = List.map2
- (fun {CAst.v=id} (c,u) -> (id, Universes.constr_of_global_univ (c,u)))
+ (fun {CAst.v=id} (c,u) -> (id, UnivGen.constr_of_global_univ (c,u)))
idl refs
in
subst'@subst, status' && status, next_uctx uctx)
diff --git a/vernac/comAssumption.mli b/vernac/comAssumption.mli
index 56e324376..56932360a 100644
--- a/vernac/comAssumption.mli
+++ b/vernac/comAssumption.mli
@@ -11,7 +11,6 @@
open Names
open Constr
open Entries
-open Globnames
open Vernacexpr
open Constrexpr
open Decl_kinds
@@ -19,7 +18,7 @@ open Decl_kinds
(** {6 Parameters/Assumptions} *)
val do_assumptions : locality * polymorphic * assumption_object_kind ->
- Vernacexpr.inline -> (ident_decl list * constr_expr) with_coercion list -> bool
+ Declaremods.inline -> (ident_decl list * constr_expr) with_coercion list -> bool
(************************************************************************)
(** Internal API *)
@@ -31,6 +30,6 @@ val do_assumptions : locality * polymorphic * assumption_object_kind ->
nor in a module type and meant to be instantiated. *)
val declare_assumption : coercion_flag -> assumption_kind ->
types in_constant_universes_entry ->
- Universes.universe_binders -> Impargs.manual_implicits ->
- bool (** implicit *) -> Vernacexpr.inline -> variable CAst.t ->
- global_reference * Univ.Instance.t * bool
+ UnivNames.universe_binders -> Impargs.manual_implicits ->
+ bool (** implicit *) -> Declaremods.inline -> variable CAst.t ->
+ GlobRef.t * Univ.Instance.t * bool
diff --git a/vernac/comDefinition.ml b/vernac/comDefinition.ml
index b18a60a1f..f55c852c0 100644
--- a/vernac/comDefinition.ml
+++ b/vernac/comDefinition.ml
@@ -65,7 +65,7 @@ let interp_definition pl bl poly red_option c ctypopt =
let open EConstr in
let env = Global.env() in
(* Explicitly bound universes and constraints *)
- let evd, decl = Univdecls.interp_univ_decl_opt env pl in
+ let evd, decl = Constrexpr_ops.interp_univ_decl_opt env pl in
(* Build the parameters *)
let evd, (impls, ((env_bl, ctx), imps1)) = interp_context_evars env evd bl in
(* Build the type *)
@@ -88,9 +88,9 @@ let interp_definition pl bl poly red_option c ctypopt =
let evd = Evd.minimize_universes evd in
(* Substitute evars and universes, and add parameters.
Note: in program mode some evars may remain. *)
- let ctx = List.map (EConstr.to_rel_decl evd) ctx in
- let c = Term.it_mkLambda_or_LetIn (EConstr.to_constr evd c) ctx in
- let tyopt = Option.map (fun ty -> Term.it_mkProd_or_LetIn (EConstr.to_constr evd ty) ctx) tyopt in
+ let ctx = List.map Termops.(map_rel_decl (to_constr ~abort_on_undefined_evars:false evd)) ctx in
+ let c = Term.it_mkLambda_or_LetIn (EConstr.to_constr ~abort_on_undefined_evars:false evd c) ctx in
+ let tyopt = Option.map (fun ty -> Term.it_mkProd_or_LetIn (EConstr.to_constr ~abort_on_undefined_evars:false evd ty) ctx) tyopt in
(* Keep only useful universes. *)
let uvars_fold uvars c =
Univ.LSet.union uvars (universes_of_constr env evd (of_constr c))
@@ -104,7 +104,9 @@ let interp_definition pl bl poly red_option c ctypopt =
(red_constant_entry (Context.Rel.length ctx) ce evd red_option, evd, decl, imps)
let check_definition (ce, evd, _, imps) =
- check_evars_are_solved (Global.env ()) evd Evd.empty;
+ let env = Global.env () in
+ let empty_sigma = Evd.from_env env in
+ check_evars_are_solved env evd empty_sigma;
ce
let do_definition ~program_mode ident k univdecl bl red_option c ctypopt hook =
@@ -118,7 +120,7 @@ let do_definition ~program_mode ident k univdecl bl red_option c ctypopt hook =
assert(Univ.ContextSet.is_empty ctx);
let typ = match ce.const_entry_type with
| Some t -> t
- | None -> EConstr.to_constr evd (Retyping.get_type_of env evd (EConstr.of_constr c))
+ | None -> EConstr.to_constr ~abort_on_undefined_evars:false evd (Retyping.get_type_of env evd (EConstr.of_constr c))
in
Obligations.check_evars env evd;
let obls, _, c, cty =
diff --git a/vernac/comDefinition.mli b/vernac/comDefinition.mli
index 6f81c4575..7f1c902c0 100644
--- a/vernac/comDefinition.mli
+++ b/vernac/comDefinition.mli
@@ -29,4 +29,4 @@ val do_definition : program_mode:bool ->
val interp_definition :
universe_decl_expr option -> local_binder_expr list -> polymorphic -> red_expr option -> constr_expr ->
constr_expr option -> Safe_typing.private_constants definition_entry * Evd.evar_map *
- Univdecls.universe_decl * Impargs.manual_implicits
+ UState.universe_decl * Impargs.manual_implicits
diff --git a/vernac/comFixpoint.ml b/vernac/comFixpoint.ml
index a794c2db0..1d1cc62de 100644
--- a/vernac/comFixpoint.ml
+++ b/vernac/comFixpoint.ml
@@ -13,7 +13,6 @@ open Decl_kinds
open Pretyping
open Evarutil
open Evarconv
-open Misctypes
module RelDecl = Context.Rel.Declaration
@@ -173,11 +172,12 @@ let interp_recursive ~program_mode ~cofix fixl notations =
| None , acc -> acc
| x , None -> x
| Some ls , Some us ->
- let lsu = ls.univdecl_instance and usu = us.univdecl_instance in
+ let open UState in
+ let lsu = ls.univdecl_instance and usu = us.univdecl_instance in
if not (CList.for_all2eq (fun x y -> Id.equal x.CAst.v y.CAst.v) lsu usu) then
user_err Pp.(str "(co)-recursive definitions should all have the same universe binders");
Some us) fixl None in
- let sigma, decl = Univdecls.interp_univ_decl_opt env all_universes in
+ let sigma, decl = interp_univ_decl_opt env all_universes in
let sigma, (fixctxs, fiximppairs, fixannots) =
on_snd List.split3 @@
List.fold_left_map (fun sigma -> interp_fix_context env sigma ~cofix) sigma fixl in
@@ -199,9 +199,7 @@ let interp_recursive ~program_mode ~cofix fixl notations =
try
let sigma, h_term = fix_proto sigma in
let app = mkApp (h_term, [|sort; t|]) in
- let _evd = ref sigma in
- let res = Typing.e_solve_evars env _evd app in
- !_evd, res
+ Typing.solve_evars env sigma app
with e when CErrors.noncritical e -> sigma, t
in
sigma, LocalAssum (id,fixprot) :: env'
@@ -224,8 +222,9 @@ let interp_recursive ~program_mode ~cofix fixl notations =
(* Instantiate evars and check all are resolved *)
let sigma = solve_unif_constraints_with_heuristics env_rec sigma in
- let sigma, _ = nf_evars_and_universes sigma in
- let fixdefs = List.map (fun c -> Option.map EConstr.(to_constr sigma) c) fixdefs in
+ let sigma = Evd.minimize_universes sigma in
+ (* XXX: We still have evars here in Program *)
+ let fixdefs = List.map (fun c -> Option.map EConstr.(to_constr ~abort_on_undefined_evars:false sigma) c) fixdefs in
let fixtypes = List.map EConstr.(to_constr sigma) fixtypes in
let fixctxs = List.map (fun (_,ctx) -> ctx) fixctxs in
@@ -233,7 +232,7 @@ let interp_recursive ~program_mode ~cofix fixl notations =
(env,rec_sign,decl,sigma), (fixnames,fixdefs,fixtypes), List.combine3 fixctxs fiximps fixannots
let check_recursive isfix env evd (fixnames,fixdefs,_) =
- check_evars_are_solved env evd Evd.empty;
+ check_evars_are_solved env evd (Evd.from_env env);
if List.for_all Option.has_some fixdefs then begin
let fixdefs = List.map Option.get fixdefs in
check_mutuality env evd isfix (List.combine fixnames fixdefs)
@@ -254,7 +253,7 @@ let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ind
Some (List.map (Option.cata (EConstr.of_constr %> Tactics.exact_no_check) Tacticals.New.tclIDTAC)
fixdefs) in
let evd = Evd.from_ctx ctx in
- Lemmas.start_proof_with_initialization (Global,poly,DefinitionBody Fixpoint)
+ Lemmas.start_proof_with_initialization (local,poly,DefinitionBody Fixpoint)
evd pl (Some(false,indexes,init_tac)) thms None (Lemmas.mk_hook (fun _ _ -> ()))
else begin
(* We shortcut the proof process *)
diff --git a/vernac/comFixpoint.mli b/vernac/comFixpoint.mli
index 36c2993af..f51bfbad5 100644
--- a/vernac/comFixpoint.mli
+++ b/vernac/comFixpoint.mli
@@ -33,7 +33,7 @@ val do_cofixpoint :
type structured_fixpoint_expr = {
fix_name : Id.t;
fix_univs : Constrexpr.universe_decl_expr option;
- fix_annot : Misctypes.lident option;
+ fix_annot : lident option;
fix_binders : local_binder_expr list;
fix_body : constr_expr option;
fix_type : constr_expr
@@ -49,7 +49,7 @@ val interp_recursive :
structured_fixpoint_expr list -> decl_notation list ->
(* env / signature / univs / evar_map *)
- (Environ.env * EConstr.named_context * Univdecls.universe_decl * Evd.evar_map) *
+ (Environ.env * EConstr.named_context * UState.universe_decl * Evd.evar_map) *
(* names / defs / types *)
(Id.t list * Constr.constr option list * Constr.types list) *
(* ctx per mutual def / implicits / struct annotations *)
@@ -74,19 +74,19 @@ type recursive_preentry =
val interp_fixpoint :
cofix:bool ->
structured_fixpoint_expr list -> decl_notation list ->
- recursive_preentry * Univdecls.universe_decl * UState.t *
+ recursive_preentry * UState.universe_decl * UState.t *
(EConstr.rel_context * Impargs.manual_implicits * int option) list
(** Registering fixpoints and cofixpoints in the environment *)
(** [Not used so far] *)
val declare_fixpoint :
locality -> polymorphic ->
- recursive_preentry * Univdecls.universe_decl * UState.t *
+ recursive_preentry * UState.universe_decl * UState.t *
(Context.Rel.t * Impargs.manual_implicits * int option) list ->
Proof_global.lemma_possible_guards -> decl_notation list -> unit
val declare_cofixpoint : locality -> polymorphic ->
- recursive_preentry * Univdecls.universe_decl * UState.t *
+ recursive_preentry * UState.universe_decl * UState.t *
(Context.Rel.t * Impargs.manual_implicits * int option) list ->
decl_notation list -> unit
diff --git a/vernac/comInductive.ml b/vernac/comInductive.ml
index db2f16525..b93e8d9ac 100644
--- a/vernac/comInductive.ml
+++ b/vernac/comInductive.ml
@@ -27,9 +27,7 @@ open Impargs
open Reductionops
open Indtypes
open Pretyping
-open Evarutil
open Indschemes
-open Misctypes
open Context.Rel.Declaration
open Entries
@@ -159,7 +157,7 @@ let sign_level env evd sign =
| LocalDef _ -> lev, push_rel d env
| LocalAssum _ ->
let s = destSort (Reduction.whd_all env
- (EConstr.Unsafe.to_constr (nf_evar evd (Retyping.get_type_of env evd (EConstr.of_constr (RelDecl.get_type d))))))
+ (EConstr.to_constr evd (Retyping.get_type_of env evd (EConstr.of_constr (RelDecl.get_type d)))))
in
let u = univ_of_sort s in
(Univ.sup u lev, push_rel d env))
@@ -178,6 +176,72 @@ let is_flexible_sort evd u =
| Some l -> Evd.is_flexible_level evd l
| None -> false
+(**********************************************************************)
+(* Tools for template polymorphic inductive types *)
+
+(* Miscellaneous functions to remove or test local univ assumed to
+ occur only in the le constraints *)
+
+(*
+ Solve a system of universe constraint of the form
+
+ u_s11, ..., u_s1p1, w1 <= u1
+ ...
+ u_sn1, ..., u_snpn, wn <= un
+
+where
+
+ - the ui (1 <= i <= n) are universe variables,
+ - the sjk select subsets of the ui for each equations,
+ - the wi are arbitrary complex universes that do not mention the ui.
+*)
+
+let is_direct_sort_constraint s v = match s with
+ | Some u -> Univ.univ_level_mem u v
+ | None -> false
+
+let solve_constraints_system levels level_bounds =
+ let open Univ in
+ let levels =
+ Array.mapi (fun i o ->
+ match o with
+ | Some u ->
+ (match Universe.level u with
+ | Some u -> Some u
+ | _ -> level_bounds.(i) <- Universe.sup level_bounds.(i) u; None)
+ | None -> None)
+ levels in
+ let v = Array.copy level_bounds in
+ let nind = Array.length v in
+ let clos = Array.map (fun _ -> Int.Set.empty) levels in
+ (* First compute the transitive closure of the levels dependencies *)
+ for i=0 to nind-1 do
+ for j=0 to nind-1 do
+ if not (Int.equal i j) && is_direct_sort_constraint levels.(j) v.(i) then
+ clos.(i) <- Int.Set.add j clos.(i);
+ done;
+ done;
+ let rec closure () =
+ let continue = ref false in
+ Array.iteri (fun i deps ->
+ let deps' =
+ Int.Set.fold (fun j acc -> Int.Set.union acc clos.(j)) deps deps
+ in
+ if Int.Set.equal deps deps' then ()
+ else (clos.(i) <- deps'; continue := true))
+ clos;
+ if !continue then closure ()
+ else ()
+ in
+ closure ();
+ for i=0 to nind-1 do
+ for j=0 to nind-1 do
+ if not (Int.equal i j) && Int.Set.mem j clos.(i) then
+ (v.(i) <- Universe.sup v.(i) level_bounds.(j));
+ done;
+ done;
+ v
+
let inductive_levels env evd 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)) ->
@@ -206,8 +270,8 @@ let inductive_levels env evd poly arities inds =
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)
+ let levels' = solve_constraints_system (Array.of_list levels)
+ (Array.of_list cstrs_levels)
in
let evd, arities =
CList.fold_left3 (fun (evd, arities) cu (arity,(ctx,du)) len ->
@@ -261,14 +325,14 @@ let check_param = function
| CLocalAssum (nas, Default _, _) -> List.iter check_named nas
| CLocalAssum (nas, Generalized _, _) -> ()
| CLocalPattern {CAst.loc} ->
- Loc.raise ?loc (Stream.Error "pattern with quote not allowed here.")
+ Loc.raise ?loc (Stream.Error "pattern with quote not allowed here")
let interp_mutual_inductive (paramsl,indl) notations cum 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 sigma, decl = Univdecls.interp_univ_decl_opt env0 pl in
+ let sigma, decl = interp_univ_decl_opt env0 pl in
let sigma, (impls, ((env_params, ctx_params), userimpls)) =
interp_context_evars env0 sigma paramsl
in
@@ -302,22 +366,24 @@ let interp_mutual_inductive (paramsl,indl) notations cum poly prv finite =
() in
(* Try further to solve evars, and instantiate them *)
- let sigma = solve_remaining_evars all_and_fail_flags env_params sigma Evd.empty in
+ let sigma = solve_remaining_evars all_and_fail_flags env_params sigma (Evd.from_env env_params) in
(* Compute renewed arities *)
- let sigma, nf = nf_evars_and_universes sigma in
+ let sigma = Evd.minimize_universes sigma in
+ let nf = Evarutil.nf_evars_universes sigma in
let constructors = List.map (fun (idl,cl,impsl) -> (idl,List.map nf cl,impsl)) constructors in
let arities = List.map EConstr.(to_constr sigma) arities in
let sigma = List.fold_left2 (fun sigma ty poly -> make_conclusion_flexible sigma ty poly) sigma arities aritypoly in
let sigma, arities = inductive_levels env_ar_params sigma poly arities constructors in
- let sigma, nf' = nf_evars_and_universes sigma 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 sigma = Evd.minimize_universes sigma in
+ let nf = Evarutil.nf_evars_universes sigma 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 = List.map Termops.(map_rel_decl (EConstr.to_constr sigma)) ctx_params in
let uctx = Evd.check_univ_decl ~poly sigma decl in
- List.iter (fun c -> check_evars env_params Evd.empty sigma (EConstr.of_constr c)) arities;
- Context.Rel.iter (fun c -> check_evars env0 Evd.empty sigma (EConstr.of_constr c)) ctx_params;
+ List.iter (fun c -> check_evars env_params (Evd.from_env env_params) sigma (EConstr.of_constr c)) arities;
+ Context.Rel.iter (fun c -> check_evars env0 (Evd.from_env env0) sigma (EConstr.of_constr c)) ctx_params;
List.iter (fun (_,ctyps,_) ->
- List.iter (fun c -> check_evars env_ar_params Evd.empty sigma (EConstr.of_constr c)) ctyps)
+ List.iter (fun c -> check_evars env_ar_params (Evd.from_env env_ar_params) sigma (EConstr.of_constr c)) ctyps)
constructors;
(* Build the inductive entries *)
@@ -378,7 +444,7 @@ let extract_params indl =
let extract_inductive indl =
List.map (fun (({CAst.v=indname},pl),_,ar,lc) -> {
ind_name = indname; ind_univs = pl;
- ind_arity = Option.cata (fun x -> x) (CAst.make @@ CSort (GType [])) ar;
+ ind_arity = Option.cata (fun x -> x) (CAst.make @@ CSort (Glob_term.GType [])) ar;
ind_lc = List.map (fun (_,({CAst.v=id},t)) -> (id,t)) lc
}) indl
diff --git a/vernac/comInductive.mli b/vernac/comInductive.mli
index 833935724..7ae8e8f71 100644
--- a/vernac/comInductive.mli
+++ b/vernac/comInductive.mli
@@ -37,7 +37,7 @@ type one_inductive_impls =
Impargs.manual_implicits list (** for constrs *)
val declare_mutual_inductive_with_eliminations :
- mutual_inductive_entry -> Universes.universe_binders -> one_inductive_impls list ->
+ mutual_inductive_entry -> UnivNames.universe_binders -> one_inductive_impls list ->
MutInd.t
(** Exported for Funind *)
@@ -64,4 +64,4 @@ val extract_mutual_inductive_declaration_components :
val interp_mutual_inductive :
structured_inductive_expr -> decl_notation list -> cumulative_inductive_flag ->
polymorphic -> private_flag -> Declarations.recursivity_kind ->
- mutual_inductive_entry * Universes.universe_binders * one_inductive_impls list
+ mutual_inductive_entry * UnivNames.universe_binders * one_inductive_impls list
diff --git a/vernac/comProgramFixpoint.ml b/vernac/comProgramFixpoint.ml
index b95741ca4..a6d7fccf3 100644
--- a/vernac/comProgramFixpoint.ml
+++ b/vernac/comProgramFixpoint.ml
@@ -91,7 +91,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
let lift_rel_context n l = Termops.map_rel_context_with_binders (liftn n) l in
Coqlib.check_required_library ["Coq";"Program";"Wf"];
let env = Global.env() in
- let sigma, decl = Univdecls.interp_univ_decl_opt env pl in
+ let sigma, decl = Constrexpr_ops.interp_univ_decl_opt env pl in
let sigma, (_, ((env', binders_rel), impls)) = interp_context_evars env sigma bl in
let len = List.length binders_rel in
let top_env = push_rel_context binders_rel env in
@@ -190,9 +190,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
~src:(Loc.tag @@ Evar_kinds.QuestionMark (Evar_kinds.Define false,Anonymous)) wf_proof in
sigma, mkApp (h_a_term, [| argtyp ; wf_rel ; h_e_term; prop |])
in
- let _evd = ref sigma in
- let def = Typing.e_solve_evars env _evd def in
- let sigma = !_evd in
+ let sigma, def = Typing.solve_evars env sigma def in
let sigma = Evarutil.nf_evar_map sigma in
let def = mkApp (def, [|intern_body_lam|]) in
let binders_rel = nf_evar_context sigma binders_rel in
@@ -214,7 +212,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
(** FIXME: include locality *)
let c = Declare.declare_constant recname (DefinitionEntry ce, IsDefinition Definition) in
let gr = ConstRef c in
- let () = Universes.register_universe_binders gr (Evd.universe_binders sigma) in
+ let () = UnivNames.register_universe_binders gr (Evd.universe_binders sigma) in
if Impargs.is_implicit_args () || not (List.is_empty impls) then
Impargs.declare_manual_implicits false gr [impls]
in
@@ -229,7 +227,8 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
in
(* XXX: Capturing sigma here... bad bad *)
let hook = Lemmas.mk_hook (hook sigma) in
- let fullcoqc = EConstr.to_constr sigma def in
+ (* XXX: Grounding non-ground terms here... bad bad *)
+ let fullcoqc = EConstr.to_constr ~abort_on_undefined_evars:false sigma def in
let fullctyp = EConstr.to_constr sigma typ in
Obligations.check_evars env sigma;
let evars, _, evars_def, evars_typ =
@@ -261,9 +260,10 @@ let do_program_recursive local poly fixkind fixl ntns =
let collect_evars id def typ imps =
(* Generalize by the recursive prototypes *)
let def =
- EConstr.to_constr evd (Termops.it_mkNamedLambda_or_LetIn (EConstr.of_constr def) rec_sign)
+ EConstr.to_constr ~abort_on_undefined_evars:false evd (Termops.it_mkNamedLambda_or_LetIn (EConstr.of_constr def) rec_sign)
and typ =
- EConstr.to_constr evd (Termops.it_mkNamedProd_or_LetIn (EConstr.of_constr typ) rec_sign)
+ (* Worrying... *)
+ EConstr.to_constr ~abort_on_undefined_evars:false evd (Termops.it_mkNamedProd_or_LetIn (EConstr.of_constr typ) rec_sign)
in
let evm = collect_evars_of_term evd def typ in
let evars, _, def, typ =
diff --git a/vernac/declareDef.mli b/vernac/declareDef.mli
index 010874e23..c5e704a5e 100644
--- a/vernac/declareDef.mli
+++ b/vernac/declareDef.mli
@@ -8,17 +8,17 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
-open Decl_kinds
open Names
+open Decl_kinds
val get_locality : Id.t -> kind:string -> Decl_kinds.locality -> bool
val declare_definition : Id.t -> definition_kind ->
- Safe_typing.private_constants Entries.definition_entry -> Universes.universe_binders -> Impargs.manual_implicits ->
- Globnames.global_reference Lemmas.declaration_hook -> Globnames.global_reference
+ Safe_typing.private_constants Entries.definition_entry -> UnivNames.universe_binders -> Impargs.manual_implicits ->
+ GlobRef.t Lemmas.declaration_hook -> GlobRef.t
val declare_fix : ?opaque:bool -> definition_kind ->
- Universes.universe_binders -> Entries.constant_universes_entry ->
+ UnivNames.universe_binders -> Entries.constant_universes_entry ->
Id.t -> Safe_typing.private_constants Entries.proof_output ->
Constr.types -> Impargs.manual_implicits ->
- Globnames.global_reference
+ GlobRef.t
diff --git a/vernac/egramcoq.ml b/vernac/egramcoq.ml
new file mode 100644
index 000000000..434e836d8
--- /dev/null
+++ b/vernac/egramcoq.ml
@@ -0,0 +1,483 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+open Util
+open CErrors
+open Names
+open Libnames
+open Constrexpr
+open Extend
+open Notation_gram
+open Pcoq
+
+(**********************************************************************)
+(* This determines (depending on the associativity of the current
+ level and on the expected associativity) if a reference to constr_n is
+ a reference to the current level (to be translated into "SELF" on the
+ left border and into "constr LEVEL n" elsewhere), to the level below
+ (to be translated into "NEXT") or to an below wrt associativity (to be
+ translated in camlp5 into "constr" without level) or to another level
+ (to be translated into "constr LEVEL n")
+
+ The boolean is true if the entry was existing _and_ empty; this to
+ circumvent a weakness of camlp5 whose undo mechanism is not the
+ converse of the extension mechanism *)
+
+let constr_level = string_of_int
+
+let default_levels =
+ [200,Extend.RightA,false;
+ 100,Extend.RightA,false;
+ 99,Extend.RightA,true;
+ 90,Extend.RightA,true;
+ 10,Extend.LeftA,false;
+ 9,Extend.RightA,false;
+ 8,Extend.RightA,true;
+ 1,Extend.LeftA,false;
+ 0,Extend.RightA,false]
+
+let default_pattern_levels =
+ [200,Extend.RightA,true;
+ 100,Extend.RightA,false;
+ 99,Extend.RightA,true;
+ 90,Extend.RightA,true;
+ 10,Extend.LeftA,false;
+ 1,Extend.LeftA,false;
+ 0,Extend.RightA,false]
+
+let default_constr_levels = (default_levels, default_pattern_levels)
+
+(* At a same level, LeftA takes precedence over RightA and NoneA *)
+(* In case, several associativity exists for a level, we make two levels, *)
+(* first LeftA, then RightA and NoneA together *)
+
+let admissible_assoc = function
+ | Extend.LeftA, Some (Extend.RightA | Extend.NonA) -> false
+ | Extend.RightA, Some Extend.LeftA -> false
+ | _ -> true
+
+let create_assoc = function
+ | None -> Extend.RightA
+ | Some a -> a
+
+let error_level_assoc p current expected =
+ let open Pp in
+ let pr_assoc = function
+ | Extend.LeftA -> str "left"
+ | Extend.RightA -> str "right"
+ | Extend.NonA -> str "non" in
+ user_err
+ (str "Level " ++ int p ++ str " is already declared " ++
+ pr_assoc current ++ str " associative while it is now expected to be " ++
+ pr_assoc expected ++ str " associative.")
+
+let create_pos = function
+ | None -> Extend.First
+ | Some lev -> Extend.After (constr_level lev)
+
+let find_position_gen current ensure assoc lev =
+ match lev with
+ | None ->
+ current, (None, None, None, None)
+ | Some n ->
+ let after = ref None in
+ let init = ref None in
+ let rec add_level q = function
+ | (p,_,_ as pa)::l when p > n -> pa :: add_level (Some p) l
+ | (p,a,reinit)::l when Int.equal p n ->
+ if reinit then
+ let a' = create_assoc assoc in
+ (init := Some (a',create_pos q); (p,a',false)::l)
+ else if admissible_assoc (a,assoc) then
+ raise Exit
+ else
+ error_level_assoc p a (Option.get assoc)
+ | l -> after := q; (n,create_assoc assoc,ensure)::l
+ in
+ try
+ let updated = add_level None current in
+ let assoc = create_assoc assoc in
+ begin match !init with
+ | None ->
+ (* Create the entry *)
+ updated, (Some (create_pos !after), Some assoc, Some (constr_level n), None)
+ | _ ->
+ (* The reinit flag has been updated *)
+ updated, (Some (Extend.Level (constr_level n)), None, None, !init)
+ end
+ with
+ (* Nothing has changed *)
+ Exit ->
+ (* Just inherit the existing associativity and name (None) *)
+ current, (Some (Extend.Level (constr_level n)), None, None, None)
+
+let rec list_mem_assoc_triple x = function
+ | [] -> false
+ | (a,b,c) :: l -> Int.equal a x || list_mem_assoc_triple x l
+
+let register_empty_levels accu forpat levels =
+ let rec filter accu = function
+ | [] -> ([], accu)
+ | n :: rem ->
+ let rem, accu = filter accu rem in
+ let (clev, plev) = accu in
+ let levels = if forpat then plev else clev in
+ if not (list_mem_assoc_triple n levels) then
+ let nlev, ans = find_position_gen levels true None (Some n) in
+ let nlev = if forpat then (clev, nlev) else (nlev, plev) in
+ ans :: rem, nlev
+ else rem, accu
+ in
+ filter accu levels
+
+let find_position accu forpat assoc level =
+ let (clev, plev) = accu in
+ let levels = if forpat then plev else clev in
+ let nlev, ans = find_position_gen levels false assoc level in
+ let nlev = if forpat then (clev, nlev) else (nlev, plev) in
+ (ans, nlev)
+
+(**************************************************************************)
+(*
+ * --- Note on the mapping of grammar productions to camlp5 actions ---
+ *
+ * Translation of environments: a production
+ * [ nt1(x1) ... nti(xi) ] -> act(x1..xi)
+ * is written (with camlp5 conventions):
+ * (fun vi -> .... (fun v1 -> act(v1 .. vi) )..)
+ * where v1..vi are the values generated by non-terminals nt1..nti.
+ * Since the actions are executed by substituting an environment,
+ * the make_*_action family build the following closure:
+ *
+ * ((fun env ->
+ * (fun vi ->
+ * (fun env -> ...
+ *
+ * (fun v1 ->
+ * (fun env -> gram_action .. env act)
+ * ((x1,v1)::env))
+ * ...)
+ * ((xi,vi)::env)))
+ * [])
+ *)
+
+(**********************************************************************)
+(** Declare Notations grammar rules *)
+
+(**********************************************************************)
+(* Binding constr entry keys to entries *)
+
+(* Camlp5 levels do not treat NonA: use RightA with a NEXT on the left *)
+let camlp5_assoc = function
+ | Some NonA | Some RightA -> RightA
+ | None | Some LeftA -> LeftA
+
+let assoc_eq al ar = match al, ar with
+| NonA, NonA
+| RightA, RightA
+| LeftA, LeftA -> true
+| _, _ -> false
+
+(* [adjust_level assoc from prod] where [assoc] and [from] are the name
+ and associativity of the level where to add the rule; the meaning of
+ the result is
+
+ None = SELF
+ Some None = NEXT
+ Some (Some (n,cur)) = constr LEVEL n
+ s.t. if [cur] is set then [n] is the same as the [from] level *)
+let adjust_level assoc from = function
+(* Associativity is None means force the level *)
+ | (NumLevel n,BorderProd (_,None)) -> Some (Some (n,true))
+(* Compute production name on the right side *)
+ (* If NonA or LeftA on the right-hand side, set to NEXT *)
+ | (NumLevel n,BorderProd (Right,Some (NonA|LeftA))) ->
+ Some None
+ (* If RightA on the right-hand side, set to the explicit (current) level *)
+ | (NumLevel n,BorderProd (Right,Some RightA)) ->
+ Some (Some (n,true))
+(* Compute production name on the left side *)
+ (* If NonA on the left-hand side, adopt the current assoc ?? *)
+ | (NumLevel n,BorderProd (Left,Some NonA)) -> None
+ (* If the expected assoc is the current one, set to SELF *)
+ | (NumLevel n,BorderProd (Left,Some a)) when assoc_eq a (camlp5_assoc assoc) ->
+ None
+ (* Otherwise, force the level, n or n-1, according to expected assoc *)
+ | (NumLevel n,BorderProd (Left,Some a)) ->
+ begin match a with
+ | LeftA -> Some (Some (n, true))
+ | _ -> Some None
+ end
+ (* None means NEXT *)
+ | (NextLevel,_) -> Some None
+(* Compute production name elsewhere *)
+ | (NumLevel n,InternalProd) ->
+ if from = n + 1 then Some None else Some (Some (n, Int.equal n from))
+
+type _ target =
+| ForConstr : constr_expr target
+| ForPattern : cases_pattern_expr target
+
+type prod_info = production_level * production_position
+
+type (_, _) entry =
+| TTName : ('self, lname) entry
+| TTReference : ('self, reference) entry
+| TTBigint : ('self, Constrexpr.raw_natural_number) entry
+| TTConstr : prod_info * 'r target -> ('r, 'r) entry
+| TTConstrList : prod_info * Tok.t list * 'r target -> ('r, 'r list) entry
+| TTPattern : int -> ('self, cases_pattern_expr) entry
+| TTOpenBinderList : ('self, local_binder_expr list) entry
+| TTClosedBinderList : Tok.t list -> ('self, local_binder_expr list list) entry
+
+type _ any_entry = TTAny : ('s, 'r) entry -> 's any_entry
+
+(* This computes the name of the level where to add a new rule *)
+let interp_constr_entry_key : type r. r target -> int -> r Gram.entry * int option =
+ fun forpat level -> match forpat with
+ | ForConstr ->
+ if level = 200 then Constr.binder_constr, None
+ else Constr.operconstr, Some level
+ | ForPattern -> Constr.pattern, Some level
+
+let target_entry : type s. s target -> s Gram.entry = function
+| ForConstr -> Constr.operconstr
+| ForPattern -> Constr.pattern
+
+let is_self from e = match e with
+| (NumLevel n, BorderProd (Right, _ (* Some(NonA|LeftA) *))) -> false
+| (NumLevel n, BorderProd (Left, _)) -> Int.equal from n
+| _ -> false
+
+let is_binder_level from e = match e with
+| (NumLevel 200, (BorderProd (Right, _) | InternalProd)) -> from = 200
+| _ -> false
+
+let make_sep_rules = function
+ | [tk] -> Atoken tk
+ | tkl ->
+ let rec mkrule : Tok.t list -> string rules = function
+ | [] -> Rules ({ norec_rule = Stop }, fun _ -> (* dropped anyway: *) "")
+ | tkn :: rem ->
+ let Rules ({ norec_rule = r }, f) = mkrule rem in
+ let r = { norec_rule = Next (r, Atoken tkn) } in
+ Rules (r, fun _ -> f)
+ in
+ let r = mkrule (List.rev tkl) in
+ Arules [r]
+
+let symbol_of_target : type s. _ -> _ -> _ -> s target -> (s, s) symbol = fun p assoc from forpat ->
+ if is_binder_level from p then Aentryl (target_entry forpat, 200)
+ else if is_self from p then Aself
+ else
+ let g = target_entry forpat in
+ let lev = adjust_level assoc from p in
+ begin match lev with
+ | None -> Aentry g
+ | Some None -> Anext
+ | Some (Some (lev, cur)) -> Aentryl (g, lev)
+ end
+
+let symbol_of_entry : type s r. _ -> _ -> (s, r) entry -> (s, r) symbol = fun assoc from typ -> match typ with
+| TTConstr (p, forpat) -> symbol_of_target p assoc from forpat
+| TTConstrList (typ', [], forpat) ->
+ Alist1 (symbol_of_target typ' assoc from forpat)
+| TTConstrList (typ', tkl, forpat) ->
+ Alist1sep (symbol_of_target typ' assoc from forpat, make_sep_rules tkl)
+| TTPattern p -> Aentryl (Constr.pattern, p)
+| TTClosedBinderList [] -> Alist1 (Aentry Constr.binder)
+| TTClosedBinderList tkl -> Alist1sep (Aentry Constr.binder, make_sep_rules tkl)
+| TTName -> Aentry Prim.name
+| TTOpenBinderList -> Aentry Constr.open_binders
+| TTBigint -> Aentry Prim.bigint
+| TTReference -> Aentry Constr.global
+
+let interp_entry forpat e = match e with
+| ETProdName -> TTAny TTName
+| ETProdReference -> TTAny TTReference
+| ETProdBigint -> TTAny TTBigint
+| ETProdConstr p -> TTAny (TTConstr (p, forpat))
+| ETProdPattern p -> TTAny (TTPattern p)
+| ETProdOther _ -> assert false (** not used *)
+| ETProdConstrList (p, tkl) -> TTAny (TTConstrList (p, tkl, forpat))
+| ETProdBinderList ETBinderOpen -> TTAny TTOpenBinderList
+| ETProdBinderList (ETBinderClosed tkl) -> TTAny (TTClosedBinderList tkl)
+
+let cases_pattern_expr_of_name { CAst.loc; v = na } = CAst.make ?loc @@ match na with
+ | Anonymous -> CPatAtom None
+ | Name id -> CPatAtom (Some (CAst.make ?loc @@ Ident id))
+
+type 'r env = {
+ constrs : 'r list;
+ constrlists : 'r list list;
+ binders : cases_pattern_expr list;
+ binderlists : local_binder_expr list list;
+}
+
+let push_constr subst v = { subst with constrs = v :: subst.constrs }
+
+let push_item : type s r. s target -> (s, r) entry -> s env -> r -> s env = fun forpat e subst v ->
+match e with
+| TTConstr _ -> push_constr subst v
+| TTName ->
+ begin match forpat with
+ | ForConstr -> { subst with binders = cases_pattern_expr_of_name v :: subst.binders }
+ | ForPattern -> push_constr subst (cases_pattern_expr_of_name v)
+ end
+| TTPattern _ ->
+ begin match forpat with
+ | ForConstr -> { subst with binders = v :: subst.binders }
+ | ForPattern -> push_constr subst v
+ end
+| TTOpenBinderList -> { subst with binderlists = v :: subst.binderlists }
+| TTClosedBinderList _ -> { subst with binderlists = List.flatten v :: subst.binderlists }
+| TTBigint ->
+ begin match forpat with
+ | ForConstr -> push_constr subst (CAst.make @@ CPrim (Numeral (v,true)))
+ | ForPattern -> push_constr subst (CAst.make @@ CPatPrim (Numeral (v,true)))
+ end
+| TTReference ->
+ begin match forpat with
+ | ForConstr -> push_constr subst (CAst.make @@ CRef (v, None))
+ | ForPattern -> push_constr subst (CAst.make @@ CPatAtom (Some v))
+ end
+| TTConstrList _ -> { subst with constrlists = v :: subst.constrlists }
+
+type (_, _) ty_symbol =
+| TyTerm : Tok.t -> ('s, string) ty_symbol
+| TyNonTerm : 's target * ('s, 'a) entry * ('s, 'a) symbol * bool -> ('s, 'a) ty_symbol
+
+type ('self, _, 'r) ty_rule =
+| TyStop : ('self, 'r, 'r) ty_rule
+| TyNext : ('self, 'a, 'r) ty_rule * ('self, 'b) ty_symbol -> ('self, 'b -> 'a, 'r) ty_rule
+| TyMark : int * bool * int * ('self, 'a, 'r) ty_rule -> ('self, 'a, 'r) ty_rule
+
+type 'r gen_eval = Loc.t -> 'r env -> 'r
+
+let rec ty_eval : type s a. (s, a, Loc.t -> s) ty_rule -> s gen_eval -> s env -> a = function
+| TyStop ->
+ fun f env loc -> f loc env
+| TyNext (rem, TyTerm _) ->
+ fun f env _ -> ty_eval rem f env
+| TyNext (rem, TyNonTerm (_, _, _, false)) ->
+ fun f env _ -> ty_eval rem f env
+| TyNext (rem, TyNonTerm (forpat, e, _, true)) ->
+ fun f env v ->
+ ty_eval rem f (push_item forpat e env v)
+| TyMark (n, b, p, rem) ->
+ fun f env ->
+ let heads, constrs = List.chop n env.constrs in
+ let constrlists, constrs =
+ if b then
+ (* We rearrange constrs = c1..cn rem and constrlists = [d1..dr e1..ep] rem' into
+ constrs = e1..ep rem and constrlists [c1..cn d1..dr] rem' *)
+ let constrlist = List.hd env.constrlists in
+ let constrlist, tail = List.chop (List.length constrlist - p) constrlist in
+ (heads @ constrlist) :: List.tl env.constrlists, tail @ constrs
+ else
+ (* We rearrange constrs = c1..cn e1..ep rem into
+ constrs = e1..ep rem and add a constr list [c1..cn] *)
+ let constrlist, tail = List.chop (n - p) heads in
+ constrlist :: env.constrlists, tail @ constrs
+ in
+ ty_eval rem f { env with constrs; constrlists; }
+
+let rec ty_erase : type s a r. (s, a, r) ty_rule -> (s, a, r) Extend.rule = function
+| TyStop -> Stop
+| TyMark (_, _, _, r) -> ty_erase r
+| TyNext (rem, TyTerm tok) -> Next (ty_erase rem, Atoken tok)
+| TyNext (rem, TyNonTerm (_, _, s, _)) -> Next (ty_erase rem, s)
+
+type ('self, 'r) any_ty_rule =
+| AnyTyRule : ('self, 'act, Loc.t -> 'r) ty_rule -> ('self, 'r) any_ty_rule
+
+let make_ty_rule assoc from forpat prods =
+ let rec make_ty_rule = function
+ | [] -> AnyTyRule TyStop
+ | GramConstrTerminal tok :: rem ->
+ let AnyTyRule r = make_ty_rule rem in
+ AnyTyRule (TyNext (r, TyTerm tok))
+ | GramConstrNonTerminal (e, var) :: rem ->
+ let AnyTyRule r = make_ty_rule rem in
+ let TTAny e = interp_entry forpat e in
+ let s = symbol_of_entry assoc from e in
+ let bind = match var with None -> false | Some _ -> true in
+ AnyTyRule (TyNext (r, TyNonTerm (forpat, e, s, bind)))
+ | GramConstrListMark (n, b, p) :: rem ->
+ let AnyTyRule r = make_ty_rule rem in
+ AnyTyRule (TyMark (n, b, p, r))
+ in
+ make_ty_rule (List.rev prods)
+
+let target_to_bool : type r. r target -> bool = function
+| ForConstr -> false
+| ForPattern -> true
+
+let prepare_empty_levels forpat (pos,p4assoc,name,reinit) =
+ let empty = (pos, [(name, p4assoc, [])]) in
+ if forpat then ExtendRule (Constr.pattern, reinit, empty)
+ else ExtendRule (Constr.operconstr, reinit, empty)
+
+let rec pure_sublevels : type a b c. int option -> (a, b, c) rule -> int list = fun level r -> match r with
+| Stop -> []
+| Next (rem, Aentryl (_, i)) ->
+ let rem = pure_sublevels level rem in
+ begin match level with
+ | Some j when Int.equal i j -> rem
+ | _ -> i :: rem
+ end
+| Next (rem, _) -> pure_sublevels level rem
+
+let make_act : type r. r target -> _ -> r gen_eval = function
+| ForConstr -> fun notation loc env ->
+ let env = (env.constrs, env.constrlists, env.binders, env.binderlists) in
+ CAst.make ~loc @@ CNotation (notation, env)
+| ForPattern -> fun notation loc env ->
+ let env = (env.constrs, env.constrlists) in
+ CAst.make ~loc @@ CPatNotation (notation, env, [])
+
+let extend_constr state forpat ng =
+ let n,_,_ = ng.notgram_level in
+ let assoc = ng.notgram_assoc in
+ let (entry, level) = interp_constr_entry_key forpat n in
+ let fold (accu, state) pt =
+ let AnyTyRule r = make_ty_rule assoc n forpat pt in
+ let symbs = ty_erase r in
+ let pure_sublevels = pure_sublevels level symbs in
+ let isforpat = target_to_bool forpat in
+ let needed_levels, state = register_empty_levels state isforpat pure_sublevels in
+ let (pos,p4assoc,name,reinit), state = find_position state isforpat assoc level in
+ let empty_rules = List.map (prepare_empty_levels isforpat) needed_levels in
+ let empty = { constrs = []; constrlists = []; binders = []; binderlists = [] } in
+ let act = ty_eval r (make_act forpat ng.notgram_notation) empty in
+ let rule = (name, p4assoc, [Rule (symbs, act)]) in
+ let r = ExtendRule (entry, reinit, (pos, [rule])) in
+ (accu @ empty_rules @ [r], state)
+ in
+ List.fold_left fold ([], state) ng.notgram_prods
+
+let constr_levels = GramState.field ()
+
+let extend_constr_notation ng state =
+ let levels = match GramState.get state constr_levels with
+ | None -> default_constr_levels
+ | Some lev -> lev
+ in
+ (* Add the notation in constr *)
+ let (r, levels) = extend_constr levels ForConstr ng in
+ (* Add the notation in cases_pattern *)
+ let (r', levels) = extend_constr levels ForPattern ng in
+ let state = GramState.set state constr_levels levels in
+ (r @ r', state)
+
+let constr_grammar : one_notation_grammar grammar_command =
+ create_grammar_command "Notation" extend_constr_notation
+
+let extend_constr_grammar ntn = extend_grammar_command constr_grammar ntn
diff --git a/vernac/egramcoq.mli b/vernac/egramcoq.mli
new file mode 100644
index 000000000..b0341e6a1
--- /dev/null
+++ b/vernac/egramcoq.mli
@@ -0,0 +1,19 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+(** Mapping of grammar productions to camlp5 actions *)
+
+(** This is the part specific to Coq-level Notation and Tactic Notation.
+ For the ML-level tactic and vernac extensions, see Egramml. *)
+
+(** {5 Adding notations} *)
+
+val extend_constr_grammar : Notation_gram.one_notation_grammar -> unit
+(** Add a term notation rule to the parsing system. *)
diff --git a/vernac/egramml.ml b/vernac/egramml.ml
new file mode 100644
index 000000000..048d4d93a
--- /dev/null
+++ b/vernac/egramml.ml
@@ -0,0 +1,84 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+open Util
+open Extend
+open Pcoq
+open Genarg
+open Vernacexpr
+
+(** Grammar extensions declared at ML level *)
+
+type 's grammar_prod_item =
+ | GramTerminal of string
+ | GramNonTerminal :
+ ('a raw_abstract_argument_type option * ('s, 'a) symbol) Loc.located -> 's grammar_prod_item
+
+type 'a ty_arg = ('a -> raw_generic_argument)
+
+type ('self, _, 'r) ty_rule =
+| TyStop : ('self, 'r, 'r) ty_rule
+| TyNext : ('self, 'a, 'r) ty_rule * ('self, 'b) Extend.symbol * 'b ty_arg option ->
+ ('self, 'b -> 'a, 'r) ty_rule
+
+type ('self, 'r) any_ty_rule =
+| AnyTyRule : ('self, 'act, Loc.t -> 'r) ty_rule -> ('self, 'r) any_ty_rule
+
+let rec ty_rule_of_gram = function
+| [] -> AnyTyRule TyStop
+| GramTerminal s :: rem ->
+ let AnyTyRule rem = ty_rule_of_gram rem in
+ let tok = Atoken (CLexer.terminal s) in
+ let r = TyNext (rem, tok, None) in
+ AnyTyRule r
+| GramNonTerminal (_, (t, tok)) :: rem ->
+ let AnyTyRule rem = ty_rule_of_gram rem in
+ let inj = Option.map (fun t obj -> Genarg.in_gen t obj) t in
+ let r = TyNext (rem, tok, inj) in
+ AnyTyRule r
+
+let rec ty_erase : type s a r. (s, a, r) ty_rule -> (s, a, r) Extend.rule = function
+| TyStop -> Extend.Stop
+| TyNext (rem, tok, _) -> Extend.Next (ty_erase rem, tok)
+
+type 'r gen_eval = Loc.t -> raw_generic_argument list -> 'r
+
+let rec ty_eval : type s a. (s, a, Loc.t -> s) ty_rule -> s gen_eval -> a = function
+| TyStop -> fun f loc -> f loc []
+| TyNext (rem, tok, None) -> fun f _ -> ty_eval rem f
+| TyNext (rem, tok, Some inj) -> fun f x ->
+ let f loc args = f loc (inj x :: args) in
+ ty_eval rem f
+
+let make_rule f prod =
+ let AnyTyRule ty_rule = ty_rule_of_gram (List.rev prod) in
+ let symb = ty_erase ty_rule in
+ let f loc l = f loc (List.rev l) in
+ let act = ty_eval ty_rule f in
+ Extend.Rule (symb, act)
+
+(** Vernac grammar extensions *)
+
+let vernac_exts = ref []
+
+let get_extend_vernac_rule (s, i) =
+ try
+ let find ((name, j), _) = String.equal name s && Int.equal i j in
+ let (_, rules) = List.find find !vernac_exts in
+ rules
+ with
+ | Failure _ -> raise Not_found
+
+let extend_vernac_command_grammar s nt gl =
+ let nt = Option.default Pvernac.Vernac_.command nt in
+ vernac_exts := (s,gl) :: !vernac_exts;
+ let mkact loc l = VernacExtend (s, l) in
+ let rules = [make_rule mkact gl] in
+ grammar_extend nt None (None, [None, None, rules])
diff --git a/vernac/egramml.mli b/vernac/egramml.mli
new file mode 100644
index 000000000..31aa1a989
--- /dev/null
+++ b/vernac/egramml.mli
@@ -0,0 +1,33 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+open Vernacexpr
+
+(** Mapping of grammar productions to camlp5 actions. *)
+
+(** This is the part specific to vernac extensions.
+ For the Coq-level Notation and Tactic Notation, see Egramcoq. *)
+
+type 's grammar_prod_item =
+ | GramTerminal of string
+ | GramNonTerminal : ('a Genarg.raw_abstract_argument_type option *
+ ('s, 'a) Extend.symbol) Loc.located -> 's grammar_prod_item
+
+val extend_vernac_command_grammar :
+ Vernacexpr.extend_name -> vernac_expr Pcoq.Gram.entry option ->
+ vernac_expr grammar_prod_item list -> unit
+
+val get_extend_vernac_rule : Vernacexpr.extend_name -> vernac_expr grammar_prod_item list
+
+(** Utility function reused in Egramcoq : *)
+
+val make_rule :
+ (Loc.t -> Genarg.raw_generic_argument list -> 'a) ->
+ 'a grammar_prod_item list -> 'a Extend.production_rule
diff --git a/vernac/explainErr.ml b/vernac/explainErr.ml
index f9167f969..504e7095b 100644
--- a/vernac/explainErr.ml
+++ b/vernac/explainErr.ml
@@ -55,7 +55,7 @@ let process_vernac_interp_error exn = match fst exn with
let msg =
if !Constrextern.print_universes then
str "." ++ spc() ++
- Univ.explain_universe_inconsistency Universes.pr_with_global_universes i
+ Univ.explain_universe_inconsistency UnivNames.pr_with_global_universes i
else
mt() in
wrap_vernac_error exn (str "Universe inconsistency" ++ msg ++ str ".")
@@ -66,6 +66,8 @@ let process_vernac_interp_error exn = match fst exn with
wrap_vernac_error exn (Himsg.explain_pretype_error ctx sigma te)
| Typeclasses_errors.TypeClassError(env, te) ->
wrap_vernac_error exn (Himsg.explain_typeclass_error env te)
+ | Implicit_quantifiers.MismatchedContextInstance(e,c,l,x) ->
+ wrap_vernac_error exn (Himsg.explain_mismatched_contexts e c l x)
| InductiveError e ->
wrap_vernac_error exn (Himsg.explain_inductive_error e)
| Modops.ModuleTypingError e ->
diff --git a/vernac/g_proofs.ml4 b/vernac/g_proofs.ml4
new file mode 100644
index 000000000..4b11276af
--- /dev/null
+++ b/vernac/g_proofs.ml4
@@ -0,0 +1,131 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+open Glob_term
+open Constrexpr
+open Vernacexpr
+open Proof_global
+
+open Pcoq
+open Pcoq.Prim
+open Pcoq.Constr
+open Pvernac.Vernac_
+
+let thm_token = G_vernac.thm_token
+
+let hint = Gram.entry_create "hint"
+
+let warn_deprecated_focus =
+ CWarnings.create ~name:"deprecated-focus" ~category:"deprecated"
+ (fun () ->
+ Pp.strbrk
+ "The Focus command is deprecated; use bullets or focusing brackets instead"
+ )
+
+let warn_deprecated_focus_n n =
+ CWarnings.create ~name:"deprecated-focus" ~category:"deprecated"
+ (fun () ->
+ Pp.(str "The Focus command is deprecated;" ++ spc ()
+ ++ str "use '" ++ int n ++ str ": {' instead")
+ )
+
+let warn_deprecated_unfocus =
+ CWarnings.create ~name:"deprecated-unfocus" ~category:"deprecated"
+ (fun () -> Pp.strbrk "The Unfocus command is deprecated")
+
+(* Proof commands *)
+GEXTEND Gram
+ GLOBAL: hint command;
+
+ opt_hintbases:
+ [ [ -> []
+ | ":"; l = LIST1 [id = IDENT -> id ] -> l ] ]
+ ;
+ command:
+ [ [ IDENT "Goal"; c = lconstr ->
+ VernacDefinition (Decl_kinds.(NoDischarge, Definition), ((CAst.make ~loc:!@loc Names.Anonymous), None), ProveBody ([], c))
+ | IDENT "Proof" -> VernacProof (None,None)
+ | IDENT "Proof" ; IDENT "Mode" ; mn = string -> VernacProofMode mn
+ | IDENT "Proof"; c = lconstr -> VernacExactProof c
+ | IDENT "Abort" -> VernacAbort None
+ | IDENT "Abort"; IDENT "All" -> VernacAbortAll
+ | IDENT "Abort"; id = identref -> VernacAbort (Some id)
+ | IDENT "Existential"; n = natural; c = constr_body ->
+ VernacSolveExistential (n,c)
+ | IDENT "Admitted" -> VernacEndProof Admitted
+ | IDENT "Qed" -> VernacEndProof (Proved (Opaque,None))
+ | IDENT "Save"; id = identref ->
+ VernacEndProof (Proved (Opaque, Some id))
+ | IDENT "Defined" -> VernacEndProof (Proved (Transparent,None))
+ | IDENT "Defined"; id=identref ->
+ VernacEndProof (Proved (Transparent,Some id))
+ | IDENT "Restart" -> VernacRestart
+ | IDENT "Undo" -> VernacUndo 1
+ | IDENT "Undo"; n = natural -> VernacUndo n
+ | IDENT "Undo"; IDENT "To"; n = natural -> VernacUndoTo n
+ | IDENT "Focus" ->
+ warn_deprecated_focus ~loc:!@loc ();
+ VernacFocus None
+ | IDENT "Focus"; n = natural ->
+ warn_deprecated_focus_n n ~loc:!@loc ();
+ VernacFocus (Some n)
+ | IDENT "Unfocus" ->
+ warn_deprecated_unfocus ~loc:!@loc ();
+ VernacUnfocus
+ | IDENT "Unfocused" -> VernacUnfocused
+ | IDENT "Show" -> VernacShow (ShowGoal OpenSubgoals)
+ | IDENT "Show"; n = natural -> VernacShow (ShowGoal (NthGoal n))
+ | IDENT "Show"; id = ident -> VernacShow (ShowGoal (GoalId id))
+ | IDENT "Show"; IDENT "Script" -> VernacShow ShowScript
+ | IDENT "Show"; IDENT "Existentials" -> VernacShow ShowExistentials
+ | IDENT "Show"; IDENT "Universes" -> VernacShow ShowUniverses
+ | IDENT "Show"; IDENT "Conjectures" -> VernacShow ShowProofNames
+ | IDENT "Show"; IDENT "Proof" -> VernacShow ShowProof
+ | IDENT "Show"; IDENT "Intro" -> VernacShow (ShowIntros false)
+ | IDENT "Show"; IDENT "Intros" -> VernacShow (ShowIntros true)
+ | IDENT "Show"; IDENT "Match"; id = reference -> VernacShow (ShowMatch id)
+ | IDENT "Guarded" -> VernacCheckGuard
+ (* Hints for Auto and EAuto *)
+ | IDENT "Create"; IDENT "HintDb" ;
+ id = IDENT ; b = [ "discriminated" -> true | -> false ] ->
+ VernacCreateHintDb (id, b)
+ | IDENT "Remove"; IDENT "Hints"; ids = LIST1 global; dbnames = opt_hintbases ->
+ VernacRemoveHints (dbnames, ids)
+ | IDENT "Hint"; h = hint; dbnames = opt_hintbases ->
+ VernacHints (dbnames, h)
+ ] ];
+ reference_or_constr:
+ [ [ r = global -> HintsReference r
+ | c = constr -> HintsConstr c ] ]
+ ;
+ hint:
+ [ [ IDENT "Resolve"; lc = LIST1 reference_or_constr; info = hint_info ->
+ HintsResolve (List.map (fun x -> (info, true, x)) lc)
+ | IDENT "Resolve"; "->"; lc = LIST1 global; n = OPT natural ->
+ HintsResolveIFF (true, lc, n)
+ | IDENT "Resolve"; "<-"; lc = LIST1 global; n = OPT natural ->
+ HintsResolveIFF (false, lc, n)
+ | IDENT "Immediate"; lc = LIST1 reference_or_constr -> HintsImmediate lc
+ | IDENT "Transparent"; lc = LIST1 global -> HintsTransparency (lc, true)
+ | IDENT "Opaque"; lc = LIST1 global -> HintsTransparency (lc, false)
+ | IDENT "Mode"; l = global; m = mode -> HintsMode (l, m)
+ | IDENT "Unfold"; lqid = LIST1 global -> HintsUnfold lqid
+ | IDENT "Constructors"; lc = LIST1 global -> HintsConstructors lc ] ]
+ ;
+ constr_body:
+ [ [ ":="; c = lconstr -> c
+ | ":"; t = lconstr; ":="; c = lconstr -> CAst.make ~loc:!@loc @@ CCast(c,CastConv t) ] ]
+ ;
+ mode:
+ [ [ l = LIST1 [ "+" -> ModeInput
+ | "!" -> ModeNoHeadEvar
+ | "-" -> ModeOutput ] -> l ] ]
+ ;
+END
diff --git a/vernac/g_vernac.ml4 b/vernac/g_vernac.ml4
new file mode 100644
index 000000000..3a59242de
--- /dev/null
+++ b/vernac/g_vernac.ml4
@@ -0,0 +1,1156 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+open Pp
+open CErrors
+open Util
+open Names
+open Glob_term
+open Vernacexpr
+open Constrexpr
+open Constrexpr_ops
+open Extend
+open Decl_kinds
+open Declaremods
+open Declarations
+open Namegen
+open Tok (* necessary for camlp5 *)
+
+open Pcoq
+open Pcoq.Prim
+open Pcoq.Constr
+open Pcoq.Module
+open Pvernac.Vernac_
+
+let vernac_kw = [ ";"; ","; ">->"; ":<"; "<:"; "where"; "at" ]
+let _ = List.iter CLexer.add_keyword vernac_kw
+
+(* Rem: do not join the different GEXTEND into one, it breaks native *)
+(* compilation on PowerPC and Sun architectures *)
+
+let query_command = Gram.entry_create "vernac:query_command"
+
+let subprf = Gram.entry_create "vernac:subprf"
+
+let class_rawexpr = Gram.entry_create "vernac:class_rawexpr"
+let thm_token = Gram.entry_create "vernac:thm_token"
+let def_body = Gram.entry_create "vernac:def_body"
+let decl_notation = Gram.entry_create "vernac:decl_notation"
+let record_field = Gram.entry_create "vernac:record_field"
+let of_type_with_opt_coercion = Gram.entry_create "vernac:of_type_with_opt_coercion"
+let instance_name = Gram.entry_create "vernac:instance_name"
+let section_subset_expr = Gram.entry_create "vernac:section_subset_expr"
+
+let make_bullet s =
+ let open Proof_bullet in
+ let n = String.length s in
+ match s.[0] with
+ | '-' -> Dash n
+ | '+' -> Plus n
+ | '*' -> Star n
+ | _ -> assert false
+
+let parse_compat_version ?(allow_old = true) = let open Flags in function
+ | "8.8" -> Current
+ | "8.7" -> V8_7
+ | "8.6" -> V8_6
+ | ("8.5" | "8.4" | "8.3" | "8.2" | "8.1" | "8.0") as s ->
+ CErrors.user_err ~hdr:"get_compat_version"
+ Pp.(str "Compatibility with version " ++ str s ++ str " not supported.")
+ | s ->
+ CErrors.user_err ~hdr:"get_compat_version"
+ Pp.(str "Unknown compatibility version \"" ++ str s ++ str "\".")
+
+GEXTEND Gram
+ GLOBAL: vernac_control gallina_ext noedit_mode subprf;
+ vernac_control: FIRST
+ [ [ IDENT "Time"; c = located_vernac -> VernacTime (false,c)
+ | IDENT "Redirect"; s = ne_string; c = located_vernac -> VernacRedirect (s, c)
+ | IDENT "Timeout"; n = natural; v = vernac_control -> VernacTimeout(n,v)
+ | IDENT "Fail"; v = vernac_control -> VernacFail v
+ | (f, v) = vernac -> VernacExpr(f, v) ]
+ ]
+ ;
+ vernac:
+ [ [ IDENT "Local"; (f, v) = vernac_poly -> (VernacLocal true :: f, v)
+ | IDENT "Global"; (f, v) = vernac_poly -> (VernacLocal false :: f, v)
+
+ | v = vernac_poly -> v ]
+ ]
+ ;
+ vernac_poly:
+ [ [ IDENT "Polymorphic"; (f, v) = vernac_aux -> (VernacPolymorphic true :: f, v)
+ | IDENT "Monomorphic"; (f, v) = vernac_aux -> (VernacPolymorphic false :: f, v)
+ | v = vernac_aux -> v ]
+ ]
+ ;
+ vernac_aux:
+ (* Better to parse "." here: in case of failure (e.g. in coerce_to_var), *)
+ (* "." is still in the stream and discard_to_dot works correctly *)
+ [ [ IDENT "Program"; g = gallina; "." -> ([VernacProgram], g)
+ | IDENT "Program"; g = gallina_ext; "." -> ([VernacProgram], g)
+ | g = gallina; "." -> ([], g)
+ | g = gallina_ext; "." -> ([], g)
+ | c = command; "." -> ([], c)
+ | c = syntax; "." -> ([], c)
+ | c = subprf -> ([], c)
+ ] ]
+ ;
+ vernac_aux: LAST
+ [ [ prfcom = command_entry -> ([], prfcom) ] ]
+ ;
+ noedit_mode:
+ [ [ c = query_command -> c None] ]
+ ;
+
+ subprf:
+ [ [ s = BULLET -> VernacBullet (make_bullet s)
+ | "{" -> VernacSubproof None
+ | "}" -> VernacEndSubproof
+ ] ]
+ ;
+
+ located_vernac:
+ [ [ v = vernac_control -> CAst.make ~loc:!@loc v ] ]
+ ;
+END
+
+let warn_plural_command =
+ CWarnings.create ~name:"plural-command" ~category:"pedantic" ~default:CWarnings.Disabled
+ (fun kwd -> strbrk (Printf.sprintf "Command \"%s\" expects more than one assumption." kwd))
+
+let test_plural_form loc kwd = function
+ | [(_,([_],_))] ->
+ warn_plural_command ~loc:!@loc kwd
+ | _ -> ()
+
+let test_plural_form_types loc kwd = function
+ | [([_],_)] ->
+ warn_plural_command ~loc:!@loc kwd
+ | _ -> ()
+
+let lname_of_lident : lident -> lname =
+ CAst.map (fun s -> Name s)
+
+let name_of_ident_decl : ident_decl -> name_decl =
+ on_fst lname_of_lident
+
+(* Gallina declarations *)
+GEXTEND Gram
+ GLOBAL: gallina gallina_ext thm_token def_body of_type_with_opt_coercion
+ record_field decl_notation rec_definition ident_decl univ_decl;
+
+ gallina:
+ (* Definition, Theorem, Variable, Axiom, ... *)
+ [ [ thm = thm_token; id = ident_decl; bl = binders; ":"; c = lconstr;
+ l = LIST0
+ [ "with"; id = ident_decl; bl = binders; ":"; c = lconstr ->
+ (id,(bl,c)) ] ->
+ VernacStartTheoremProof (thm, (id,(bl,c))::l)
+ | stre = assumption_token; nl = inline; bl = assum_list ->
+ VernacAssumption (stre, nl, bl)
+ | (kwd,stre) = assumptions_token; nl = inline; bl = assum_list ->
+ test_plural_form loc kwd bl;
+ VernacAssumption (stre, nl, bl)
+ | d = def_token; id = ident_decl; b = def_body ->
+ VernacDefinition (d, name_of_ident_decl id, b)
+ | IDENT "Let"; id = identref; b = def_body ->
+ VernacDefinition ((DoDischarge, Let), (lname_of_lident id, None), b)
+ (* Gallina inductive declarations *)
+ | cum = OPT cumulativity_token; priv = private_token; f = finite_token;
+ indl = LIST1 inductive_definition SEP "with" ->
+ let (k,f) = f in
+ let indl=List.map (fun ((a,b,c,d),e) -> ((a,b,c,k,d),e)) indl in
+ VernacInductive (cum, priv,f,indl)
+ | "Fixpoint"; recs = LIST1 rec_definition SEP "with" ->
+ VernacFixpoint (NoDischarge, recs)
+ | IDENT "Let"; "Fixpoint"; recs = LIST1 rec_definition SEP "with" ->
+ VernacFixpoint (DoDischarge, recs)
+ | "CoFixpoint"; corecs = LIST1 corec_definition SEP "with" ->
+ VernacCoFixpoint (NoDischarge, corecs)
+ | IDENT "Let"; "CoFixpoint"; corecs = LIST1 corec_definition SEP "with" ->
+ VernacCoFixpoint (DoDischarge, corecs)
+ | IDENT "Scheme"; l = LIST1 scheme SEP "with" -> VernacScheme l
+ | IDENT "Combined"; IDENT "Scheme"; id = identref; IDENT "from";
+ l = LIST1 identref SEP "," -> VernacCombinedScheme (id, l)
+ | IDENT "Register"; IDENT "Inline"; id = identref ->
+ VernacRegister(id, RegisterInline)
+ | IDENT "Universe"; l = LIST1 identref -> VernacUniverse l
+ | IDENT "Universes"; l = LIST1 identref -> VernacUniverse l
+ | IDENT "Constraint"; l = LIST1 univ_constraint SEP "," -> VernacConstraint l
+ ] ]
+ ;
+
+ thm_token:
+ [ [ "Theorem" -> Theorem
+ | IDENT "Lemma" -> Lemma
+ | IDENT "Fact" -> Fact
+ | IDENT "Remark" -> Remark
+ | IDENT "Corollary" -> Corollary
+ | IDENT "Proposition" -> Proposition
+ | IDENT "Property" -> Property ] ]
+ ;
+ def_token:
+ [ [ "Definition" -> (NoDischarge,Definition)
+ | IDENT "Example" -> (NoDischarge,Example)
+ | IDENT "SubClass" -> (NoDischarge,SubClass) ] ]
+ ;
+ assumption_token:
+ [ [ "Hypothesis" -> (DoDischarge, Logical)
+ | "Variable" -> (DoDischarge, Definitional)
+ | "Axiom" -> (NoDischarge, Logical)
+ | "Parameter" -> (NoDischarge, Definitional)
+ | IDENT "Conjecture" -> (NoDischarge, Conjectural) ] ]
+ ;
+ assumptions_token:
+ [ [ IDENT "Hypotheses" -> ("Hypotheses", (DoDischarge, Logical))
+ | IDENT "Variables" -> ("Variables", (DoDischarge, Definitional))
+ | IDENT "Axioms" -> ("Axioms", (NoDischarge, Logical))
+ | IDENT "Parameters" -> ("Parameters", (NoDischarge, Definitional))
+ | IDENT "Conjectures" -> ("Conjectures", (NoDischarge, Conjectural)) ] ]
+ ;
+ inline:
+ [ [ IDENT "Inline"; "("; i = INT; ")" -> InlineAt (int_of_string i)
+ | IDENT "Inline" -> DefaultInline
+ | -> NoInline] ]
+ ;
+ univ_constraint:
+ [ [ l = universe_level; ord = [ "<" -> Univ.Lt | "=" -> Univ.Eq | "<=" -> Univ.Le ];
+ r = universe_level -> (l, ord, r) ] ]
+ ;
+ univ_decl :
+ [ [ "@{" ; l = LIST0 identref; ext = [ "+" -> true | -> false ];
+ cs = [ "|"; l' = LIST0 univ_constraint SEP ",";
+ ext = [ "+" -> true | -> false ]; "}" -> (l',ext)
+ | ext = [ "}" -> true | "|}" -> false ] -> ([], ext) ]
+ ->
+ let open UState in
+ { univdecl_instance = l;
+ univdecl_extensible_instance = ext;
+ univdecl_constraints = fst cs;
+ univdecl_extensible_constraints = snd cs }
+ ] ]
+ ;
+ ident_decl:
+ [ [ i = identref; l = OPT univ_decl -> (i, l)
+ ] ]
+ ;
+ finite_token:
+ [ [ IDENT "Inductive" -> (Inductive_kw,Finite)
+ | IDENT "CoInductive" -> (CoInductive,CoFinite)
+ | IDENT "Variant" -> (Variant,BiFinite)
+ | IDENT "Record" -> (Record,BiFinite)
+ | IDENT "Structure" -> (Structure,BiFinite)
+ | IDENT "Class" -> (Class true,BiFinite) ] ]
+ ;
+ cumulativity_token:
+ [ [ IDENT "Cumulative" -> VernacCumulative
+ | IDENT "NonCumulative" -> VernacNonCumulative ] ]
+ ;
+ private_token:
+ [ [ IDENT "Private" -> true | -> false ] ]
+ ;
+ (* Simple definitions *)
+ def_body:
+ [ [ bl = binders; ":="; red = reduce; c = lconstr ->
+ if List.exists (function CLocalPattern _ -> true | _ -> false) bl
+ then
+ (* FIXME: "red" will be applied to types in bl and Cast with remain *)
+ let c = mkCLambdaN ~loc:!@loc bl c in
+ DefineBody ([], red, c, None)
+ else
+ (match c with
+ | { CAst.v = CCast(c, CastConv t) } -> DefineBody (bl, red, c, Some t)
+ | _ -> DefineBody (bl, red, c, None))
+ | bl = binders; ":"; t = lconstr; ":="; red = reduce; c = lconstr ->
+ let ((bl, c), tyo) =
+ if List.exists (function CLocalPattern _ -> true | _ -> false) bl
+ then
+ (* FIXME: "red" will be applied to types in bl and Cast with remain *)
+ let c = CAst.make ~loc:!@loc @@ CCast (c, CastConv t) in
+ (([],mkCLambdaN ~loc:!@loc bl c), None)
+ else ((bl, c), Some t)
+ in
+ DefineBody (bl, red, c, tyo)
+ | bl = binders; ":"; t = lconstr ->
+ ProveBody (bl, t) ] ]
+ ;
+ reduce:
+ [ [ IDENT "Eval"; r = red_expr; "in" -> Some r
+ | -> None ] ]
+ ;
+ one_decl_notation:
+ [ [ ntn = ne_lstring; ":="; c = constr;
+ scopt = OPT [ ":"; sc = IDENT -> sc] -> (ntn,c,scopt) ] ]
+ ;
+ decl_notation:
+ [ [ "where"; l = LIST1 one_decl_notation SEP IDENT "and" -> l
+ | -> [] ] ]
+ ;
+ (* Inductives and records *)
+ opt_constructors_or_fields:
+ [ [ ":="; lc = constructor_list_or_record_decl -> lc
+ | -> RecordDecl (None, []) ] ]
+ ;
+ inductive_definition:
+ [ [ oc = opt_coercion; id = ident_decl; indpar = binders;
+ c = OPT [ ":"; c = lconstr -> c ];
+ lc=opt_constructors_or_fields; ntn = decl_notation ->
+ (((oc,id),indpar,c,lc),ntn) ] ]
+ ;
+ constructor_list_or_record_decl:
+ [ [ "|"; l = LIST1 constructor SEP "|" -> Constructors l
+ | id = identref ; c = constructor_type; "|"; l = LIST0 constructor SEP "|" ->
+ Constructors ((c id)::l)
+ | id = identref ; c = constructor_type -> Constructors [ c id ]
+ | cstr = identref; "{"; fs = record_fields; "}" ->
+ RecordDecl (Some cstr,fs)
+ | "{";fs = record_fields; "}" -> RecordDecl (None,fs)
+ | -> Constructors [] ] ]
+ ;
+(*
+ csort:
+ [ [ s = sort -> CSort (loc,s) ] ]
+ ;
+*)
+ opt_coercion:
+ [ [ ">" -> true
+ | -> false ] ]
+ ;
+ (* (co)-fixpoints *)
+ rec_definition:
+ [ [ id = ident_decl;
+ bl = binders_fixannot;
+ ty = type_cstr;
+ def = OPT [":="; def = lconstr -> def]; ntn = decl_notation ->
+ let bl, annot = bl in ((id,annot,bl,ty,def),ntn) ] ]
+ ;
+ corec_definition:
+ [ [ id = ident_decl; bl = binders; ty = type_cstr;
+ def = OPT [":="; def = lconstr -> def]; ntn = decl_notation ->
+ ((id,bl,ty,def),ntn) ] ]
+ ;
+ type_cstr:
+ [ [ ":"; c=lconstr -> c
+ | -> CAst.make ~loc:!@loc @@ CHole (None, IntroAnonymous, None) ] ]
+ ;
+ (* Inductive schemes *)
+ scheme:
+ [ [ kind = scheme_kind -> (None,kind)
+ | id = identref; ":="; kind = scheme_kind -> (Some id,kind) ] ]
+ ;
+ scheme_kind:
+ [ [ IDENT "Induction"; "for"; ind = smart_global;
+ IDENT "Sort"; s = sort_family-> InductionScheme(true,ind,s)
+ | IDENT "Minimality"; "for"; ind = smart_global;
+ IDENT "Sort"; s = sort_family-> InductionScheme(false,ind,s)
+ | IDENT "Elimination"; "for"; ind = smart_global;
+ IDENT "Sort"; s = sort_family-> CaseScheme(true,ind,s)
+ | IDENT "Case"; "for"; ind = smart_global;
+ IDENT "Sort"; s = sort_family-> CaseScheme(false,ind,s)
+ | IDENT "Equality"; "for" ; ind = smart_global -> EqualityScheme(ind) ] ]
+ ;
+ (* Various Binders *)
+(*
+ (* ... without coercions *)
+ binder_nodef:
+ [ [ b = binder_let ->
+ (match b with
+ CLocalAssum(l,ty) -> (l,ty)
+ | CLocalDef _ ->
+ Util.user_err_loc
+ (loc,"fix_param",Pp.str"defined binder not allowed here.")) ] ]
+ ;
+*)
+ (* ... with coercions *)
+ record_field:
+ [ [ bd = record_binder; pri = OPT [ "|"; n = natural -> n ];
+ ntn = decl_notation -> (bd,pri),ntn ] ]
+ ;
+ record_fields:
+ [ [ f = record_field; ";"; fs = record_fields -> f :: fs
+ | f = record_field; ";" -> [f]
+ | f = record_field -> [f]
+ | -> []
+ ] ]
+ ;
+ record_binder_body:
+ [ [ l = binders; oc = of_type_with_opt_coercion;
+ t = lconstr -> fun id -> (oc,AssumExpr (id,mkCProdN ~loc:!@loc l t))
+ | l = binders; oc = of_type_with_opt_coercion;
+ t = lconstr; ":="; b = lconstr -> fun id ->
+ (oc,DefExpr (id,mkCLambdaN ~loc:!@loc l b,Some (mkCProdN ~loc:!@loc l t)))
+ | l = binders; ":="; b = lconstr -> fun id ->
+ match b.CAst.v with
+ | CCast(b', (CastConv t|CastVM t|CastNative t)) ->
+ (None,DefExpr(id,mkCLambdaN ~loc:!@loc l b',Some (mkCProdN ~loc:!@loc l t)))
+ | _ ->
+ (None,DefExpr(id,mkCLambdaN ~loc:!@loc l b,None)) ] ]
+ ;
+ record_binder:
+ [ [ id = name -> (None,AssumExpr(id, CAst.make ~loc:!@loc @@ CHole (None, IntroAnonymous, None)))
+ | id = name; f = record_binder_body -> f id ] ]
+ ;
+ assum_list:
+ [ [ bl = LIST1 assum_coe -> bl | b = simple_assum_coe -> [b] ] ]
+ ;
+ assum_coe:
+ [ [ "("; a = simple_assum_coe; ")" -> a ] ]
+ ;
+ simple_assum_coe:
+ [ [ idl = LIST1 ident_decl; oc = of_type_with_opt_coercion; c = lconstr ->
+ (not (Option.is_empty oc),(idl,c)) ] ]
+ ;
+
+ constructor_type:
+ [[ l = binders;
+ t= [ coe = of_type_with_opt_coercion; c = lconstr ->
+ fun l id -> (not (Option.is_empty coe),(id,mkCProdN ~loc:!@loc l c))
+ | ->
+ fun l id -> (false,(id,mkCProdN ~loc:!@loc l (CAst.make ~loc:!@loc @@ CHole (None, IntroAnonymous, None)))) ]
+ -> t l
+ ]]
+;
+
+ constructor:
+ [ [ id = identref; c=constructor_type -> c id ] ]
+ ;
+ of_type_with_opt_coercion:
+ [ [ ":>>" -> Some false
+ | ":>"; ">" -> Some false
+ | ":>" -> Some true
+ | ":"; ">"; ">" -> Some false
+ | ":"; ">" -> Some true
+ | ":" -> None ] ]
+ ;
+END
+
+let only_starredidentrefs =
+ Gram.Entry.of_parser "test_only_starredidentrefs"
+ (fun strm ->
+ let rec aux n =
+ match Util.stream_nth n strm with
+ | KEYWORD "." -> ()
+ | KEYWORD ")" -> ()
+ | (IDENT _ | KEYWORD "Type" | KEYWORD "*") -> aux (n+1)
+ | _ -> raise Stream.Failure in
+ aux 0)
+let starredidentreflist_to_expr l =
+ match l with
+ | [] -> SsEmpty
+ | x :: xs -> List.fold_right (fun i acc -> SsUnion(i,acc)) xs x
+
+let warn_deprecated_include_type =
+ CWarnings.create ~name:"deprecated-include-type" ~category:"deprecated"
+ (fun () -> strbrk "Include Type is deprecated; use Include instead")
+
+(* Modules and Sections *)
+GEXTEND Gram
+ GLOBAL: gallina_ext module_expr module_type section_subset_expr;
+
+ gallina_ext:
+ [ [ (* Interactive module declaration *)
+ IDENT "Module"; export = export_token; id = identref;
+ bl = LIST0 module_binder; sign = of_module_type;
+ body = is_module_expr ->
+ VernacDefineModule (export, id, bl, sign, body)
+ | IDENT "Module"; "Type"; id = identref;
+ bl = LIST0 module_binder; sign = check_module_types;
+ body = is_module_type ->
+ VernacDeclareModuleType (id, bl, sign, body)
+ | IDENT "Declare"; IDENT "Module"; export = export_token; id = identref;
+ bl = LIST0 module_binder; ":"; mty = module_type_inl ->
+ VernacDeclareModule (export, id, bl, mty)
+ (* Section beginning *)
+ | IDENT "Section"; id = identref -> VernacBeginSection id
+ | IDENT "Chapter"; id = identref -> VernacBeginSection id
+
+ (* This end a Section a Module or a Module Type *)
+ | IDENT "End"; id = identref -> VernacEndSegment id
+
+ (* Naming a set of section hyps *)
+ | IDENT "Collection"; id = identref; ":="; expr = section_subset_expr ->
+ VernacNameSectionHypSet (id, expr)
+
+ (* Requiring an already compiled module *)
+ | IDENT "Require"; export = export_token; qidl = LIST1 global ->
+ VernacRequire (None, export, qidl)
+ | IDENT "From" ; ns = global ; IDENT "Require"; export = export_token
+ ; qidl = LIST1 global ->
+ VernacRequire (Some ns, export, qidl)
+ | IDENT "Import"; qidl = LIST1 global -> VernacImport (false,qidl)
+ | IDENT "Export"; qidl = LIST1 global -> VernacImport (true,qidl)
+ | IDENT "Include"; e = module_type_inl; l = LIST0 ext_module_expr ->
+ VernacInclude(e::l)
+ | IDENT "Include"; "Type"; e = module_type_inl; l = LIST0 ext_module_type ->
+ warn_deprecated_include_type ~loc:!@loc ();
+ VernacInclude(e::l) ] ]
+ ;
+ export_token:
+ [ [ IDENT "Import" -> Some false
+ | IDENT "Export" -> Some true
+ | -> None ] ]
+ ;
+ ext_module_type:
+ [ [ "<+"; mty = module_type_inl -> mty ] ]
+ ;
+ ext_module_expr:
+ [ [ "<+"; mexpr = module_expr_inl -> mexpr ] ]
+ ;
+ check_module_type:
+ [ [ "<:"; mty = module_type_inl -> mty ] ]
+ ;
+ check_module_types:
+ [ [ mtys = LIST0 check_module_type -> mtys ] ]
+ ;
+ of_module_type:
+ [ [ ":"; mty = module_type_inl -> Enforce mty
+ | mtys = check_module_types -> Check mtys ] ]
+ ;
+ is_module_type:
+ [ [ ":="; mty = module_type_inl ; l = LIST0 ext_module_type -> (mty::l)
+ | -> [] ] ]
+ ;
+ is_module_expr:
+ [ [ ":="; mexpr = module_expr_inl; l = LIST0 ext_module_expr -> (mexpr::l)
+ | -> [] ] ]
+ ;
+ functor_app_annot:
+ [ [ "["; IDENT "inline"; "at"; IDENT "level"; i = INT; "]" ->
+ InlineAt (int_of_string i)
+ | "["; IDENT "no"; IDENT "inline"; "]" -> NoInline
+ | -> DefaultInline
+ ] ]
+ ;
+ module_expr_inl:
+ [ [ "!"; me = module_expr -> (me,NoInline)
+ | me = module_expr; a = functor_app_annot -> (me,a) ] ]
+ ;
+ module_type_inl:
+ [ [ "!"; me = module_type -> (me,NoInline)
+ | me = module_type; a = functor_app_annot -> (me,a) ] ]
+ ;
+ (* Module binder *)
+ module_binder:
+ [ [ "("; export = export_token; idl = LIST1 identref; ":";
+ mty = module_type_inl; ")" -> (export,idl,mty) ] ]
+ ;
+ (* Module expressions *)
+ module_expr:
+ [ [ me = module_expr_atom -> me
+ | me1 = module_expr; me2 = module_expr_atom -> CAst.make ~loc:!@loc @@ CMapply (me1,me2)
+ ] ]
+ ;
+ module_expr_atom:
+ [ [ qid = qualid -> CAst.make ~loc:!@loc @@ CMident (qid.CAst.v) | "("; me = module_expr; ")" -> me ] ]
+ ;
+ with_declaration:
+ [ [ "Definition"; fqid = fullyqualid; udecl = OPT univ_decl; ":="; c = Constr.lconstr ->
+ CWith_Definition (fqid,udecl,c)
+ | IDENT "Module"; fqid = fullyqualid; ":="; qid = qualid ->
+ CWith_Module (fqid,qid)
+ ] ]
+ ;
+ module_type:
+ [ [ qid = qualid -> CAst.make ~loc:!@loc @@ CMident (qid.CAst.v)
+ | "("; mt = module_type; ")" -> mt
+ | mty = module_type; me = module_expr_atom ->
+ CAst.make ~loc:!@loc @@ CMapply (mty,me)
+ | mty = module_type; "with"; decl = with_declaration ->
+ CAst.make ~loc:!@loc @@ CMwith (mty,decl)
+ ] ]
+ ;
+ (* Proof using *)
+ section_subset_expr:
+ [ [ only_starredidentrefs; l = LIST0 starredidentref ->
+ starredidentreflist_to_expr l
+ | e = ssexpr -> e ]]
+ ;
+ starredidentref:
+ [ [ i = identref -> SsSingl i
+ | i = identref; "*" -> SsFwdClose(SsSingl i)
+ | "Type" -> SsType
+ | "Type"; "*" -> SsFwdClose SsType ]]
+ ;
+ ssexpr:
+ [ "35"
+ [ "-"; e = ssexpr -> SsCompl e ]
+ | "50"
+ [ e1 = ssexpr; "-"; e2 = ssexpr->SsSubstr(e1,e2)
+ | e1 = ssexpr; "+"; e2 = ssexpr->SsUnion(e1,e2)]
+ | "0"
+ [ i = starredidentref -> i
+ | "("; only_starredidentrefs; l = LIST0 starredidentref; ")"->
+ starredidentreflist_to_expr l
+ | "("; only_starredidentrefs; l = LIST0 starredidentref; ")"; "*" ->
+ SsFwdClose(starredidentreflist_to_expr l)
+ | "("; e = ssexpr; ")"-> e
+ | "("; e = ssexpr; ")"; "*" -> SsFwdClose e ] ]
+ ;
+END
+
+(* Extensions: implicits, coercions, etc. *)
+GEXTEND Gram
+ GLOBAL: gallina_ext instance_name hint_info;
+
+ gallina_ext:
+ [ [ (* Transparent and Opaque *)
+ IDENT "Transparent"; l = LIST1 smart_global ->
+ VernacSetOpacity (Conv_oracle.transparent, l)
+ | IDENT "Opaque"; l = LIST1 smart_global ->
+ VernacSetOpacity (Conv_oracle.Opaque, l)
+ | IDENT "Strategy"; l =
+ LIST1 [ v=strategy_level; "["; q=LIST1 smart_global; "]" -> (v,q)] ->
+ VernacSetStrategy l
+ (* Canonical structure *)
+ | IDENT "Canonical"; IDENT "Structure"; qid = global ->
+ VernacCanonical CAst.(make ~loc:!@loc @@ AN qid)
+ | IDENT "Canonical"; IDENT "Structure"; ntn = by_notation ->
+ VernacCanonical CAst.(make ~loc:!@loc @@ ByNotation ntn)
+ | IDENT "Canonical"; IDENT "Structure"; qid = global; d = def_body ->
+ let s = coerce_reference_to_id qid in
+ VernacDefinition ((NoDischarge,CanonicalStructure),((CAst.make (Name s)),None),d)
+
+ (* Coercions *)
+ | IDENT "Coercion"; qid = global; d = def_body ->
+ let s = coerce_reference_to_id qid in
+ VernacDefinition ((NoDischarge,Coercion),((CAst.make (Name s)),None),d)
+ | IDENT "Identity"; IDENT "Coercion"; f = identref; ":";
+ s = class_rawexpr; ">->"; t = class_rawexpr ->
+ VernacIdentityCoercion (f, s, t)
+ | IDENT "Coercion"; qid = global; ":"; s = class_rawexpr; ">->";
+ t = class_rawexpr ->
+ VernacCoercion (CAst.make ~loc:!@loc @@ AN qid, s, t)
+ | IDENT "Coercion"; ntn = by_notation; ":"; s = class_rawexpr; ">->";
+ t = class_rawexpr ->
+ VernacCoercion (CAst.make ~loc:!@loc @@ ByNotation ntn, s, t)
+
+ | IDENT "Context"; c = LIST1 binder ->
+ VernacContext (List.flatten c)
+
+ | IDENT "Instance"; namesup = instance_name; ":";
+ expl = [ "!" -> Decl_kinds.Implicit | -> Decl_kinds.Explicit ] ; t = operconstr LEVEL "200";
+ info = hint_info ;
+ props = [ ":="; "{"; r = record_declaration; "}" -> Some (true,r) |
+ ":="; c = lconstr -> Some (false,c) | -> None ] ->
+ VernacInstance (false,snd namesup,(fst namesup,expl,t),props,info)
+
+ | IDENT "Existing"; IDENT "Instance"; id = global;
+ info = hint_info ->
+ VernacDeclareInstances [id, info]
+
+ | IDENT "Existing"; IDENT "Instances"; ids = LIST1 global;
+ pri = OPT [ "|"; i = natural -> i ] ->
+ let info = { Typeclasses.hint_priority = pri; hint_pattern = None } in
+ let insts = List.map (fun i -> (i, info)) ids in
+ VernacDeclareInstances insts
+
+ | IDENT "Existing"; IDENT "Class"; is = global -> VernacDeclareClass is
+
+ (* Arguments *)
+ | IDENT "Arguments"; qid = smart_global;
+ args = LIST0 argument_spec_block;
+ more_implicits = OPT
+ [ ","; impl = LIST1
+ [ impl = LIST0 more_implicits_block -> List.flatten impl]
+ SEP "," -> impl
+ ];
+ mods = OPT [ ":"; l = LIST1 arguments_modifier SEP "," -> l ] ->
+ let mods = match mods with None -> [] | Some l -> List.flatten l in
+ let slash_position = ref None in
+ let rec parse_args i = function
+ | [] -> []
+ | `Id x :: args -> x :: parse_args (i+1) args
+ | `Slash :: args ->
+ if Option.is_empty !slash_position then
+ (slash_position := Some i; parse_args i args)
+ else
+ user_err Pp.(str "The \"/\" modifier can occur only once")
+ in
+ let args = parse_args 0 (List.flatten args) in
+ let more_implicits = Option.default [] more_implicits in
+ VernacArguments (qid, args, more_implicits, !slash_position, mods)
+
+ | IDENT "Implicit"; "Type"; bl = reserv_list ->
+ VernacReserve bl
+
+ | IDENT "Implicit"; IDENT "Types"; bl = reserv_list ->
+ test_plural_form_types loc "Implicit Types" bl;
+ VernacReserve bl
+
+ | IDENT "Generalizable";
+ gen = [IDENT "All"; IDENT "Variables" -> Some []
+ | IDENT "No"; IDENT "Variables" -> None
+ | ["Variable" | IDENT "Variables"];
+ idl = LIST1 identref -> Some idl ] ->
+ VernacGeneralizable gen ] ]
+ ;
+ arguments_modifier:
+ [ [ IDENT "simpl"; IDENT "nomatch" -> [`ReductionDontExposeCase]
+ | IDENT "simpl"; IDENT "never" -> [`ReductionNeverUnfold]
+ | IDENT "default"; IDENT "implicits" -> [`DefaultImplicits]
+ | IDENT "clear"; IDENT "implicits" -> [`ClearImplicits]
+ | IDENT "clear"; IDENT "scopes" -> [`ClearScopes]
+ | IDENT "rename" -> [`Rename]
+ | IDENT "assert" -> [`Assert]
+ | IDENT "extra"; IDENT "scopes" -> [`ExtraScopes]
+ | IDENT "clear"; IDENT "scopes"; IDENT "and"; IDENT "implicits" ->
+ [`ClearImplicits; `ClearScopes]
+ | IDENT "clear"; IDENT "implicits"; IDENT "and"; IDENT "scopes" ->
+ [`ClearImplicits; `ClearScopes]
+ ] ]
+ ;
+ scope:
+ [ [ "%"; key = IDENT -> key ] ]
+ ;
+ argument_spec: [
+ [ b = OPT "!"; id = name ; s = OPT scope ->
+ id.CAst.v, not (Option.is_empty b), Option.map (fun x -> CAst.make ~loc:!@loc x) s
+ ]
+ ];
+ (* List of arguments implicit status, scope, modifiers *)
+ argument_spec_block: [
+ [ item = argument_spec ->
+ let name, recarg_like, notation_scope = item in
+ [`Id { name=name; recarg_like=recarg_like;
+ notation_scope=notation_scope;
+ implicit_status = NotImplicit}]
+ | "/" -> [`Slash]
+ | "("; items = LIST1 argument_spec; ")"; sc = OPT scope ->
+ let f x = match sc, x with
+ | None, x -> x | x, None -> Option.map (fun y -> CAst.make ~loc:!@loc y) x
+ | Some _, Some _ -> user_err Pp.(str "scope declared twice") in
+ List.map (fun (name,recarg_like,notation_scope) ->
+ `Id { name=name; recarg_like=recarg_like;
+ notation_scope=f notation_scope;
+ implicit_status = NotImplicit}) items
+ | "["; items = LIST1 argument_spec; "]"; sc = OPT scope ->
+ let f x = match sc, x with
+ | None, x -> x | x, None -> Option.map (fun y -> CAst.make ~loc:!@loc y) x
+ | Some _, Some _ -> user_err Pp.(str "scope declared twice") in
+ List.map (fun (name,recarg_like,notation_scope) ->
+ `Id { name=name; recarg_like=recarg_like;
+ notation_scope=f notation_scope;
+ implicit_status = Implicit}) items
+ | "{"; items = LIST1 argument_spec; "}"; sc = OPT scope ->
+ let f x = match sc, x with
+ | None, x -> x | x, None -> Option.map (fun y -> CAst.make ~loc:!@loc y) x
+ | Some _, Some _ -> user_err Pp.(str "scope declared twice") in
+ List.map (fun (name,recarg_like,notation_scope) ->
+ `Id { name=name; recarg_like=recarg_like;
+ notation_scope=f notation_scope;
+ implicit_status = MaximallyImplicit}) items
+ ]
+ ];
+ (* Same as [argument_spec_block], but with only implicit status and names *)
+ more_implicits_block: [
+ [ name = name -> [(name.CAst.v, Vernacexpr.NotImplicit)]
+ | "["; items = LIST1 name; "]" ->
+ List.map (fun name -> (name.CAst.v, Vernacexpr.Implicit)) items
+ | "{"; items = LIST1 name; "}" ->
+ List.map (fun name -> (name.CAst.v, Vernacexpr.MaximallyImplicit)) items
+ ]
+ ];
+ strategy_level:
+ [ [ IDENT "expand" -> Conv_oracle.Expand
+ | IDENT "opaque" -> Conv_oracle.Opaque
+ | n=INT -> Conv_oracle.Level (int_of_string n)
+ | "-"; n=INT -> Conv_oracle.Level (- int_of_string n)
+ | IDENT "transparent" -> Conv_oracle.transparent ] ]
+ ;
+ instance_name:
+ [ [ name = ident_decl; sup = OPT binders ->
+ (CAst.map (fun id -> Name id) (fst name), snd name),
+ (Option.default [] sup)
+ | -> ((CAst.make ~loc:!@loc Anonymous), None), [] ] ]
+ ;
+ hint_info:
+ [ [ "|"; i = OPT natural; pat = OPT constr_pattern ->
+ { Typeclasses.hint_priority = i; hint_pattern = pat }
+ | -> { Typeclasses.hint_priority = None; hint_pattern = None } ] ]
+ ;
+ reserv_list:
+ [ [ bl = LIST1 reserv_tuple -> bl | b = simple_reserv -> [b] ] ]
+ ;
+ reserv_tuple:
+ [ [ "("; a = simple_reserv; ")" -> a ] ]
+ ;
+ simple_reserv:
+ [ [ idl = LIST1 identref; ":"; c = lconstr -> (idl,c) ] ]
+ ;
+
+END
+
+GEXTEND Gram
+ GLOBAL: command query_command class_rawexpr gallina_ext;
+
+ gallina_ext:
+ [ [ IDENT "Export"; "Set"; table = option_table; v = option_value ->
+ VernacSetOption (true, table, v)
+ | IDENT "Export"; IDENT "Unset"; table = option_table ->
+ VernacUnsetOption (true, table)
+ ] ];
+
+ command:
+ [ [ IDENT "Comments"; l = LIST0 comment -> VernacComments l
+
+ (* Hack! Should be in grammar_ext, but camlp5 factorizes badly *)
+ | IDENT "Declare"; IDENT "Instance"; namesup = instance_name; ":";
+ expl = [ "!" -> Decl_kinds.Implicit | -> Decl_kinds.Explicit ] ; t = operconstr LEVEL "200";
+ info = hint_info ->
+ VernacInstance (true, snd namesup, (fst namesup, expl, t), None, info)
+
+ (* System directory *)
+ | IDENT "Pwd" -> VernacChdir None
+ | IDENT "Cd" -> VernacChdir None
+ | IDENT "Cd"; dir = ne_string -> VernacChdir (Some dir)
+
+ | IDENT "Load"; verbosely = [ IDENT "Verbose" -> true | -> false ];
+ s = [ s = ne_string -> s | s = IDENT -> s ] ->
+ VernacLoad (verbosely, s)
+ | IDENT "Declare"; IDENT "ML"; IDENT "Module"; l = LIST1 ne_string ->
+ VernacDeclareMLModule l
+
+ | IDENT "Locate"; l = locatable -> VernacLocate l
+
+ (* Managing load paths *)
+ | IDENT "Add"; IDENT "LoadPath"; dir = ne_string; alias = as_dirpath ->
+ VernacAddLoadPath (false, dir, alias)
+ | IDENT "Add"; IDENT "Rec"; IDENT "LoadPath"; dir = ne_string;
+ alias = as_dirpath -> VernacAddLoadPath (true, dir, alias)
+ | IDENT "Remove"; IDENT "LoadPath"; dir = ne_string ->
+ VernacRemoveLoadPath dir
+
+ (* For compatibility *)
+ | IDENT "AddPath"; dir = ne_string; "as"; alias = as_dirpath ->
+ VernacAddLoadPath (false, dir, alias)
+ | IDENT "AddRecPath"; dir = ne_string; "as"; alias = as_dirpath ->
+ VernacAddLoadPath (true, dir, alias)
+ | IDENT "DelPath"; dir = ne_string ->
+ VernacRemoveLoadPath dir
+
+ (* Type-Checking (pas dans le refman) *)
+ | "Type"; c = lconstr -> VernacGlobalCheck c
+
+ (* Printing (careful factorization of entries) *)
+ | IDENT "Print"; p = printable -> VernacPrint p
+ | IDENT "Print"; qid = smart_global; l = OPT univ_name_list -> VernacPrint (PrintName (qid,l))
+ | IDENT "Print"; IDENT "Module"; "Type"; qid = global ->
+ VernacPrint (PrintModuleType qid)
+ | IDENT "Print"; IDENT "Module"; qid = global ->
+ VernacPrint (PrintModule qid)
+ | IDENT "Print"; IDENT "Namespace" ; ns = dirpath ->
+ VernacPrint (PrintNamespace ns)
+ | IDENT "Inspect"; n = natural -> VernacPrint (PrintInspect n)
+
+ | IDENT "Add"; IDENT "ML"; IDENT "Path"; dir = ne_string ->
+ VernacAddMLPath (false, dir)
+ | IDENT "Add"; IDENT "Rec"; IDENT "ML"; IDENT "Path"; dir = ne_string ->
+ VernacAddMLPath (true, dir)
+
+ (* For acting on parameter tables *)
+ | "Set"; table = option_table; v = option_value ->
+ VernacSetOption (false, table, v)
+ | IDENT "Unset"; table = option_table ->
+ VernacUnsetOption (false, table)
+
+ | IDENT "Print"; IDENT "Table"; table = option_table ->
+ VernacPrintOption table
+
+ | IDENT "Add"; table = IDENT; field = IDENT; v = LIST1 option_ref_value
+ -> VernacAddOption ([table;field], v)
+ (* A global value below will be hidden by a field above! *)
+ (* In fact, we give priority to secondary tables *)
+ (* No syntax for tertiary tables due to conflict *)
+ (* (but they are unused anyway) *)
+ | IDENT "Add"; table = IDENT; v = LIST1 option_ref_value ->
+ VernacAddOption ([table], v)
+
+ | IDENT "Test"; table = option_table; "for"; v = LIST1 option_ref_value
+ -> VernacMemOption (table, v)
+ | IDENT "Test"; table = option_table ->
+ VernacPrintOption table
+
+ | IDENT "Remove"; table = IDENT; field = IDENT; v= LIST1 option_ref_value
+ -> VernacRemoveOption ([table;field], v)
+ | IDENT "Remove"; table = IDENT; v = LIST1 option_ref_value ->
+ VernacRemoveOption ([table], v) ]]
+ ;
+ query_command: (* TODO: rapprocher Eval et Check *)
+ [ [ IDENT "Eval"; r = red_expr; "in"; c = lconstr; "." ->
+ fun g -> VernacCheckMayEval (Some r, g, c)
+ | IDENT "Compute"; c = lconstr; "." ->
+ fun g -> VernacCheckMayEval (Some (Genredexpr.CbvVm None), g, c)
+ | IDENT "Check"; c = lconstr; "." ->
+ fun g -> VernacCheckMayEval (None, g, c)
+ (* Searching the environment *)
+ | IDENT "About"; qid = smart_global; l = OPT univ_name_list; "." ->
+ fun g -> VernacPrint (PrintAbout (qid,l,g))
+ | IDENT "SearchHead"; c = constr_pattern; l = in_or_out_modules; "." ->
+ fun g -> VernacSearch (SearchHead c,g, l)
+ | IDENT "SearchPattern"; c = constr_pattern; l = in_or_out_modules; "." ->
+ fun g -> VernacSearch (SearchPattern c,g, l)
+ | IDENT "SearchRewrite"; c = constr_pattern; l = in_or_out_modules; "." ->
+ fun g -> VernacSearch (SearchRewrite c,g, l)
+ | IDENT "Search"; s = searchabout_query; l = searchabout_queries; "." ->
+ let (sl,m) = l in fun g -> VernacSearch (SearchAbout (s::sl),g, m)
+ (* compatibility: SearchAbout *)
+ | IDENT "SearchAbout"; s = searchabout_query; l = searchabout_queries; "." ->
+ fun g -> let (sl,m) = l in VernacSearch (SearchAbout (s::sl),g, m)
+ (* compatibility: SearchAbout with "[ ... ]" *)
+ | IDENT "SearchAbout"; "["; sl = LIST1 searchabout_query; "]";
+ l = in_or_out_modules; "." ->
+ fun g -> VernacSearch (SearchAbout sl,g, l)
+ ] ]
+ ;
+ printable:
+ [ [ IDENT "Term"; qid = smart_global; l = OPT univ_name_list -> PrintName (qid,l)
+ | IDENT "All" -> PrintFullContext
+ | IDENT "Section"; s = global -> PrintSectionContext s
+ | IDENT "Grammar"; ent = IDENT ->
+ (* This should be in "syntax" section but is here for factorization*)
+ PrintGrammar ent
+ | IDENT "LoadPath"; dir = OPT dirpath -> PrintLoadPath dir
+ | IDENT "Modules" ->
+ user_err Pp.(str "Print Modules is obsolete; use Print Libraries instead")
+ | IDENT "Libraries" -> PrintModules
+
+ | IDENT "ML"; IDENT "Path" -> PrintMLLoadPath
+ | IDENT "ML"; IDENT "Modules" -> PrintMLModules
+ | IDENT "Debug"; IDENT "GC" -> PrintDebugGC
+ | IDENT "Graph" -> PrintGraph
+ | IDENT "Classes" -> PrintClasses
+ | IDENT "TypeClasses" -> PrintTypeClasses
+ | IDENT "Instances"; qid = smart_global -> PrintInstances qid
+ | IDENT "Coercions" -> PrintCoercions
+ | IDENT "Coercion"; IDENT "Paths"; s = class_rawexpr; t = class_rawexpr
+ -> PrintCoercionPaths (s,t)
+ | IDENT "Canonical"; IDENT "Projections" -> PrintCanonicalConversions
+ | IDENT "Tables" -> PrintTables
+ | IDENT "Options" -> PrintTables (* A Synonymous to Tables *)
+ | IDENT "Hint" -> PrintHintGoal
+ | IDENT "Hint"; qid = smart_global -> PrintHint qid
+ | IDENT "Hint"; "*" -> PrintHintDb
+ | IDENT "HintDb"; s = IDENT -> PrintHintDbName s
+ | IDENT "Scopes" -> PrintScopes
+ | IDENT "Scope"; s = IDENT -> PrintScope s
+ | IDENT "Visibility"; s = OPT [x = IDENT -> x ] -> PrintVisibility s
+ | IDENT "Implicit"; qid = smart_global -> PrintImplicit qid
+ | IDENT "Universes"; fopt = OPT ne_string -> PrintUniverses (false, fopt)
+ | IDENT "Sorted"; IDENT "Universes"; fopt = OPT ne_string -> PrintUniverses (true, fopt)
+ | IDENT "Assumptions"; qid = smart_global -> PrintAssumptions (false, false, qid)
+ | IDENT "Opaque"; IDENT "Dependencies"; qid = smart_global -> PrintAssumptions (true, false, qid)
+ | IDENT "Transparent"; IDENT "Dependencies"; qid = smart_global -> PrintAssumptions (false, true, qid)
+ | IDENT "All"; IDENT "Dependencies"; qid = smart_global -> PrintAssumptions (true, true, qid)
+ | IDENT "Strategy"; qid = smart_global -> PrintStrategy (Some qid)
+ | IDENT "Strategies" -> PrintStrategy None ] ]
+ ;
+ class_rawexpr:
+ [ [ IDENT "Funclass" -> FunClass
+ | IDENT "Sortclass" -> SortClass
+ | qid = smart_global -> RefClass qid ] ]
+ ;
+ locatable:
+ [ [ qid = smart_global -> LocateAny qid
+ | IDENT "Term"; qid = smart_global -> LocateTerm qid
+ | IDENT "File"; f = ne_string -> LocateFile f
+ | IDENT "Library"; qid = global -> LocateLibrary qid
+ | IDENT "Module"; qid = global -> LocateModule qid ] ]
+ ;
+ option_value:
+ [ [ -> BoolValue true
+ | n = integer -> IntValue (Some n)
+ | s = STRING -> StringValue s ] ]
+ ;
+ option_ref_value:
+ [ [ id = global -> QualidRefValue id
+ | s = STRING -> StringRefValue s ] ]
+ ;
+ option_table:
+ [ [ fl = LIST1 [ x = IDENT -> x ] -> fl ]]
+ ;
+ as_dirpath:
+ [ [ d = OPT [ "as"; d = dirpath -> d ] -> d ] ]
+ ;
+ ne_in_or_out_modules:
+ [ [ IDENT "inside"; l = LIST1 global -> SearchInside l
+ | IDENT "outside"; l = LIST1 global -> SearchOutside l ] ]
+ ;
+ in_or_out_modules:
+ [ [ m = ne_in_or_out_modules -> m
+ | -> SearchOutside [] ] ]
+ ;
+ comment:
+ [ [ c = constr -> CommentConstr c
+ | s = STRING -> CommentString s
+ | n = natural -> CommentInt n ] ]
+ ;
+ positive_search_mark:
+ [ [ "-" -> false | -> true ] ]
+ ;
+ scope:
+ [ [ "%"; key = IDENT -> key ] ]
+ ;
+ searchabout_query:
+ [ [ b = positive_search_mark; s = ne_string; sc = OPT scope ->
+ (b, SearchString (s,sc))
+ | b = positive_search_mark; p = constr_pattern ->
+ (b, SearchSubPattern p)
+ ] ]
+ ;
+ searchabout_queries:
+ [ [ m = ne_in_or_out_modules -> ([],m)
+ | s = searchabout_query; l = searchabout_queries ->
+ let (sl,m) = l in (s::sl,m)
+ | -> ([],SearchOutside [])
+ ] ]
+ ;
+ univ_name_list:
+ [ [ "@{" ; l = LIST0 name; "}" -> l ] ]
+ ;
+END;
+
+GEXTEND Gram
+ command:
+ [ [
+(* State management *)
+ IDENT "Write"; IDENT "State"; s = IDENT -> VernacWriteState s
+ | IDENT "Write"; IDENT "State"; s = ne_string -> VernacWriteState s
+ | IDENT "Restore"; IDENT "State"; s = IDENT -> VernacRestoreState s
+ | IDENT "Restore"; IDENT "State"; s = ne_string -> VernacRestoreState s
+
+(* Resetting *)
+ | IDENT "Reset"; IDENT "Initial" -> VernacResetInitial
+ | IDENT "Reset"; id = identref -> VernacResetName id
+ | IDENT "Back" -> VernacBack 1
+ | IDENT "Back"; n = natural -> VernacBack n
+ | IDENT "BackTo"; n = natural -> VernacBackTo n
+
+(* Tactic Debugger *)
+ | IDENT "Debug"; IDENT "On" ->
+ VernacSetOption (false, ["Ltac";"Debug"], BoolValue true)
+
+ | IDENT "Debug"; IDENT "Off" ->
+ VernacSetOption (false, ["Ltac";"Debug"], BoolValue false)
+
+(* registration of a custom reduction *)
+
+ | IDENT "Declare"; IDENT "Reduction"; s = IDENT; ":=";
+ r = red_expr ->
+ VernacDeclareReduction (s,r)
+
+ ] ];
+ END
+;;
+
+(* Grammar extensions *)
+
+GEXTEND Gram
+ GLOBAL: syntax;
+
+ syntax:
+ [ [ IDENT "Open"; IDENT "Scope"; sc = IDENT ->
+ VernacOpenCloseScope (true,sc)
+
+ | IDENT "Close"; IDENT "Scope"; sc = IDENT ->
+ VernacOpenCloseScope (false,sc)
+
+ | IDENT "Delimit"; IDENT "Scope"; sc = IDENT; "with"; key = IDENT ->
+ VernacDelimiters (sc, Some key)
+ | IDENT "Undelimit"; IDENT "Scope"; sc = IDENT ->
+ VernacDelimiters (sc, None)
+
+ | IDENT "Bind"; IDENT "Scope"; sc = IDENT; "with";
+ refl = LIST1 class_rawexpr -> VernacBindScope (sc,refl)
+
+ | IDENT "Infix"; op = ne_lstring; ":="; p = constr;
+ modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ];
+ sc = OPT [ ":"; sc = IDENT -> sc ] ->
+ VernacInfix ((op,modl),p,sc)
+ | IDENT "Notation"; id = identref;
+ idl = LIST0 ident; ":="; c = constr; b = only_parsing ->
+ VernacSyntacticDefinition
+ (id,(idl,c),b)
+ | IDENT "Notation"; s = lstring; ":=";
+ c = constr;
+ modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ];
+ sc = OPT [ ":"; sc = IDENT -> sc ] ->
+ VernacNotation (c,(s,modl),sc)
+ | IDENT "Format"; IDENT "Notation"; n = STRING; s = STRING; fmt = STRING ->
+ VernacNotationAddFormat (n,s,fmt)
+
+ | IDENT "Reserved"; IDENT "Infix"; s = ne_lstring;
+ l = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ] ->
+ let s = CAst.map (fun s -> "x '"^s^"' y") s in
+ VernacSyntaxExtension (true,(s,l))
+
+ | IDENT "Reserved"; IDENT "Notation";
+ s = ne_lstring;
+ l = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ]
+ -> VernacSyntaxExtension (false, (s,l))
+
+ (* "Print" "Grammar" should be here but is in "command" entry in order
+ to factorize with other "Print"-based vernac entries *)
+ ] ]
+ ;
+ only_parsing:
+ [ [ "("; IDENT "only"; IDENT "parsing"; ")" ->
+ Some Flags.Current
+ | "("; IDENT "compat"; s = STRING; ")" ->
+ Some (parse_compat_version s)
+ | -> None ] ]
+ ;
+ level:
+ [ [ IDENT "level"; n = natural -> NumLevel n
+ | IDENT "next"; IDENT "level" -> NextLevel ] ]
+ ;
+ syntax_modifier:
+ [ [ "at"; IDENT "level"; n = natural -> SetLevel n
+ | IDENT "left"; IDENT "associativity" -> SetAssoc LeftA
+ | IDENT "right"; IDENT "associativity" -> SetAssoc RightA
+ | IDENT "no"; IDENT "associativity" -> SetAssoc NonA
+ | IDENT "only"; IDENT "printing" -> SetOnlyPrinting
+ | IDENT "only"; IDENT "parsing" -> SetOnlyParsing
+ | IDENT "compat"; s = STRING ->
+ SetCompatVersion (parse_compat_version s)
+ | IDENT "format"; s1 = [s = STRING -> CAst.make ~loc:!@loc s];
+ s2 = OPT [s = STRING -> CAst.make ~loc:!@loc s] ->
+ begin match s1, s2 with
+ | { CAst.v = k }, Some s -> SetFormat(k,s)
+ | s, None -> SetFormat ("text",s) end
+ | x = IDENT; ","; l = LIST1 [id = IDENT -> id ] SEP ","; "at";
+ lev = level -> SetItemLevel (x::l,lev)
+ | x = IDENT; "at"; lev = level -> SetItemLevel ([x],lev)
+ | x = IDENT; "at"; lev = level; b = constr_as_binder_kind -> SetItemLevelAsBinder ([x],b,Some lev)
+ | x = IDENT; b = constr_as_binder_kind -> SetItemLevelAsBinder ([x],b,None)
+ | x = IDENT; typ = syntax_extension_type -> SetEntryType (x,typ)
+ ] ]
+ ;
+ syntax_extension_type:
+ [ [ IDENT "ident" -> ETName | IDENT "global" -> ETReference
+ | IDENT "bigint" -> ETBigint
+ | IDENT "binder" -> ETBinder true
+ | IDENT "constr"; n = OPT at_level; b = constr_as_binder_kind -> ETConstrAsBinder (b,n)
+ | IDENT "pattern" -> ETPattern (false,None)
+ | IDENT "pattern"; "at"; IDENT "level"; n = natural -> ETPattern (false,Some n)
+ | IDENT "strict"; IDENT "pattern" -> ETPattern (true,None)
+ | IDENT "strict"; IDENT "pattern"; "at"; IDENT "level"; n = natural -> ETPattern (true,Some n)
+ | IDENT "closed"; IDENT "binder" -> ETBinder false
+ ] ]
+ ;
+ at_level:
+ [ [ "at"; n = level -> n ] ]
+ ;
+ constr_as_binder_kind:
+ [ [ "as"; IDENT "ident" -> Notation_term.AsIdent
+ | "as"; IDENT "pattern" -> Notation_term.AsIdentOrPattern
+ | "as"; IDENT "strict"; IDENT "pattern" -> Notation_term.AsStrictPattern ] ]
+ ;
+END
diff --git a/vernac/himsg.ml b/vernac/himsg.ml
index 698ee4703..5d671ef52 100644
--- a/vernac/himsg.ml
+++ b/vernac/himsg.ml
@@ -75,11 +75,7 @@ let rec contract3' env sigma a b c = function
| MetaOccurInBody _ | InstanceNotSameType _ | ProblemBeyondCapabilities
| UnifUnivInconsistency _ as x -> contract3 env sigma a b c, x
| CannotSolveConstraint ((pb,env',t,u),x) ->
- let t = EConstr.of_constr t in
- let u = EConstr.of_constr u in
let env',t,u = contract2 env' sigma t u in
- let t = EConstr.Unsafe.to_constr t in
- let u = EConstr.Unsafe.to_constr u in
let y,x = contract3' env sigma a b c x in
y,CannotSolveConstraint ((pb,env',t,u),x)
@@ -90,7 +86,7 @@ let j_nf_betaiotaevar env sigma j =
uj_type = Reductionops.nf_betaiota env sigma j.uj_type }
let jv_nf_betaiotaevar env sigma jl =
- Array.map (fun j -> j_nf_betaiotaevar env sigma j) jl
+ Array.Smart.map (fun j -> j_nf_betaiotaevar env sigma j) jl
(** Printers *)
@@ -201,7 +197,7 @@ let rec pr_disjunction pr = function
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"*)"
+ str"(*" ++ Univ.Instance.pr UnivNames.pr_with_global_universes u ++ str"*)"
else mt())
let explain_elim_arity env sigma ind sorts c pj okinds =
@@ -318,12 +314,10 @@ let explain_unification_error env sigma p1 p2 = function
| UnifUnivInconsistency p ->
if !Constrextern.print_universes then
[str "universe inconsistency: " ++
- Univ.explain_universe_inconsistency Universes.pr_with_global_universes p]
+ Univ.explain_universe_inconsistency UnivNames.pr_with_global_universes p]
else
[str "universe inconsistency"]
| CannotSolveConstraint ((pb,env,t,u),e) ->
- let t = EConstr.of_constr t in
- let u = EConstr.of_constr u in
let env = make_all_name_different env sigma in
(strbrk "cannot satisfy constraint " ++ pr_leconstr_env env sigma t ++
str " == " ++ pr_leconstr_env env sigma u)
@@ -562,9 +556,9 @@ let rec explain_evar_kind env sigma evk ty = function
| Evar_kinds.SubEvar (where,evk') ->
let evi = Evd.find sigma evk' in
let pc = match evi.evar_body with
- | Evar_defined c -> pr_leconstr_env env sigma (EConstr.of_constr c)
+ | Evar_defined c -> pr_leconstr_env env sigma c
| Evar_empty -> assert false in
- let ty' = EConstr.of_constr evi.evar_concl in
+ let ty' = evi.evar_concl in
(match where with
| Some Evar_kinds.Body -> str "the body of "
| Some Evar_kinds.Domain -> str "the domain of "
@@ -577,11 +571,11 @@ let rec explain_evar_kind env sigma evk ty = function
(pr_leconstr_env env sigma ty') (snd evi.evar_source)
let explain_typeclass_resolution env sigma evi k =
- match Typeclasses.class_of_constr sigma (EConstr.of_constr evi.evar_concl) with
+ match Typeclasses.class_of_constr sigma 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_leconstr_env env sigma evi.evar_concl ++
pr_trailing_ne_context_of env sigma
| _ -> mt()
@@ -590,14 +584,14 @@ let explain_placeholder_kind env sigma c e =
| Some (SeveralInstancesFound n) ->
strbrk " (several distinct possible type class instances found)"
| None ->
- match Typeclasses.class_of_constr sigma (EConstr.of_constr c) with
+ match Typeclasses.class_of_constr sigma 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 type_of_hole = pr_leconstr_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) ++
@@ -640,8 +634,7 @@ let explain_refiner_cannot_generalize env sigma ty =
pr_leconstr_env env sigma ty ++ str "."
let explain_no_occurrence_found env sigma c id =
- let c = EConstr.to_constr sigma c in
- str "Found no subterm matching " ++ pr_lconstr_env env sigma c ++
+ str "Found no subterm matching " ++ pr_leconstr_env env sigma c ++
str " in " ++
(match id with
| Some id -> Id.print id
@@ -766,7 +759,7 @@ let pr_constraints printenv env sigma evars cstrs =
let evs =
prlist
(fun (ev, evi) -> fnl () ++ pr_existential_key sigma ev ++
- str " : " ++ pr_lconstr_env env' sigma evi.evar_concl ++ fnl ()) l
+ str " : " ++ pr_leconstr_env env' sigma evi.evar_concl ++ fnl ()) l
in
h 0 (pe ++ evs ++ pr_evar_constraints sigma cstrs)
else
@@ -855,9 +848,9 @@ let explain_not_match_error = function
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 () ++
+ quote (Printer.safe_pr_lconstr_env env (Evd.from_env env) typ2) ++ spc () ++
str "but found type" ++ spc () ++
- quote (Printer.safe_pr_lconstr_env env Evd.empty typ1)
+ quote (Printer.safe_pr_lconstr_env env (Evd.from_env env) typ1)
| NotSameConstructorNamesField ->
str "constructor names differ"
| NotSameInductiveNameInBlockField ->
@@ -893,12 +886,12 @@ let explain_not_match_error = function
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
+ Univ.explain_universe_inconsistency UnivNames.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 () ++
+ quote (Printer.safe_pr_lconstr_env env (Evd.from_env env) t1) ++ spc () ++
str "compared to " ++ spc () ++
- quote (Printer.safe_pr_lconstr_env env Evd.empty t2)
+ quote (Printer.safe_pr_lconstr_env env (Evd.from_env env) t2)
| IncompatibleConstraints cst ->
str " the expected (polymorphic) constraints do not imply " ++
let cst = Univ.AUContext.instantiate (Univ.AUContext.instance cst) cst in
@@ -1018,8 +1011,9 @@ let explain_module_internalization_error = function
(* Typeclass errors *)
let explain_not_a_class env c =
- let c = EConstr.to_constr Evd.empty c in
- pr_constr_env env Evd.empty c ++ str" is not a declared type class."
+ let sigma = Evd.from_env env in
+ let c = EConstr.to_constr sigma c in
+ pr_constr_env env sigma c ++ str" is not a declared type class."
let explain_unbound_method env cid { CAst.v = id } =
str "Unbound method name " ++ Id.print (id) ++ spc () ++
@@ -1032,14 +1026,13 @@ let pr_constr_exprs exprs =
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) ++
+ hov 1 (str"Expected:" ++ brk (1, 1) ++ pr_rel_context env (Evd.from_env env) 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 *)
@@ -1094,19 +1087,19 @@ let explain_refiner_error env sigma = function
(* 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
+ let pc = pr_lconstr_env env (Evd.from_env env) c in
+ let pv = pr_lconstr_env env (Evd.from_env env) 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
+ let pc = pr_lconstr_env env (Evd.from_env env) c in
+ let pv = pr_lconstr_env env (Evd.from_env env) 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 pv = pr_lconstr_env env (Evd.from_env env) v in
let atomic = Int.equal (nb_prod Evd.empty (EConstr.of_constr c)) (** FIXME *) 0 in
str "The type of constructor" ++ brk(1,1) ++ Id.print id ++ brk(1,1) ++
str "is not valid;" ++ brk(1,1) ++
@@ -1126,12 +1119,12 @@ let error_ill_formed_constructor env id c v nparams nargs =
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 (EConstr.of_constr c))
+ quote (pr_goal_concl_style_env env (Evd.from_env env) (EConstr.of_constr 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
+ let pv1 = pr_lconstr_env env (Evd.from_env env) v1 in
+ let pv2 = pr_lconstr_env env (Evd.from_env env) v2 in
str "Last occurrence of " ++ pv2 ++ str " must have " ++ pv1 ++
str " as " ++ pr_nth n ++ str " argument in" ++ brk(1,1) ++ pc ++ str "."
@@ -1149,7 +1142,7 @@ let error_same_names_overlap idl =
prlist_with_sep pr_comma Id.print idl ++ str "."
let error_not_an_arity env c =
- str "The type" ++ spc () ++ pr_lconstr_env env Evd.empty c ++ spc () ++
+ str "The type" ++ spc () ++ pr_lconstr_env env (Evd.from_env env) c ++ spc () ++
str "is not an arity."
let error_bad_entry () =
@@ -1323,4 +1316,4 @@ let explain_reduction_tactic_error = function
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
+ explain_type_error env' (Evd.from_env env') e
diff --git a/vernac/himsg.mli b/vernac/himsg.mli
index 0e20d18c6..1d3807502 100644
--- a/vernac/himsg.mli
+++ b/vernac/himsg.mli
@@ -25,6 +25,8 @@ val explain_pretype_error : env -> Evd.evar_map -> pretype_error -> Pp.t
val explain_inductive_error : inductive_error -> Pp.t
+val explain_mismatched_contexts : env -> contexts -> Constrexpr.constr_expr list -> Context.Rel.t -> Pp.t
+
val explain_typeclass_error : env -> typeclass_error -> Pp.t
val explain_recursion_scheme_error : recursion_scheme_error -> Pp.t
diff --git a/vernac/indschemes.ml b/vernac/indschemes.ml
index 32885ab88..2deca1e06 100644
--- a/vernac/indschemes.ml
+++ b/vernac/indschemes.ml
@@ -380,7 +380,7 @@ let do_mutual_induction_scheme lnamedepindsort =
match inst with
| None ->
let _, ctx = Global.type_of_global_in_context env0 (IndRef ind) in
- let u, ctx = Universes.fresh_instance_from ctx None in
+ let u, ctx = UnivGen.fresh_instance_from ctx None in
let evd = Evd.from_ctx (UState.of_context_set ctx) in
evd, (ind,u), Some u
| Some ui -> evd, (ind, ui), inst
diff --git a/vernac/indschemes.mli b/vernac/indschemes.mli
index bd4249cac..261c3d813 100644
--- a/vernac/indschemes.mli
+++ b/vernac/indschemes.mli
@@ -32,17 +32,17 @@ val declare_rewriting_schemes : inductive -> unit
(** Mutual Minimality/Induction scheme *)
val do_mutual_induction_scheme :
- (Misctypes.lident * bool * inductive * Sorts.family) list -> unit
+ (lident * bool * inductive * Sorts.family) list -> unit
(** Main calls to interpret the Scheme command *)
-val do_scheme : (Misctypes.lident option * scheme) list -> unit
+val do_scheme : (lident option * scheme) list -> unit
(** Combine a list of schemes into a conjunction of them *)
val build_combined_scheme : env -> Constant.t list -> Evd.evar_map * constr * types
-val do_combined_scheme : Misctypes.lident -> Misctypes.lident list -> unit
+val do_combined_scheme : lident -> lident list -> unit
(** Hook called at each inductive type definition *)
diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml
index 30dd6ec74..ce74f2344 100644
--- a/vernac/lemmas.ml
+++ b/vernac/lemmas.ml
@@ -34,7 +34,7 @@ open Impargs
module RelDecl = Context.Rel.Declaration
module NamedDecl = Context.Named.Declaration
-type 'a declaration_hook = Decl_kinds.locality -> Globnames.global_reference -> 'a
+type 'a declaration_hook = Decl_kinds.locality -> GlobRef.t -> 'a
let mk_hook hook = hook
let call_hook fix_exn hook l c =
try hook l c
@@ -334,8 +334,8 @@ let universe_proof_terminator compute_guard hook =
Feedback.feedback Feedback.AddedAxiom
| Proved (opaque,idopt,proof) ->
let is_opaque, export_seff = match opaque with
- | Vernacexpr.Transparent -> false, true
- | Vernacexpr.Opaque -> true, false
+ | Transparent -> false, true
+ | Opaque -> true, false
in
let proof = get_proof proof compute_guard
(hook (Some (proof.Proof_global.universes))) is_opaque in
@@ -436,7 +436,7 @@ let start_proof_with_initialization kind sigma decl recguard thms snl hook =
let start_proof_com ?inference_hook kind thms hook =
let env0 = Global.env () in
let decl = fst (List.hd thms) in
- let evd, decl = Univdecls.interp_univ_decl_opt env0 (snd decl) in
+ let evd, decl = Constrexpr_ops.interp_univ_decl_opt env0 (snd decl) in
let evd, thms = List.fold_left_map (fun evd ((id, _), (bl, t)) ->
let evd, (impls, ((env, ctx), imps)) = interp_context_evars env0 evd bl in
let evd, (t', imps') = interp_type_evars_impls ~impls env evd t in
@@ -451,12 +451,12 @@ let start_proof_com ?inference_hook kind thms hook =
(ids, imps @ lift_implicits (Context.Rel.nhyps ctx) imps'))))
evd thms in
let recguard,thms,snl = look_for_possibly_mutual_statements evd thms in
- let evd, _nf = Evarutil.nf_evars_and_universes evd in
+ let evd = Evd.minimize_universes evd in
(* XXX: This nf_evar is critical too!! We are normalizing twice if
you look at the previous lines... *)
let thms = List.map (fun (n, (t, info)) -> (n, (nf_evar evd t, info))) thms in
let () =
- let open Misctypes in
+ let open UState in
if not (decl.univdecl_extensible_instance && decl.univdecl_extensible_constraints) then
ignore (Evd.check_univ_decl ~poly:(pi2 kind) evd decl)
in
diff --git a/vernac/lemmas.mli b/vernac/lemmas.mli
index ad4c278e0..c9e4876ee 100644
--- a/vernac/lemmas.mli
+++ b/vernac/lemmas.mli
@@ -13,21 +13,21 @@ open Decl_kinds
type 'a declaration_hook
val mk_hook :
- (Decl_kinds.locality -> Globnames.global_reference -> 'a) -> 'a declaration_hook
+ (Decl_kinds.locality -> GlobRef.t -> 'a) -> 'a declaration_hook
val call_hook :
- Future.fix_exn -> 'a declaration_hook -> Decl_kinds.locality -> Globnames.global_reference -> 'a
+ Future.fix_exn -> 'a declaration_hook -> Decl_kinds.locality -> GlobRef.t -> 'a
(** A hook start_proof calls on the type of the definition being started *)
val set_start_hook : (EConstr.types -> unit) -> unit
-val start_proof : Id.t -> ?pl:Univdecls.universe_decl -> goal_kind -> Evd.evar_map ->
+val start_proof : Id.t -> ?pl:UState.universe_decl -> goal_kind -> Evd.evar_map ->
?terminator:(Proof_global.lemma_possible_guards -> unit declaration_hook -> Proof_global.proof_terminator) ->
?sign:Environ.named_context_val -> EConstr.types ->
?init_tac:unit Proofview.tactic -> ?compute_guard:Proof_global.lemma_possible_guards ->
unit declaration_hook -> unit
-val start_proof_univs : Id.t -> ?pl:Univdecls.universe_decl -> goal_kind -> Evd.evar_map ->
+val start_proof_univs : Id.t -> ?pl:UState.universe_decl -> goal_kind -> Evd.evar_map ->
?terminator:(Proof_global.lemma_possible_guards -> (UState.t option -> unit declaration_hook) -> Proof_global.proof_terminator) ->
?sign:Environ.named_context_val -> EConstr.types ->
?init_tac:unit Proofview.tactic -> ?compute_guard:Proof_global.lemma_possible_guards ->
@@ -39,7 +39,7 @@ val start_proof_com :
unit declaration_hook -> unit
val start_proof_with_initialization :
- goal_kind -> Evd.evar_map -> Univdecls.universe_decl ->
+ goal_kind -> Evd.evar_map -> UState.universe_decl ->
(bool * Proof_global.lemma_possible_guards * unit Proofview.tactic list option) option ->
(Id.t (* name of thm *) *
(EConstr.types (* type of thm *) * (Name.t list (* names to pre-introduce *) * Impargs.manual_explicitation list))) list
diff --git a/vernac/metasyntax.ml b/vernac/metasyntax.ml
index feeca6075..8f64f5519 100644
--- a/vernac/metasyntax.ml
+++ b/vernac/metasyntax.ml
@@ -15,6 +15,7 @@ open Names
open Constrexpr
open Constrexpr_ops
open Notation_term
+open Notation_gram
open Notation_ops
open Ppextend
open Extend
@@ -76,22 +77,22 @@ let pr_grammar = function
pr_entry Pcoq.Constr.pattern
| "vernac" ->
str "Entry vernac_control is" ++ fnl () ++
- pr_entry Pcoq.Vernac_.vernac_control ++
+ pr_entry Pvernac.Vernac_.vernac_control ++
str "Entry command is" ++ fnl () ++
- pr_entry Pcoq.Vernac_.command ++
+ pr_entry Pvernac.Vernac_.command ++
str "Entry syntax is" ++ fnl () ++
- pr_entry Pcoq.Vernac_.syntax ++
+ pr_entry Pvernac.Vernac_.syntax ++
str "Entry gallina is" ++ fnl () ++
- pr_entry Pcoq.Vernac_.gallina ++
+ pr_entry Pvernac.Vernac_.gallina ++
str "Entry gallina_ext is" ++ fnl () ++
- pr_entry Pcoq.Vernac_.gallina_ext
+ pr_entry Pvernac.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 ({CAst.loc;v=str} : Misctypes.lstring) =
+let parse_format ({CAst.loc;v=str} : lstring) =
let len = String.length str in
(* TODO: update the line of the location when the string contains newlines *)
let make_loc i j = Option.map (Loc.shift_loc (i+1) (j-len)) loc in
@@ -709,7 +710,7 @@ let error_parsing_incompatible_level ntn ntn' oldprec prec =
pr_level ntn prec ++ str ".")
type syntax_extension = {
- synext_level : Notation_term.level;
+ synext_level : Notation_gram.level;
synext_notation : notation;
synext_notgram : notation_grammar;
synext_unparsing : unparsing list;
@@ -728,8 +729,8 @@ let check_and_extend_constr_grammar ntn rule =
let ntn_for_grammar = rule.notgram_notation in
if String.equal ntn ntn_for_grammar then raise Not_found;
let prec = rule.notgram_level in
- let oldprec = Notation.level_of_notation ntn_for_grammar in
- if not (Notation.level_eq prec oldprec) then error_parsing_incompatible_level ntn ntn_for_grammar oldprec prec;
+ let oldprec = Notgram_ops.level_of_notation ntn_for_grammar in
+ if not (Notgram_ops.level_eq prec oldprec) then error_parsing_incompatible_level ntn ntn_for_grammar oldprec prec;
with Not_found ->
Egramcoq.extend_constr_grammar rule
@@ -738,16 +739,16 @@ let cache_one_syntax_extension se =
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;
+ let oldprec = Notgram_ops.level_of_notation ~onlyprint ntn in
+ if not (Notgram_ops.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;
+ Notgram_ops.declare_notation_level ntn prec ~onlyprint;
(* Declare the parsing rule *)
if not onlyprint then List.iter (check_and_extend_constr_grammar ntn) se.synext_notgram.notgram_rules;
(* Declare the notation rule *)
- Notation.declare_notation_rule ntn
+ declare_notation_rule ntn
~extra:se.synext_extra (se.synext_unparsing, pi1 prec) se.synext_notgram
end
@@ -791,7 +792,7 @@ type notation_modifier = {
only_parsing : bool;
only_printing : bool;
compat : Flags.compat_version option;
- format : Misctypes.lstring option;
+ format : lstring option;
extra : (string * string) list;
}
@@ -1061,7 +1062,7 @@ let find_precedence lev etyps symbols onlyprint =
[],Option.get lev
let check_curly_brackets_notation_exists () =
- try let _ = Notation.level_of_notation "{ _ }" in ()
+ try let _ = Notgram_ops.level_of_notation "{ _ }" in ()
with Not_found ->
user_err Pp.(str "Notations involving patterns of the form \"{ _ }\" are treated \n\
specially and require that the notation \"{ _ }\" is already reserved.")
@@ -1103,7 +1104,7 @@ module SynData = struct
only_parsing : bool;
only_printing : bool;
compat : Flags.compat_version option;
- format : Misctypes.lstring option;
+ format : lstring option;
extra : (string * string) list;
(* XXX: Callback to printing, must remove *)
@@ -1274,10 +1275,10 @@ exception NoSyntaxRule
let recover_notation_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
+ let prec = Notgram_ops.level_of_notation ~onlyprint:true ntn (* Be as little restrictive as possible *) in
+ let pp_rule,_ = find_notation_printing_rule ntn in
+ let pp_extra_rules = find_notation_extra_printing_rules ntn in
+ let pa_rule = find_notation_parsing_rules ntn in
{ synext_level = prec;
synext_notation = ntn;
synext_notgram = pa_rule;
@@ -1444,7 +1445,7 @@ let add_notation_extra_printing_rule df k v =
let notk =
let _,_, symbs = analyze_notation_tokens ~onlyprint:true df in
make_notation_key symbs in
- Notation.add_notation_extra_printing_rule notk k v
+ add_notation_extra_printing_rule notk k v
(* Infix notations *)
diff --git a/vernac/metasyntax.mli b/vernac/metasyntax.mli
index a6c12e089..f6de75b07 100644
--- a/vernac/metasyntax.mli
+++ b/vernac/metasyntax.mli
@@ -14,7 +14,6 @@ open Notation
open Constrexpr
open Notation_term
open Environ
-open Misctypes
val add_token_obj : string -> unit
diff --git a/vernac/misctypes.ml b/vernac/misctypes.ml
new file mode 100644
index 000000000..ae725efaa
--- /dev/null
+++ b/vernac/misctypes.ml
@@ -0,0 +1,75 @@
+(* Compat module, to be removed in 8.10 *)
+open Names
+
+type lident = Names.lident
+[@@ocaml.deprecated "use [Names.lident"]
+type lname = Names.lname
+[@@ocaml.deprecated "use [Names.lname]"]
+type lstring = Names.lstring
+[@@ocaml.deprecated "use [Names.lstring]"]
+
+type 'a or_by_notation_r = 'a Constrexpr.or_by_notation_r =
+ | AN of 'a [@ocaml.deprecated "use version in [Constrexpr]"]
+ | ByNotation of (string * string option) [@ocaml.deprecated "use version in [Constrexpr]"]
+[@@ocaml.deprecated "use [Constrexpr.or_by_notation_r]"]
+
+type 'a or_by_notation = 'a Constrexpr.or_by_notation
+[@@ocaml.deprecated "use [Constrexpr.or_by_notation]"]
+
+type intro_pattern_naming_expr = Namegen.intro_pattern_naming_expr =
+ | IntroIdentifier of Id.t [@ocaml.deprecated "Use version in [Evarutil]"]
+ | IntroFresh of Id.t [@ocaml.deprecated "Use version in [Evarutil]"]
+ | IntroAnonymous [@ocaml.deprecated "Use version in [Evarutil]"]
+[@@ocaml.deprecated "use [Evarutil.intro_pattern_naming_expr]"]
+
+type 'a or_var = 'a Locus.or_var =
+ | ArgArg of 'a [@ocaml.deprecated "Use version in [Locus]"]
+ | ArgVar of Names.lident [@ocaml.deprecated "Use version in [Locus]"]
+[@@ocaml.deprecated "use [Locus.or_var]"]
+
+type quantified_hypothesis = Tactypes.quantified_hypothesis =
+ AnonHyp of int [@ocaml.deprecated "Use version in [Tactypes]"]
+ | NamedHyp of Id.t [@ocaml.deprecated "Use version in [Tactypes]"]
+[@@ocaml.deprecated "use [Tactypes.quantified_hypothesis]"]
+
+type multi = Equality.multi =
+ | Precisely of int [@ocaml.deprecated "use version in [Equality]"]
+ | UpTo of int [@ocaml.deprecated "use version in [Equality]"]
+ | RepeatStar [@ocaml.deprecated "use version in [Equality]"]
+ | RepeatPlus [@ocaml.deprecated "use version in [Equality]"]
+[@@ocaml.deprecated "use [Equality.multi]"]
+
+type 'a bindings = 'a Tactypes.bindings =
+ | ImplicitBindings of 'a list [@ocaml.deprecated "use version in [Tactypes]"]
+ | ExplicitBindings of 'a Tactypes.explicit_bindings [@ocaml.deprecated "use version in [Tactypes]"]
+ | NoBindings [@ocaml.deprecated "use version in [Tactypes]"]
+[@@ocaml.deprecated "use [Tactypes.bindings]"]
+
+type 'constr intro_pattern_expr = 'constr Tactypes.intro_pattern_expr =
+ | IntroForthcoming of bool [@ocaml.deprecated "use version in [Tactypes]"]
+ | IntroNaming of Namegen.intro_pattern_naming_expr [@ocaml.deprecated "use version in [Tactypes]"]
+ | IntroAction of 'constr Tactypes.intro_pattern_action_expr [@ocaml.deprecated "use version in [Tactypes]"]
+and 'constr intro_pattern_action_expr = 'constr Tactypes.intro_pattern_action_expr =
+ | IntroWildcard [@ocaml.deprecated "use [Tactypes]"]
+ | IntroOrAndPattern of 'constr Tactypes.or_and_intro_pattern_expr [@ocaml.deprecated "use [Tactypes]"]
+ | IntroInjection of ('constr intro_pattern_expr) CAst.t list [@ocaml.deprecated "use [Tactypes]"]
+ | IntroApplyOn of 'constr CAst.t * 'constr intro_pattern_expr CAst.t [@ocaml.deprecated "use [Tactypes]"]
+ | IntroRewrite of bool [@ocaml.deprecated "use [Tactypes]"]
+and 'constr or_and_intro_pattern_expr = 'constr Tactypes.or_and_intro_pattern_expr =
+ | IntroOrPattern of ('constr intro_pattern_expr) CAst.t list list [@ocaml.deprecated "use [Tactypes]"]
+ | IntroAndPattern of ('constr intro_pattern_expr) CAst.t list [@ocaml.deprecated "use [Tactypes]"]
+[@@ocaml.deprecated "use version in [Tactypes]"]
+
+type 'id move_location = 'id Logic.move_location =
+ | MoveAfter of 'id [@ocaml.deprecated "use version in [Logic]"]
+ | MoveBefore of 'id [@ocaml.deprecated "use version in [Logic]"]
+ | MoveFirst [@ocaml.deprecated "use version in [Logic]"]
+ | MoveLast [@ocaml.deprecated "use version in [Logic]"]
+[@@ocaml.deprecated "use version in [Logic]"]
+
+type 'a cast_type = 'a Glob_term.cast_type =
+ | CastConv of 'a [@ocaml.deprecated "use version in [Glob_term]"]
+ | CastVM of 'a [@ocaml.deprecated "use version in [Glob_term]"]
+ | CastCoerce [@ocaml.deprecated "use version in [Glob_term]"]
+ | CastNative of 'a [@ocaml.deprecated "use version in [Glob_term]"]
+[@@ocaml.deprecated "use version in [Glob_term]"]
diff --git a/vernac/mltop.ml b/vernac/mltop.ml
index 343b0925d..d25dea141 100644
--- a/vernac/mltop.ml
+++ b/vernac/mltop.ml
@@ -345,13 +345,6 @@ let load_ml_object mname ?path fname=
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 *)
@@ -396,8 +389,6 @@ let trigger_ml_object verb cache reinit ?path name =
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
diff --git a/vernac/mltop.mli b/vernac/mltop.mli
index da195f4fc..ed1f9a12d 100644
--- a/vernac/mltop.mli
+++ b/vernac/mltop.mli
@@ -68,9 +68,6 @@ val add_coq_path : coq_path -> 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} *)
diff --git a/vernac/obligations.ml b/vernac/obligations.ml
index 064e40b9b..1ab24b670 100644
--- a/vernac/obligations.ml
+++ b/vernac/obligations.ml
@@ -209,8 +209,10 @@ let eterm_obligations env name evm fs ?status t ty =
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 hyps = trunc_named_context nc_len hyps in
+ let hyps = EConstr.Unsafe.to_named_context hyps in
+ let concl = EConstr.Unsafe.to_constr ev.evar_concl in
+ let evtyp, deps, transp = etype_of_evar l hyps concl in
let evtyp, hyps, chop =
match chop_product fs evtyp with
| Some t -> t, trunc_named_context fs hyps, fs
@@ -257,14 +259,16 @@ let eterm_obligations env name evm fs ?status t ty =
let tactics_module = ["Program";"Tactics"]
let safe_init_constant md name () =
Coqlib.check_required_library ("Coq"::md);
- Universes.constr_of_global (Coqlib.coq_reference "Obligations" md name)
+ UnivGen.constr_of_global (Coqlib.coq_reference "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 =
- EConstr.Unsafe.to_constr (Reductionops.clos_norm_flags CClosure.betaiota (Global.env ()) Evd.empty (EConstr.of_constr c))
+ let env = Global.env () in
+ let sigma = Evd.from_env env in
+ EConstr.Unsafe.to_constr (Reductionops.clos_norm_flags CClosure.betaiota env sigma (EConstr.of_constr c))
exception NoObligations of Id.t option
@@ -294,17 +298,17 @@ type obligation =
type obligations = (obligation array * int)
type fixpoint_kind =
- | IsFixpoint of (Misctypes.lident option * Constrexpr.recursion_order_expr) list
+ | IsFixpoint of (lident option * Constrexpr.recursion_order_expr) list
| IsCoFixpoint
-type notations = (Misctypes.lstring * Constrexpr.constr_expr * Notation_term.scope_name option) list
+type notations = (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: UState.t;
- prg_univdecl: Univdecls.universe_decl;
+ prg_univdecl: UState.universe_decl;
prg_obligations: obligations;
prg_deps : Id.t list;
prg_fixkind : fixpoint_kind option ;
@@ -356,7 +360,7 @@ let _ =
optread = get_shrink_obligations;
optwrite = set_shrink_obligations; }
-let evar_of_obligation o = make_evar (Global.named_context_val ()) o.obl_type
+let evar_of_obligation o = make_evar (Global.named_context_val ()) (EConstr.of_constr o.obl_type)
let get_obligation_body expand obl =
match obl.obl_body with
@@ -470,7 +474,7 @@ let subst_body expand 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)
+ let nf = UnivSubst.nf_evars_and_universes_opt_subst (fun x -> None)
(UState.subst prg.prg_ctx) in
let opaque = prg.prg_opaque in
let fix_exn = Hook.get get_fix_exn () in
@@ -519,8 +523,10 @@ let declare_mutual_definition l =
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 (EConstr.of_constr subs)) in
- let typ = snd (Reductionops.splay_prod_n (Global.env ()) Evd.empty len (EConstr.of_constr typ)) in
+ let env = Global.env () in
+ let sigma = Evd.from_env env in
+ let term = snd (Reductionops.splay_lam_n env sigma len (EConstr.of_constr subs)) in
+ let typ = snd (Reductionops.splay_prod_n env sigma len (EConstr.of_constr typ)) in
let term = EConstr.Unsafe.to_constr term in
let typ = EConstr.Unsafe.to_constr typ in
x.prg_reduce term, x.prg_reduce typ, x.prg_implicits) l)
@@ -553,15 +559,14 @@ let declare_mutual_definition l =
(* Declare the recursive definitions *)
let univs = UState.const_univ_entry ~poly first.prg_ctx in
let fix_exn = Hook.get get_fix_exn () in
- let kns = List.map4 (DeclareDef.declare_fix ~opaque (local, poly, kind) Universes.empty_binders univs)
+ let kns = List.map4 (DeclareDef.declare_fix ~opaque (local, poly, kind) UnivNames.empty_binders univs)
fixnames fixdecls fixtypes fiximps in
(* Declare notations *)
List.iter (Metasyntax.add_notation_interpretation (Global.env())) 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
+ List.iter progmap_remove l; gr
let decompose_lam_prod c ty =
let open Context.Rel.Declaration in
@@ -610,7 +615,7 @@ let shrink_body c ty =
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)
+ Hints.add_hints ~local [Id.to_string prg.prg_name] (unfold_entry cst)
let it_mkLambda_or_LetIn_or_clean t ctx =
let open Context.Rel.Declaration in
@@ -742,7 +747,7 @@ let all_programs () =
type progress =
| Remain of int
| Dependent
- | Defined of global_reference
+ | Defined of GlobRef.t
let obligations_message rem =
if rem > 0 then
@@ -768,8 +773,8 @@ let update_obls prg obls rem =
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)
+ Defined kn
+ else Dependent)
let is_defined obls x = not (Option.is_empty obls.(x).obl_body)
@@ -813,10 +818,9 @@ let rec string_of_list sep f = function
let solve_by_tac name evi t poly ctx =
let id = name in
- let concl = EConstr.of_constr 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
+ id ~goal_kind:(goal_kind poly) ctx evi.evar_hyps evi.evar_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, () = Future.force entry.const_entry_body in
@@ -848,12 +852,12 @@ let obligation_terminator name num guard hook auto pf =
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.Expand), Opaque -> err_not_transp ()
+ | (true, _), Opaque -> err_not_transp ()
+ | (false, _), Opaque -> Evar_kinds.Define true
+ | (_, Evar_kinds.Define true), Transparent ->
Evar_kinds.Define false
- | (_, status), Vernacexpr.Transparent -> status
+ | (_, status), Transparent -> status
in
let obl = { obl with obl_status = false, status } in
let ctx =
@@ -890,7 +894,7 @@ let obligation_terminator name num guard hook auto pf =
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 cst = match gr with GlobRef.ConstRef cst -> cst | _ -> assert false in
let transparent = evaluable_constant cst (Global.env ()) in
let () = match obl.obl_status with
(true, Evar_kinds.Expand)
@@ -957,7 +961,7 @@ 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
+ if num >= 0 && num < Array.length obls then
let obl = obls.(num) in
match obl.obl_body with
None -> solve_obligation prg num tac
@@ -1068,9 +1072,11 @@ let show_obligations_of_prg ?(msg=true) prg =
if !showed > 0 then (
decr showed;
let x = subst_deps_obl obls x in
+ let env = Global.env () in
+ let sigma = Evd.from_env env 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 ++
+ hov 1 (Printer.pr_constr_env env sigma x.obl_type ++
str "." ++ fnl ())))
| Some _ -> ())
obls
@@ -1086,11 +1092,13 @@ let show_obligations ?(msg=true) n =
let show_term n =
let prg = get_prog_err n in
let n = prg.prg_name in
+ let env = Global.env () in
+ let sigma = Evd.from_env env 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)
+ Printer.pr_constr_env env sigma prg.prg_type ++ spc () ++ str ":=" ++ fnl ()
+ ++ Printer.pr_constr_env env sigma prg.prg_body)
-let add_definition n ?term t ctx ?(univdecl=Univdecls.default_univ_decl)
+let add_definition n ?term t ctx ?(univdecl=UState.default_univ_decl)
?(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
@@ -1110,7 +1118,7 @@ let add_definition n ?term t ctx ?(univdecl=Univdecls.default_univ_decl)
| Remain rem -> Flags.if_verbose (fun () -> show_obligations ~msg:false (Some n)) (); res
| _ -> res)
-let add_mutual_definitions l ctx ?(univdecl=Univdecls.default_univ_decl) ?tactic
+let add_mutual_definitions l ctx ?(univdecl=UState.default_univ_decl) ?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
diff --git a/vernac/obligations.mli b/vernac/obligations.mli
index cc2cacd86..a37c30aaf 100644
--- a/vernac/obligations.mli
+++ b/vernac/obligations.mli
@@ -12,7 +12,6 @@ open Environ
open Constr
open Evd
open Names
-open Globnames
(* This is a hack to make it possible for Obligations to craft a Qed
* behind the scenes. The fix_exn the Stm attaches to the Future proof
@@ -49,13 +48,13 @@ type obligation_info =
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 *)
+ | Defined of GlobRef.t (* Defined as id *)
val default_tactic : unit Proofview.tactic ref
val add_definition : Names.Id.t -> ?term:constr -> types ->
UState.t ->
- ?univdecl:Univdecls.universe_decl -> (* Universe binders and constraints *)
+ ?univdecl:UState.universe_decl -> (* Universe binders and constraints *)
?implicits:(Constrexpr.explicitation * (bool * bool * bool)) list ->
?kind:Decl_kinds.definition_kind ->
?tactic:unit Proofview.tactic ->
@@ -63,17 +62,17 @@ val add_definition : Names.Id.t -> ?term:constr -> types ->
?hook:(UState.t -> unit) Lemmas.declaration_hook -> ?opaque:bool -> obligation_info -> progress
type notations =
- (Misctypes.lstring * Constrexpr.constr_expr * Notation_term.scope_name option) list
+ (lstring * Constrexpr.constr_expr * Notation_term.scope_name option) list
type fixpoint_kind =
- | IsFixpoint of (Misctypes.lident option * Constrexpr.recursion_order_expr) list
+ | IsFixpoint of (lident option * Constrexpr.recursion_order_expr) list
| IsCoFixpoint
val add_mutual_definitions :
(Names.Id.t * constr * types *
(Constrexpr.explicitation * (bool * bool * bool)) list * obligation_info) list ->
UState.t ->
- ?univdecl:Univdecls.universe_decl -> (* Universe binders and constraints *)
+ ?univdecl:UState.universe_decl -> (* Universe binders and constraints *)
?tactic:unit Proofview.tactic ->
?kind:Decl_kinds.definition_kind ->
?reduce:(constr -> constr) ->
diff --git a/vernac/ppvernac.ml b/vernac/ppvernac.ml
new file mode 100644
index 000000000..d0c423650
--- /dev/null
+++ b/vernac/ppvernac.ml
@@ -0,0 +1,1222 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+open Pp
+open Names
+
+open CErrors
+open Util
+open CAst
+
+open Extend
+open Libnames
+open Decl_kinds
+open Constrexpr
+open Constrexpr_ops
+open Vernacexpr
+open Declaremods
+open Pputils
+
+ open Ppconstr
+
+ let do_not_tag _ x = x
+ let tag_keyword = do_not_tag ()
+ let tag_vernac = do_not_tag
+
+ let keyword s = tag_keyword (str s)
+
+ let pr_constr = pr_constr_expr
+ let pr_lconstr = pr_lconstr_expr
+ let pr_spc_lconstr = pr_sep_com spc pr_lconstr_expr
+
+ let pr_uconstraint (l, d, r) =
+ pr_glob_level l ++ spc () ++ Univ.pr_constraint_type d ++ spc () ++
+ pr_glob_level r
+
+ let pr_univ_name_list = function
+ | None -> mt ()
+ | Some l ->
+ str "@{" ++ prlist_with_sep spc pr_lname l ++ str"}"
+
+ let pr_univdecl_instance l extensible =
+ prlist_with_sep spc pr_lident l ++
+ (if extensible then str"+" else mt ())
+
+ let pr_univdecl_constraints l extensible =
+ if List.is_empty l && extensible then mt ()
+ else str"|" ++ spc () ++ prlist_with_sep (fun () -> str",") pr_uconstraint l ++
+ (if extensible then str"+" else mt())
+
+ let pr_universe_decl l =
+ let open UState in
+ match l with
+ | None -> mt ()
+ | Some l ->
+ str"@{" ++ pr_univdecl_instance l.univdecl_instance l.univdecl_extensible_instance ++
+ pr_univdecl_constraints l.univdecl_constraints l.univdecl_extensible_constraints ++ str "}"
+
+ let pr_ident_decl (lid, l) =
+ pr_lident lid ++ pr_universe_decl l
+
+ let string_of_fqid fqid =
+ String.concat "." (List.map Id.to_string fqid)
+
+ let pr_fqid fqid = str (string_of_fqid fqid)
+
+ let pr_lfqid {CAst.loc;v=fqid} =
+ match loc with
+ | None -> pr_fqid fqid
+ | Some loc -> let (b,_) = Loc.unloc loc in
+ pr_located pr_fqid @@ Loc.tag ~loc:(Loc.make_loc (b,b + String.length (string_of_fqid fqid))) fqid
+
+ let pr_lname_decl (n, u) =
+ pr_lname n ++ pr_universe_decl u
+
+ let pr_smart_global = Pputils.pr_or_by_notation pr_reference
+
+ let pr_ltac_ref = Libnames.pr_reference
+
+ let pr_module = Libnames.pr_reference
+
+ let pr_import_module = Libnames.pr_reference
+
+ let sep_end = function
+ | VernacBullet _
+ | VernacSubproof _
+ | VernacEndSubproof -> str""
+ | _ -> str"."
+
+ let pr_gen t = Pputils.pr_raw_generic (Global.env ()) t
+
+ let sep = fun _ -> spc()
+ let sep_v2 = fun _ -> str"," ++ spc()
+
+ let pr_at_level = function
+ | NumLevel n -> keyword "at" ++ spc () ++ keyword "level" ++ spc () ++ int n
+ | NextLevel -> keyword "at" ++ spc () ++ keyword "next" ++ spc () ++ keyword "level"
+
+ let pr_constr_as_binder_kind = let open Notation_term in function
+ | AsIdent -> keyword "as ident"
+ | AsIdentOrPattern -> keyword "as pattern"
+ | AsStrictPattern -> keyword "as strict pattern"
+
+ let pr_strict b = if b then str "strict " else mt ()
+
+ let pr_set_entry_type pr = function
+ | ETName -> str"ident"
+ | ETReference -> str"global"
+ | ETPattern (b,None) -> pr_strict b ++ str"pattern"
+ | ETPattern (b,Some n) -> pr_strict b ++ str"pattern" ++ spc () ++ pr_at_level (NumLevel n)
+ | ETConstr lev -> str"constr" ++ pr lev
+ | ETOther (_,e) -> str e
+ | ETConstrAsBinder (bk,lev) -> pr lev ++ spc () ++ pr_constr_as_binder_kind bk
+ | ETBigint -> str "bigint"
+ | ETBinder true -> str "binder"
+ | ETBinder false -> str "closed binder"
+
+ let pr_at_level_opt = function
+ | None -> mt ()
+ | Some n -> spc () ++ pr_at_level n
+
+ let pr_set_simple_entry_type =
+ pr_set_entry_type pr_at_level_opt
+
+ let pr_comment pr_c = function
+ | CommentConstr c -> pr_c c
+ | CommentString s -> qs s
+ | CommentInt n -> int n
+
+ let pr_in_out_modules = function
+ | SearchInside l -> spc() ++ keyword "inside" ++ spc() ++ prlist_with_sep sep pr_module l
+ | SearchOutside [] -> mt()
+ | SearchOutside l -> spc() ++ keyword "outside" ++ spc() ++ prlist_with_sep sep pr_module l
+
+ let pr_search_about (b,c) =
+ (if b then str "-" else mt()) ++
+ match c with
+ | SearchSubPattern p -> pr_constr_pattern_expr p
+ | SearchString (s,sc) -> qs s ++ pr_opt (fun sc -> str "%" ++ str sc) sc
+
+ let pr_search a gopt b pr_p =
+ pr_opt (fun g -> Goal_select.pr_goal_selector g ++ str ":"++ spc()) gopt
+ ++
+ match a with
+ | SearchHead c -> keyword "SearchHead" ++ spc() ++ pr_p c ++ pr_in_out_modules b
+ | SearchPattern c -> keyword "SearchPattern" ++ spc() ++ pr_p c ++ pr_in_out_modules b
+ | SearchRewrite c -> keyword "SearchRewrite" ++ spc() ++ pr_p c ++ pr_in_out_modules b
+ | SearchAbout sl ->
+ keyword "Search" ++ spc() ++ prlist_with_sep spc pr_search_about sl ++ pr_in_out_modules b
+
+ let pr_locality local = if local then keyword "Local" else keyword "Global"
+
+ let pr_option_ref_value = function
+ | QualidRefValue id -> pr_reference id
+ | StringRefValue s -> qs s
+
+ let pr_printoption table b =
+ prlist_with_sep spc str table ++
+ pr_opt (prlist_with_sep sep pr_option_ref_value) b
+
+ let pr_set_option a b =
+ let pr_opt_value = function
+ | IntValue None -> assert false
+ (* This should not happen because of the grammar *)
+ | IntValue (Some n) -> spc() ++ int n
+ | StringValue s -> spc() ++ str s
+ | StringOptValue None -> mt()
+ | StringOptValue (Some s) -> spc() ++ str s
+ | BoolValue b -> mt()
+ in pr_printoption a None ++ pr_opt_value b
+
+ let pr_opt_hintbases l = match l with
+ | [] -> mt()
+ | _ as z -> str":" ++ spc() ++ prlist_with_sep sep str z
+
+ let pr_reference_or_constr pr_c = function
+ | HintsReference r -> pr_reference r
+ | HintsConstr c -> pr_c c
+
+ let pr_hint_mode = function
+ | ModeInput -> str"+"
+ | ModeNoHeadEvar -> str"!"
+ | ModeOutput -> str"-"
+
+ let pr_hint_info pr_pat { Typeclasses.hint_priority = pri; hint_pattern = pat } =
+ pr_opt (fun x -> str"|" ++ int x) pri ++
+ pr_opt (fun y -> (if Option.is_empty pri then str"| " else mt()) ++ pr_pat y) pat
+
+ let pr_hints db h pr_c pr_pat =
+ let opth = pr_opt_hintbases db in
+ let pph =
+ match h with
+ | HintsResolve l ->
+ keyword "Resolve " ++ prlist_with_sep sep
+ (fun (info, _, c) -> pr_reference_or_constr pr_c c ++ pr_hint_info pr_pat info)
+ l
+ | HintsResolveIFF (l2r, l, n) ->
+ keyword "Resolve " ++ str (if l2r then "->" else "<-")
+ ++ prlist_with_sep sep pr_reference l
+ | HintsImmediate l ->
+ keyword "Immediate" ++ spc() ++
+ prlist_with_sep sep (fun c -> pr_reference_or_constr pr_c c) l
+ | HintsUnfold l ->
+ keyword "Unfold" ++ spc () ++ prlist_with_sep sep pr_reference l
+ | HintsTransparency (l, b) ->
+ keyword (if b then "Transparent" else "Opaque")
+ ++ spc ()
+ ++ prlist_with_sep sep pr_reference l
+ | HintsMode (m, l) ->
+ keyword "Mode"
+ ++ spc ()
+ ++ pr_reference m ++ spc() ++
+ prlist_with_sep spc pr_hint_mode l
+ | HintsConstructors c ->
+ keyword "Constructors"
+ ++ spc() ++ prlist_with_sep spc pr_reference c
+ | HintsExtern (n,c,tac) ->
+ let pat = match c with None -> mt () | Some pat -> pr_pat pat in
+ keyword "Extern" ++ spc() ++ int n ++ spc() ++ pat ++ str" =>" ++
+ spc() ++ Pputils.pr_raw_generic (Global.env ()) tac
+ in
+ hov 2 (keyword "Hint "++ pph ++ opth)
+
+ let pr_with_declaration pr_c = function
+ | CWith_Definition (id,udecl,c) ->
+ let p = pr_c c in
+ keyword "Definition" ++ spc() ++ pr_lfqid id ++ pr_universe_decl udecl ++ str" := " ++ p
+ | CWith_Module (id,qid) ->
+ keyword "Module" ++ spc() ++ pr_lfqid id ++ str" := " ++
+ pr_ast pr_qualid qid
+
+ let rec pr_module_ast leading_space pr_c = function
+ | { loc ; v = CMident qid } ->
+ if leading_space then
+ spc () ++ pr_located pr_qualid (loc, qid)
+ else
+ pr_located pr_qualid (loc,qid)
+ | { v = CMwith (mty,decl) } ->
+ let m = pr_module_ast leading_space pr_c mty in
+ let p = pr_with_declaration pr_c decl in
+ m ++ spc() ++ keyword "with" ++ spc() ++ p
+ | { v = CMapply (me1, ( { v = CMident _ } as me2 ) ) } ->
+ pr_module_ast leading_space pr_c me1 ++ spc() ++ pr_module_ast false pr_c me2
+ | { v = CMapply (me1,me2) } ->
+ pr_module_ast leading_space pr_c me1 ++ spc() ++
+ hov 1 (str"(" ++ pr_module_ast false pr_c me2 ++ str")")
+
+ let pr_inline = function
+ | DefaultInline -> mt ()
+ | NoInline -> str "[no inline]"
+ | InlineAt i -> str "[inline at level " ++ int i ++ str "]"
+
+ let pr_assumption_inline = function
+ | DefaultInline -> str "Inline"
+ | NoInline -> mt ()
+ | InlineAt i -> str "Inline(" ++ int i ++ str ")"
+
+ let pr_module_ast_inl leading_space pr_c (mast,inl) =
+ pr_module_ast leading_space pr_c mast ++ pr_inline inl
+
+ let pr_of_module_type prc = function
+ | Enforce mty -> str ":" ++ pr_module_ast_inl true prc mty
+ | Check mtys ->
+ prlist_strict (fun m -> str "<:" ++ pr_module_ast_inl true prc m) mtys
+
+ let pr_require_token = function
+ | Some true ->
+ keyword "Export" ++ spc ()
+ | Some false ->
+ keyword "Import" ++ spc ()
+ | None -> mt()
+
+ let pr_module_vardecls pr_c (export,idl,(mty,inl)) =
+ let m = pr_module_ast true pr_c mty in
+ spc() ++
+ hov 1 (str"(" ++ pr_require_token export ++
+ prlist_with_sep spc pr_lident idl ++ str":" ++ m ++ str")")
+
+ let pr_module_binders l pr_c =
+ prlist_strict (pr_module_vardecls pr_c) l
+
+ let pr_type_option pr_c = function
+ | { v = CHole (k, Namegen.IntroAnonymous, _) } -> mt()
+ | _ as c -> brk(0,2) ++ str" :" ++ pr_c c
+
+ let pr_decl_notation prc ({loc; v=ntn},c,scopt) =
+ fnl () ++ keyword "where " ++ qs ntn ++ str " := "
+ ++ Flags.without_option Flags.beautify prc c ++
+ pr_opt (fun sc -> str ": " ++ str sc) scopt
+
+ let pr_binders_arg =
+ pr_non_empty_arg pr_binders
+
+ let pr_and_type_binders_arg bl =
+ pr_binders_arg bl
+
+ let pr_onescheme (idop,schem) =
+ match schem with
+ | InductionScheme (dep,ind,s) ->
+ (match idop with
+ | Some id -> hov 0 (pr_lident id ++ str" :=") ++ spc()
+ | None -> spc ()
+ ) ++
+ hov 0 ((if dep then keyword "Induction for" else keyword "Minimality for")
+ ++ spc() ++ pr_smart_global ind) ++ spc() ++
+ hov 0 (keyword "Sort" ++ spc() ++ Termops.pr_sort_family s)
+ | CaseScheme (dep,ind,s) ->
+ (match idop with
+ | Some id -> hov 0 (pr_lident id ++ str" :=") ++ spc()
+ | None -> spc ()
+ ) ++
+ hov 0 ((if dep then keyword "Elimination for" else keyword "Case for")
+ ++ spc() ++ pr_smart_global ind) ++ spc() ++
+ hov 0 (keyword "Sort" ++ spc() ++ Termops.pr_sort_family s)
+ | EqualityScheme ind ->
+ (match idop with
+ | Some id -> hov 0 (pr_lident id ++ str" :=") ++ spc()
+ | None -> spc()
+ ) ++
+ hov 0 (keyword "Equality for")
+ ++ spc() ++ pr_smart_global ind
+
+ let begin_of_inductive = function
+ | [] -> 0
+ | (_,({loc},_))::_ -> Option.cata (fun loc -> fst (Loc.unloc loc)) 0 loc
+
+ let pr_class_rawexpr = function
+ | FunClass -> keyword "Funclass"
+ | SortClass -> keyword "Sortclass"
+ | RefClass qid -> pr_smart_global qid
+
+ let pr_assumption_token many discharge kind =
+ match discharge, kind with
+ | (NoDischarge,Logical) ->
+ keyword (if many then "Axioms" else "Axiom")
+ | (NoDischarge,Definitional) ->
+ keyword (if many then "Parameters" else "Parameter")
+ | (NoDischarge,Conjectural) -> str"Conjecture"
+ | (DoDischarge,Logical) ->
+ keyword (if many then "Hypotheses" else "Hypothesis")
+ | (DoDischarge,Definitional) ->
+ keyword (if many then "Variables" else "Variable")
+ | (DoDischarge,Conjectural) ->
+ anomaly (Pp.str "Don't know how to beautify a local conjecture.")
+
+ let pr_params pr_c (xl,(c,t)) =
+ hov 2 (prlist_with_sep sep pr_lident xl ++ spc() ++
+ (if c then str":>" else str":" ++
+ spc() ++ pr_c t))
+
+ let rec factorize = function
+ | [] -> []
+ | (c,(idl,t))::l ->
+ match factorize l with
+ | (xl,((c', t') as r))::l'
+ when (c : bool) == c' && Pervasives.(=) t t' ->
+ (** FIXME: we need equality on constr_expr *)
+ (idl@xl,r)::l'
+ | l' -> (idl,(c,t))::l'
+
+ let pr_ne_params_list pr_c l =
+ match factorize l with
+ | [p] -> pr_params pr_c p
+ | l ->
+ prlist_with_sep spc
+ (fun p -> hov 1 (str "(" ++ pr_params pr_c p ++ str ")")) l
+
+(*
+ prlist_with_sep pr_semicolon (pr_params pr_c)
+*)
+
+ let pr_thm_token k = keyword (Kindops.string_of_theorem_kind k)
+
+ let pr_syntax_modifier = function
+ | SetItemLevel (l,n) ->
+ prlist_with_sep sep_v2 str l ++ spc () ++ pr_at_level n
+ | SetItemLevelAsBinder (l,bk,n) ->
+ prlist_with_sep sep_v2 str l ++
+ spc() ++ pr_at_level_opt n ++ spc() ++ pr_constr_as_binder_kind bk
+ | SetLevel n -> pr_at_level (NumLevel n)
+ | SetAssoc LeftA -> keyword "left associativity"
+ | SetAssoc RightA -> keyword "right associativity"
+ | SetAssoc NonA -> keyword "no associativity"
+ | SetEntryType (x,typ) -> str x ++ spc() ++ pr_set_simple_entry_type typ
+ | SetOnlyPrinting -> keyword "only printing"
+ | SetOnlyParsing -> keyword "only parsing"
+ | SetCompatVersion v -> keyword("compat \"" ^ Flags.pr_version v ^ "\"")
+ | SetFormat("text",s) -> keyword "format " ++ pr_ast qs s
+ | SetFormat(k,s) -> keyword "format " ++ qs k ++ spc() ++ pr_ast qs s
+
+ let pr_syntax_modifiers = function
+ | [] -> mt()
+ | l -> spc() ++
+ hov 1 (str"(" ++ prlist_with_sep sep_v2 pr_syntax_modifier l ++ str")")
+
+ let pr_rec_definition ((iddecl,ro,bl,type_,def),ntn) =
+ let pr_pure_lconstr c = Flags.without_option Flags.beautify pr_lconstr c in
+ let annot = pr_guard_annot pr_lconstr_expr bl ro in
+ pr_ident_decl iddecl ++ pr_binders_arg bl ++ annot
+ ++ pr_type_option (fun c -> spc() ++ pr_lconstr_expr c) type_
+ ++ pr_opt (fun def -> str":=" ++ brk(1,2) ++ pr_pure_lconstr def) def
+ ++ prlist (pr_decl_notation pr_constr) ntn
+
+ let pr_statement head (idpl,(bl,c)) =
+ hov 2
+ (head ++ spc() ++ pr_ident_decl idpl ++ spc() ++
+ (match bl with [] -> mt() | _ -> pr_binders bl ++ spc()) ++
+ str":" ++ pr_spc_lconstr c)
+
+(**************************************)
+(* Pretty printer for vernac commands *)
+(**************************************)
+
+ let pr_constrarg c = spc () ++ pr_constr c
+ let pr_lconstrarg c = spc () ++ pr_lconstr c
+ let pr_intarg n = spc () ++ int n
+
+ let pr_oc = function
+ | None -> str" :"
+ | Some true -> str" :>"
+ | Some false -> str" :>>"
+
+ let pr_record_field ((x, pri), ntn) =
+ let prx = match x with
+ | (oc,AssumExpr (id,t)) ->
+ hov 1 (pr_lname id ++
+ pr_oc oc ++ spc() ++
+ pr_lconstr_expr t)
+ | (oc,DefExpr(id,b,opt)) -> (match opt with
+ | Some t ->
+ hov 1 (pr_lname id ++
+ pr_oc oc ++ spc() ++
+ pr_lconstr_expr t ++ str" :=" ++ pr_lconstr b)
+ | None ->
+ hov 1 (pr_lname id ++ str" :=" ++ spc() ++
+ pr_lconstr b)) in
+ let prpri = match pri with None -> mt() | Some i -> str "| " ++ int i in
+ prx ++ prpri ++ prlist (pr_decl_notation pr_constr) ntn
+
+ let pr_record_decl b c fs =
+ pr_opt pr_lident c ++ (if c = None then str"{" else str" {") ++
+ hv 0 (prlist_with_sep pr_semicolon pr_record_field fs ++ str"}")
+
+ let pr_printable = function
+ | PrintFullContext ->
+ keyword "Print All"
+ | PrintSectionContext s ->
+ keyword "Print Section" ++ spc() ++ Libnames.pr_reference s
+ | PrintGrammar ent ->
+ keyword "Print Grammar" ++ spc() ++ str ent
+ | PrintLoadPath dir ->
+ keyword "Print LoadPath" ++ pr_opt DirPath.print dir
+ | PrintModules ->
+ keyword "Print Modules"
+ | PrintMLLoadPath ->
+ keyword "Print ML Path"
+ | PrintMLModules ->
+ keyword "Print ML Modules"
+ | PrintDebugGC ->
+ keyword "Print ML GC"
+ | PrintGraph ->
+ keyword "Print Graph"
+ | PrintClasses ->
+ keyword "Print Classes"
+ | PrintTypeClasses ->
+ keyword "Print TypeClasses"
+ | PrintInstances qid ->
+ keyword "Print Instances" ++ spc () ++ pr_smart_global qid
+ | PrintCoercions ->
+ keyword "Print Coercions"
+ | PrintCoercionPaths (s,t) ->
+ keyword "Print Coercion Paths" ++ spc()
+ ++ pr_class_rawexpr s ++ spc()
+ ++ pr_class_rawexpr t
+ | PrintCanonicalConversions ->
+ keyword "Print Canonical Structures"
+ | PrintTables ->
+ keyword "Print Tables"
+ | PrintHintGoal ->
+ keyword "Print Hint"
+ | PrintHint qid ->
+ keyword "Print Hint" ++ spc () ++ pr_smart_global qid
+ | PrintHintDb ->
+ keyword "Print Hint *"
+ | PrintHintDbName s ->
+ keyword "Print HintDb" ++ spc () ++ str s
+ | PrintUniverses (b, fopt) ->
+ let cmd =
+ if b then "Print Sorted Universes"
+ else "Print Universes"
+ in
+ keyword cmd ++ pr_opt str fopt
+ | PrintName (qid,udecl) ->
+ keyword "Print" ++ spc() ++ pr_smart_global qid ++ pr_univ_name_list udecl
+ | PrintModuleType qid ->
+ keyword "Print Module Type" ++ spc() ++ pr_reference qid
+ | PrintModule qid ->
+ keyword "Print Module" ++ spc() ++ pr_reference qid
+ | PrintInspect n ->
+ keyword "Inspect" ++ spc() ++ int n
+ | PrintScopes ->
+ keyword "Print Scopes"
+ | PrintScope s ->
+ keyword "Print Scope" ++ spc() ++ str s
+ | PrintVisibility s ->
+ keyword "Print Visibility" ++ pr_opt str s
+ | PrintAbout (qid,l,gopt) ->
+ pr_opt (fun g -> Goal_select.pr_goal_selector g ++ str ":"++ spc()) gopt
+ ++ keyword "About" ++ spc() ++ pr_smart_global qid ++ pr_univ_name_list l
+ | PrintImplicit qid ->
+ keyword "Print Implicit" ++ spc() ++ pr_smart_global qid
+ (* spiwack: command printing all the axioms and section variables used in a
+ term *)
+ | PrintAssumptions (b, t, qid) ->
+ let cmd = match b, t with
+ | true, true -> "Print All Dependencies"
+ | true, false -> "Print Opaque Dependencies"
+ | false, true -> "Print Transparent Dependencies"
+ | false, false -> "Print Assumptions"
+ in
+ keyword cmd ++ spc() ++ pr_smart_global qid
+ | PrintNamespace dp ->
+ keyword "Print Namespace" ++ DirPath.print dp
+ | PrintStrategy None ->
+ keyword "Print Strategies"
+ | PrintStrategy (Some qid) ->
+ keyword "Print Strategy" ++ pr_smart_global qid
+
+ let pr_using e =
+ let rec aux = function
+ | SsEmpty -> "()"
+ | SsType -> "(Type)"
+ | SsSingl { v=id } -> "("^Id.to_string id^")"
+ | SsCompl e -> "-" ^ aux e^""
+ | SsUnion(e1,e2) -> "("^aux e1 ^" + "^ aux e2^")"
+ | SsSubstr(e1,e2) -> "("^aux e1 ^" - "^ aux e2^")"
+ | SsFwdClose e -> "("^aux e^")*"
+ in Pp.str (aux e)
+
+ let pr_extend s cl =
+ let pr_arg a =
+ try pr_gen a
+ with Failure _ -> str "<error in " ++ str (fst s) ++ str ">" in
+ try
+ let rl = Egramml.get_extend_vernac_rule s in
+ let rec aux rl cl =
+ match rl, cl with
+ | Egramml.GramNonTerminal _ :: rl, arg :: cl -> pr_arg arg :: aux rl cl
+ | Egramml.GramTerminal s :: rl, cl -> str s :: aux rl cl
+ | [], [] -> []
+ | _ -> assert false in
+ hov 1 (pr_sequence identity (aux rl cl))
+ with Not_found ->
+ hov 1 (str "TODO(" ++ str (fst s) ++ spc () ++ prlist_with_sep sep pr_arg cl ++ str ")")
+
+ let pr_vernac_expr v =
+ let return = tag_vernac v in
+ match v with
+ | VernacLoad (f,s) ->
+ return (
+ keyword "Load"
+ ++ if f then
+ (spc() ++ keyword "Verbose" ++ spc())
+ else
+ spc() ++ qs s
+ )
+
+ (* Proof management *)
+ | VernacAbortAll ->
+ return (keyword "Abort All")
+ | VernacRestart ->
+ return (keyword "Restart")
+ | VernacUnfocus ->
+ return (keyword "Unfocus")
+ | VernacUnfocused ->
+ return (keyword "Unfocused")
+ | VernacAbort id ->
+ return (keyword "Abort" ++ pr_opt pr_lident id)
+ | VernacUndo i ->
+ return (
+ if Int.equal i 1 then keyword "Undo" else keyword "Undo" ++ pr_intarg i
+ )
+ | VernacUndoTo i ->
+ return (keyword "Undo" ++ spc() ++ keyword "To" ++ pr_intarg i)
+ | VernacFocus i ->
+ return (keyword "Focus" ++ pr_opt int i)
+ | VernacShow s ->
+ let pr_goal_reference = function
+ | OpenSubgoals -> mt ()
+ | NthGoal n -> spc () ++ int n
+ | GoalId id -> spc () ++ pr_id id
+ in
+ let pr_showable = function
+ | ShowGoal n -> keyword "Show" ++ pr_goal_reference n
+ | ShowProof -> keyword "Show Proof"
+ | ShowScript -> keyword "Show Script"
+ | ShowExistentials -> keyword "Show Existentials"
+ | ShowUniverses -> keyword "Show Universes"
+ | ShowProofNames -> keyword "Show Conjectures"
+ | ShowIntros b -> keyword "Show " ++ (if b then keyword "Intros" else keyword "Intro")
+ | ShowMatch id -> keyword "Show Match " ++ pr_reference id
+ in
+ return (pr_showable s)
+ | VernacCheckGuard ->
+ return (keyword "Guarded")
+
+ (* Resetting *)
+ | VernacResetName id ->
+ return (keyword "Reset" ++ spc() ++ pr_lident id)
+ | VernacResetInitial ->
+ return (keyword "Reset Initial")
+ | VernacBack i ->
+ return (
+ if Int.equal i 1 then keyword "Back" else keyword "Back" ++ pr_intarg i
+ )
+ | VernacBackTo i ->
+ return (keyword "BackTo" ++ pr_intarg i)
+
+ (* State management *)
+ | VernacWriteState s ->
+ return (keyword "Write State" ++ spc () ++ qs s)
+ | VernacRestoreState s ->
+ return (keyword "Restore State" ++ spc() ++ qs s)
+
+ (* Syntax *)
+ | VernacOpenCloseScope (opening,sc) ->
+ return (
+ keyword (if opening then "Open " else "Close ") ++
+ keyword "Scope" ++ spc() ++ str sc
+ )
+ | VernacDelimiters (sc,Some key) ->
+ return (
+ keyword "Delimit Scope" ++ spc () ++ str sc ++
+ spc() ++ keyword "with" ++ spc () ++ str key
+ )
+ | VernacDelimiters (sc, None) ->
+ return (
+ keyword "Undelimit Scope" ++ spc () ++ str sc
+ )
+ | VernacBindScope (sc,cll) ->
+ return (
+ keyword "Bind Scope" ++ spc () ++ str sc ++
+ spc() ++ keyword "with" ++ spc () ++ prlist_with_sep spc pr_class_rawexpr cll
+ )
+ | VernacInfix (({v=s},mv),q,sn) -> (* A Verifier *)
+ return (
+ hov 0 (hov 0 (keyword "Infix "
+ ++ qs s ++ str " :=" ++ pr_constrarg q) ++
+ pr_syntax_modifiers mv ++
+ (match sn with
+ | None -> mt()
+ | Some sc -> spc() ++ str":" ++ spc() ++ str sc))
+ )
+ | VernacNotation (c,({v=s},l),opt) ->
+ return (
+ hov 2 (keyword "Notation" ++ spc() ++ qs s ++
+ str " :=" ++ Flags.without_option Flags.beautify pr_constrarg c ++ pr_syntax_modifiers l ++
+ (match opt with
+ | None -> mt()
+ | Some sc -> str" :" ++ spc() ++ str sc))
+ )
+ | VernacSyntaxExtension (_, (s, l)) ->
+ return (
+ keyword "Reserved Notation" ++ spc() ++ pr_ast qs s ++
+ pr_syntax_modifiers l
+ )
+ | VernacNotationAddFormat(s,k,v) ->
+ return (
+ keyword "Format Notation " ++ qs s ++ spc () ++ qs k ++ spc() ++ qs v
+ )
+
+ (* Gallina *)
+ | VernacDefinition ((discharge,kind),id,b) -> (* A verifier... *)
+ let pr_def_token dk =
+ keyword (
+ if Name.is_anonymous (fst id).v
+ then "Goal"
+ else Kindops.string_of_definition_object_kind dk)
+ in
+ let pr_reduce = function
+ | None -> mt()
+ | Some r ->
+ keyword "Eval" ++ spc() ++
+ pr_red_expr (pr_constr, pr_lconstr, pr_smart_global, pr_constr) keyword r ++
+ keyword " in" ++ spc()
+ in
+ let pr_def_body = function
+ | DefineBody (bl,red,body,d) ->
+ let ty = match d with
+ | None -> mt()
+ | Some ty -> spc() ++ str":" ++ pr_spc_lconstr ty
+ in
+ (pr_binders_arg bl,ty,Some (pr_reduce red ++ pr_lconstr body))
+ | ProveBody (bl,t) ->
+ let typ u = if (fst id).v = Anonymous then (assert (bl = []); u) else (str" :" ++ u) in
+ (pr_binders_arg bl, typ (pr_spc_lconstr t), None) in
+ let (binds,typ,c) = pr_def_body b in
+ return (
+ hov 2 (
+ pr_def_token kind ++ spc()
+ ++ pr_lname_decl id ++ binds ++ typ
+ ++ (match c with
+ | None -> mt()
+ | Some cc -> str" :=" ++ spc() ++ cc))
+ )
+
+ | VernacStartTheoremProof (ki,l) ->
+ return (
+ hov 1 (pr_statement (pr_thm_token ki) (List.hd l) ++
+ prlist (pr_statement (spc () ++ keyword "with")) (List.tl l))
+ )
+
+ | VernacEndProof Admitted ->
+ return (keyword "Admitted")
+
+ | VernacEndProof (Proved (opac,o)) -> return (
+ let open Proof_global in
+ match o with
+ | None -> (match opac with
+ | Transparent -> keyword "Defined"
+ | Opaque -> keyword "Qed")
+ | Some id -> (if opac <> Transparent then keyword "Save" else keyword "Defined") ++ spc() ++ pr_lident id
+ )
+ | VernacExactProof c ->
+ return (hov 2 (keyword "Proof" ++ pr_lconstrarg c))
+ | VernacAssumption ((discharge,kind),t,l) ->
+ let n = List.length (List.flatten (List.map fst (List.map snd l))) in
+ let pr_params (c, (xl, t)) =
+ hov 2 (prlist_with_sep sep pr_ident_decl xl ++ spc() ++
+ (if c then str":>" else str":" ++ spc() ++ pr_lconstr_expr t)) in
+ let assumptions = prlist_with_sep spc (fun p -> hov 1 (str "(" ++ pr_params p ++ str ")")) l in
+ return (hov 2 (pr_assumption_token (n > 1) discharge kind ++
+ pr_non_empty_arg pr_assumption_inline t ++ spc() ++ assumptions))
+ | VernacInductive (cum, p,f,l) ->
+ let pr_constructor (coe,(id,c)) =
+ hov 2 (pr_lident id ++ str" " ++
+ (if coe then str":>" else str":") ++
+ Flags.without_option Flags.beautify pr_spc_lconstr c)
+ in
+ let pr_constructor_list b l = match l with
+ | Constructors [] -> mt()
+ | Constructors l ->
+ let fst_sep = match l with [_] -> " " | _ -> " | " in
+ pr_com_at (begin_of_inductive l) ++
+ fnl() ++ str fst_sep ++
+ prlist_with_sep (fun _ -> fnl() ++ str" | ") pr_constructor l
+ | RecordDecl (c,fs) ->
+ pr_record_decl b c fs
+ in
+ let pr_oneind key (((coe,iddecl),indpar,s,k,lc),ntn) =
+ hov 0 (
+ str key ++ spc() ++
+ (if coe then str"> " else str"") ++ pr_ident_decl iddecl ++
+ pr_and_type_binders_arg indpar ++
+ pr_opt (fun s -> str":" ++ spc() ++ pr_lconstr_expr s) s ++
+ str" :=") ++ pr_constructor_list k lc ++
+ prlist (pr_decl_notation pr_constr) ntn
+ in
+ let key =
+ let (_,_,_,k,_),_ = List.hd l in
+ let kind =
+ match k with Record -> "Record" | Structure -> "Structure"
+ | Inductive_kw -> "Inductive" | CoInductive -> "CoInductive"
+ | Class _ -> "Class" | Variant -> "Variant"
+ in
+ if p then
+ let cm =
+ match cum with
+ | Some VernacCumulative -> "Cumulative"
+ | Some VernacNonCumulative -> "NonCumulative"
+ | None -> ""
+ in
+ cm ^ " " ^ kind
+ else kind
+ in
+ return (
+ hov 1 (pr_oneind key (List.hd l)) ++
+ (prlist (fun ind -> fnl() ++ hov 1 (pr_oneind "with" ind)) (List.tl l))
+ )
+
+ | VernacFixpoint (local, recs) ->
+ let local = match local with
+ | DoDischarge -> "Let "
+ | NoDischarge -> ""
+ in
+ return (
+ hov 0 (str local ++ keyword "Fixpoint" ++ spc () ++
+ prlist_with_sep (fun _ -> fnl () ++ keyword "with"
+ ++ spc ()) pr_rec_definition recs)
+ )
+
+ | VernacCoFixpoint (local, corecs) ->
+ let local = match local with
+ | DoDischarge -> keyword "Let" ++ spc ()
+ | NoDischarge -> str ""
+ in
+ let pr_onecorec ((iddecl,bl,c,def),ntn) =
+ pr_ident_decl iddecl ++ spc() ++ pr_binders bl ++ spc() ++ str":" ++
+ spc() ++ pr_lconstr_expr c ++
+ pr_opt (fun def -> str":=" ++ brk(1,2) ++ pr_lconstr def) def ++
+ prlist (pr_decl_notation pr_constr) ntn
+ in
+ return (
+ hov 0 (local ++ keyword "CoFixpoint" ++ spc() ++
+ prlist_with_sep (fun _ -> fnl() ++ keyword "with" ++ spc ()) pr_onecorec corecs)
+ )
+ | VernacScheme l ->
+ return (
+ hov 2 (keyword "Scheme" ++ spc() ++
+ prlist_with_sep (fun _ -> fnl() ++ keyword "with" ++ spc ()) pr_onescheme l)
+ )
+ | VernacCombinedScheme (id, l) ->
+ return (
+ hov 2 (keyword "Combined Scheme" ++ spc() ++
+ pr_lident id ++ spc() ++ keyword "from" ++ spc() ++
+ prlist_with_sep (fun _ -> fnl() ++ str", ") pr_lident l)
+ )
+ | VernacUniverse v ->
+ return (
+ hov 2 (keyword "Universe" ++ spc () ++
+ prlist_with_sep (fun _ -> str",") pr_lident v)
+ )
+ | VernacConstraint v ->
+ return (
+ hov 2 (keyword "Constraint" ++ spc () ++
+ prlist_with_sep (fun _ -> str",") pr_uconstraint v)
+ )
+
+ (* Gallina extensions *)
+ | VernacBeginSection id ->
+ return (hov 2 (keyword "Section" ++ spc () ++ pr_lident id))
+ | VernacEndSegment id ->
+ return (hov 2 (keyword "End" ++ spc() ++ pr_lident id))
+ | VernacNameSectionHypSet (id,set) ->
+ return (hov 2 (keyword "Package" ++ spc() ++ pr_lident id ++ spc()++
+ str ":="++spc()++pr_using set))
+ | VernacRequire (from, exp, l) ->
+ let from = match from with
+ | None -> mt ()
+ | Some r -> keyword "From" ++ spc () ++ pr_module r ++ spc ()
+ in
+ return (
+ hov 2
+ (from ++ keyword "Require" ++ spc() ++ pr_require_token exp ++
+ prlist_with_sep sep pr_module l)
+ )
+ | VernacImport (f,l) ->
+ return (
+ (if f then keyword "Export" else keyword "Import") ++ spc() ++
+ prlist_with_sep sep pr_import_module l
+ )
+ | VernacCanonical q ->
+ return (
+ keyword "Canonical Structure" ++ spc() ++ pr_smart_global q
+ )
+ | VernacCoercion (id,c1,c2) ->
+ return (
+ hov 1 (
+ keyword "Coercion" ++ spc() ++
+ pr_smart_global id ++ spc() ++ str":" ++ spc() ++ pr_class_rawexpr c1 ++
+ spc() ++ str">->" ++ spc() ++ pr_class_rawexpr c2)
+ )
+ | VernacIdentityCoercion (id,c1,c2) ->
+ return (
+ hov 1 (
+ keyword "Identity Coercion" ++ spc() ++ pr_lident id ++
+ spc() ++ str":" ++ spc() ++ pr_class_rawexpr c1 ++ spc() ++ str">->" ++
+ spc() ++ pr_class_rawexpr c2)
+ )
+
+ | VernacInstance (abst, sup, (instid, bk, cl), props, info) ->
+ return (
+ hov 1 (
+ (if abst then keyword "Declare" ++ spc () else mt ()) ++
+ keyword "Instance" ++
+ (match instid with
+ | {loc; v = Name id}, l -> spc () ++ pr_ident_decl (CAst.(make ?loc id),l) ++ spc ()
+ | { v = Anonymous }, _ -> mt ()) ++
+ pr_and_type_binders_arg sup ++
+ str":" ++ spc () ++
+ (match bk with Implicit -> str "! " | Explicit -> mt ()) ++
+ pr_constr cl ++ pr_hint_info pr_constr_pattern_expr info ++
+ (match props with
+ | Some (true, { v = CRecord l}) -> spc () ++ str":=" ++ spc () ++ str"{" ++ pr_record_body l ++ str "}"
+ | Some (true,_) -> assert false
+ | Some (false,p) -> spc () ++ str":=" ++ spc () ++ pr_constr p
+ | None -> mt()))
+ )
+
+ | VernacContext l ->
+ return (
+ hov 1 (
+ keyword "Context" ++ pr_and_type_binders_arg l)
+ )
+
+ | VernacDeclareInstances insts ->
+ let pr_inst (id, info) =
+ pr_reference id ++ pr_hint_info pr_constr_pattern_expr info
+ in
+ return (
+ hov 1 (keyword "Existing" ++ spc () ++
+ keyword(String.plural (List.length insts) "Instance") ++
+ spc () ++ prlist_with_sep (fun () -> str", ") pr_inst insts)
+ )
+
+ | VernacDeclareClass id ->
+ return (
+ hov 1 (keyword "Existing" ++ spc () ++ keyword "Class" ++ spc () ++ pr_reference id)
+ )
+
+ (* Modules and Module Types *)
+ | VernacDefineModule (export,m,bl,tys,bd) ->
+ let b = pr_module_binders bl pr_lconstr in
+ return (
+ hov 2 (keyword "Module" ++ spc() ++ pr_require_token export ++
+ pr_lident m ++ b ++
+ pr_of_module_type pr_lconstr tys ++
+ (if List.is_empty bd then mt () else str ":= ") ++
+ prlist_with_sep (fun () -> str " <+")
+ (pr_module_ast_inl true pr_lconstr) bd)
+ )
+ | VernacDeclareModule (export,id,bl,m1) ->
+ let b = pr_module_binders bl pr_lconstr in
+ return (
+ hov 2 (keyword "Declare Module" ++ spc() ++ pr_require_token export ++
+ pr_lident id ++ b ++ str " :" ++
+ pr_module_ast_inl true pr_lconstr m1)
+ )
+ | VernacDeclareModuleType (id,bl,tyl,m) ->
+ let b = pr_module_binders bl pr_lconstr in
+ let pr_mt = pr_module_ast_inl true pr_lconstr in
+ return (
+ hov 2 (keyword "Module Type " ++ pr_lident id ++ b ++
+ prlist_strict (fun m -> str " <:" ++ pr_mt m) tyl ++
+ (if List.is_empty m then mt () else str ":= ") ++
+ prlist_with_sep (fun () -> str " <+ ") pr_mt m)
+ )
+ | VernacInclude (mexprs) ->
+ let pr_m = pr_module_ast_inl false pr_lconstr in
+ return (
+ hov 2 (keyword "Include" ++ spc() ++
+ prlist_with_sep (fun () -> str " <+ ") pr_m mexprs)
+ )
+ (* Solving *)
+ | VernacSolveExistential (i,c) ->
+ return (keyword "Existential" ++ spc () ++ int i ++ pr_lconstrarg c)
+
+ (* Auxiliary file and library management *)
+ | VernacAddLoadPath (fl,s,d) ->
+ return (
+ hov 2
+ (keyword "Add" ++
+ (if fl then spc () ++ keyword "Rec" ++ spc () else spc()) ++
+ keyword "LoadPath" ++ spc() ++ qs s ++
+ (match d with
+ | None -> mt()
+ | Some dir -> spc() ++ keyword "as" ++ spc() ++ DirPath.print dir))
+ )
+ | VernacRemoveLoadPath s ->
+ return (keyword "Remove LoadPath" ++ qs s)
+ | VernacAddMLPath (fl,s) ->
+ return (
+ keyword "Add"
+ ++ (if fl then spc () ++ keyword "Rec" ++ spc () else spc())
+ ++ keyword "ML Path"
+ ++ qs s
+ )
+ | VernacDeclareMLModule (l) ->
+ return (
+ hov 2 (keyword "Declare ML Module" ++ spc() ++ prlist_with_sep sep qs l)
+ )
+ | VernacChdir s ->
+ return (keyword "Cd" ++ pr_opt qs s)
+
+ (* Commands *)
+ | VernacCreateHintDb (dbname,b) ->
+ return (
+ hov 1 (keyword "Create HintDb" ++ spc () ++
+ str dbname ++ (if b then str" discriminated" else mt ()))
+ )
+ | VernacRemoveHints (dbnames, ids) ->
+ return (
+ hov 1 (keyword "Remove Hints" ++ spc () ++
+ prlist_with_sep spc (fun r -> pr_id (coerce_reference_to_id r)) ids ++
+ pr_opt_hintbases dbnames)
+ )
+ | VernacHints (dbnames,h) ->
+ return (pr_hints dbnames h pr_constr pr_constr_pattern_expr)
+ | VernacSyntacticDefinition (id,(ids,c),compat) ->
+ return (
+ hov 2
+ (keyword "Notation" ++ spc () ++ pr_lident id ++ spc () ++
+ prlist_with_sep spc pr_id ids ++ str":=" ++ pr_constrarg c ++
+ pr_syntax_modifiers
+ (match compat with
+ | None -> []
+ | Some Flags.Current -> [SetOnlyParsing]
+ | Some v -> [SetCompatVersion v]))
+ )
+ | VernacArguments (q, args, more_implicits, nargs, mods) ->
+ return (
+ hov 2 (
+ keyword "Arguments" ++ spc() ++
+ pr_smart_global q ++
+ let pr_s = function None -> str"" | Some {v=s} -> str "%" ++ str s in
+ let pr_if b x = if b then x else str "" in
+ let pr_br imp x = match imp with
+ | Vernacexpr.Implicit -> str "[" ++ x ++ str "]"
+ | Vernacexpr.MaximallyImplicit -> str "{" ++ x ++ str "}"
+ | Vernacexpr.NotImplicit -> x in
+ let rec print_arguments n l =
+ match n, l with
+ | Some 0, l -> spc () ++ str"/" ++ print_arguments None l
+ | _, [] -> mt()
+ | n, { name = id; recarg_like = k;
+ notation_scope = s;
+ implicit_status = imp } :: tl ->
+ spc() ++ pr_br imp (pr_if k (str"!") ++ Name.print id ++ pr_s s) ++
+ print_arguments (Option.map pred n) tl
+ in
+ let rec print_implicits = function
+ | [] -> mt ()
+ | (name, impl) :: rest ->
+ spc() ++ pr_br impl (Name.print name) ++ print_implicits rest
+ in
+ print_arguments nargs args ++
+ if not (List.is_empty more_implicits) then
+ prlist (fun l -> str"," ++ print_implicits l) more_implicits
+ else (mt ()) ++
+ (if not (List.is_empty mods) then str" : " else str"") ++
+ prlist_with_sep (fun () -> str", " ++ spc()) (function
+ | `ReductionDontExposeCase -> keyword "simpl nomatch"
+ | `ReductionNeverUnfold -> keyword "simpl never"
+ | `DefaultImplicits -> keyword "default implicits"
+ | `Rename -> keyword "rename"
+ | `Assert -> keyword "assert"
+ | `ExtraScopes -> keyword "extra scopes"
+ | `ClearImplicits -> keyword "clear implicits"
+ | `ClearScopes -> keyword "clear scopes")
+ mods)
+ )
+ | VernacReserve bl ->
+ let n = List.length (List.flatten (List.map fst bl)) in
+ return (
+ hov 2 (tag_keyword (str"Implicit Type" ++ str (if n > 1 then "s " else " "))
+ ++ pr_ne_params_list pr_lconstr_expr (List.map (fun sb -> false,sb) bl))
+ )
+ | VernacGeneralizable g ->
+ return (
+ hov 1 (tag_keyword (
+ str"Generalizable Variable" ++
+ match g with
+ | None -> str "s none"
+ | Some [] -> str "s all"
+ | Some idl ->
+ str (if List.length idl > 1 then "s " else " ") ++
+ prlist_with_sep spc pr_lident idl)
+ ))
+ | VernacSetOpacity(k,l) when Conv_oracle.is_transparent k ->
+ return (
+ hov 1 (keyword "Transparent" ++
+ spc() ++ prlist_with_sep sep pr_smart_global l)
+ )
+ | VernacSetOpacity(Conv_oracle.Opaque,l) ->
+ return (
+ hov 1 (keyword "Opaque" ++
+ spc() ++ prlist_with_sep sep pr_smart_global l)
+ )
+ | VernacSetOpacity _ ->
+ return (
+ CErrors.anomaly (keyword "VernacSetOpacity used to set something else.")
+ )
+ | VernacSetStrategy l ->
+ let pr_lev = function
+ | Conv_oracle.Opaque -> keyword "opaque"
+ | Conv_oracle.Expand -> keyword "expand"
+ | l when Conv_oracle.is_transparent l -> keyword "transparent"
+ | Conv_oracle.Level n -> int n
+ in
+ let pr_line (l,q) =
+ hov 2 (pr_lev l ++ spc() ++
+ str"[" ++ prlist_with_sep sep pr_smart_global q ++ str"]")
+ in
+ return (
+ hov 1 (keyword "Strategy" ++ spc() ++
+ hv 0 (prlist_with_sep sep pr_line l))
+ )
+ | VernacUnsetOption (export, na) ->
+ let export = if export then keyword "Export" ++ spc () else mt () in
+ return (
+ hov 1 (export ++ keyword "Unset" ++ spc() ++ pr_printoption na None)
+ )
+ | VernacSetOption (export, na,v) ->
+ let export = if export then keyword "Export" ++ spc () else mt () in
+ return (
+ hov 2 (export ++ keyword "Set" ++ spc() ++ pr_set_option na v)
+ )
+ | VernacAddOption (na,l) ->
+ return (
+ hov 2 (keyword "Add" ++ spc() ++ pr_printoption na (Some l))
+ )
+ | VernacRemoveOption (na,l) ->
+ return (
+ hov 2 (keyword "Remove" ++ spc() ++ pr_printoption na (Some l))
+ )
+ | VernacMemOption (na,l) ->
+ return (
+ hov 2 (keyword "Test" ++ spc() ++ pr_printoption na (Some l))
+ )
+ | VernacPrintOption na ->
+ return (
+ hov 2 (keyword "Test" ++ spc() ++ pr_printoption na None)
+ )
+ | VernacCheckMayEval (r,io,c) ->
+ let pr_mayeval r c = match r with
+ | Some r0 ->
+ hov 2 (keyword "Eval" ++ spc() ++
+ pr_red_expr (pr_constr,pr_lconstr,pr_smart_global, pr_constr) keyword r0 ++
+ spc() ++ keyword "in" ++ spc () ++ pr_lconstr c)
+ | None -> hov 2 (keyword "Check" ++ spc() ++ pr_lconstr c)
+ in
+ let pr_i = match io with None -> mt ()
+ | Some i -> Goal_select.pr_goal_selector i ++ str ": " in
+ return (pr_i ++ pr_mayeval r c)
+ | VernacGlobalCheck c ->
+ return (hov 2 (keyword "Type" ++ pr_constrarg c))
+ | VernacDeclareReduction (s,r) ->
+ return (
+ keyword "Declare Reduction" ++ spc () ++ str s ++ str " := " ++
+ pr_red_expr (pr_constr,pr_lconstr,pr_smart_global, pr_constr) keyword r
+ )
+ | VernacPrint p ->
+ return (pr_printable p)
+ | VernacSearch (sea,g,sea_r) ->
+ return (pr_search sea g sea_r pr_constr_pattern_expr)
+ | VernacLocate loc ->
+ let pr_locate =function
+ | LocateAny qid -> pr_smart_global qid
+ | LocateTerm qid -> keyword "Term" ++ spc() ++ pr_smart_global qid
+ | LocateFile f -> keyword "File" ++ spc() ++ qs f
+ | LocateLibrary qid -> keyword "Library" ++ spc () ++ pr_module qid
+ | LocateModule qid -> keyword "Module" ++ spc () ++ pr_module qid
+ | LocateOther (s, qid) -> keyword s ++ spc () ++ pr_ltac_ref qid
+ in
+ return (keyword "Locate" ++ spc() ++ pr_locate loc)
+ | VernacRegister (id, RegisterInline) ->
+ return (
+ hov 2
+ (keyword "Register Inline" ++ spc() ++ pr_lident id)
+ )
+ | VernacComments l ->
+ return (
+ hov 2
+ (keyword "Comments" ++ spc()
+ ++ prlist_with_sep sep (pr_comment pr_constr) l)
+ )
+
+ (* For extension *)
+ | VernacExtend (s,c) ->
+ return (pr_extend s c)
+ | VernacProof (None, None) ->
+ return (keyword "Proof")
+ | VernacProof (None, Some e) ->
+ return (keyword "Proof " ++ spc () ++
+ keyword "using" ++ spc() ++ pr_using e)
+ | VernacProof (Some te, None) ->
+ return (keyword "Proof with" ++ spc() ++ Pputils.pr_raw_generic (Global.env ()) te)
+ | VernacProof (Some te, Some e) ->
+ return (
+ keyword "Proof" ++ spc () ++
+ keyword "using" ++ spc() ++ pr_using e ++ spc() ++
+ keyword "with" ++ spc() ++ Pputils.pr_raw_generic (Global.env ()) te
+ )
+ | VernacProofMode s ->
+ return (keyword "Proof Mode" ++ str s)
+ | VernacBullet b ->
+ (* XXX: Redundant with Proof_bullet.print *)
+ return (let open Proof_bullet in begin match b with
+ | Dash n -> str (String.make n '-')
+ | Star n -> str (String.make n '*')
+ | Plus n -> str (String.make n '+')
+ end)
+ | VernacSubproof None ->
+ return (str "{")
+ | VernacSubproof (Some i) ->
+ return (Goal_select.pr_goal_selector i ++ str ":" ++ spc () ++ str "{")
+ | VernacEndSubproof ->
+ return (str "}")
+
+let pr_vernac_flag =
+ function
+ | VernacPolymorphic true -> keyword "Polymorphic"
+ | VernacPolymorphic false -> keyword "Monomorphic"
+ | VernacProgram -> keyword "Program"
+ | VernacLocal local -> pr_locality local
+
+ let rec pr_vernac_control v =
+ let return = tag_vernac v in
+ match v with
+ | VernacExpr (f, v') ->
+ List.fold_right
+ (fun f a -> pr_vernac_flag f ++ spc() ++ a)
+ f
+ (pr_vernac_expr v' ++ sep_end v')
+ | VernacTime (_,{v}) ->
+ return (keyword "Time" ++ spc() ++ pr_vernac_control v)
+ | VernacRedirect (s, {v}) ->
+ return (keyword "Redirect" ++ spc() ++ qs s ++ spc() ++ pr_vernac_control v)
+ | VernacTimeout(n,v) ->
+ return (keyword "Timeout " ++ int n ++ spc() ++ pr_vernac_control v)
+ | VernacFail v ->
+ return (keyword "Fail" ++ spc() ++ pr_vernac_control v)
+
+ let pr_vernac v =
+ try pr_vernac_control v
+ with e -> CErrors.print e
diff --git a/vernac/ppvernac.mli b/vernac/ppvernac.mli
new file mode 100644
index 000000000..4aa24bf5d
--- /dev/null
+++ b/vernac/ppvernac.mli
@@ -0,0 +1,26 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+(** This module implements pretty-printers for vernac_expr syntactic
+ objects and their subcomponents. *)
+
+val pr_set_entry_type : ('a -> Pp.t) -> 'a Extend.constr_entry_key_gen -> Pp.t
+
+(** Prints a fixpoint body *)
+val pr_rec_definition : (Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) -> Pp.t
+
+(** Prints a vernac expression without dot *)
+val pr_vernac_expr : Vernacexpr.vernac_expr -> Pp.t
+
+(** Prints a "proof using X" clause. *)
+val pr_using : Vernacexpr.section_subset_expr -> Pp.t
+
+(** Prints a vernac expression and closes it with a dot. *)
+val pr_vernac : Vernacexpr.vernac_control -> Pp.t
diff --git a/vernac/pvernac.ml b/vernac/pvernac.ml
new file mode 100644
index 000000000..bac882381
--- /dev/null
+++ b/vernac/pvernac.ml
@@ -0,0 +1,56 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+open Pcoq
+
+let uncurry f (x,y) = f x y
+
+let uvernac = create_universe "vernac"
+
+module Vernac_ =
+ struct
+ let gec_vernac s = Gram.entry_create ("vernac:" ^ s)
+
+ (* The different kinds of vernacular commands *)
+ let gallina = gec_vernac "gallina"
+ let gallina_ext = gec_vernac "gallina_ext"
+ let command = gec_vernac "command"
+ let syntax = gec_vernac "syntax_command"
+ let vernac_control = gec_vernac "Vernac.vernac_control"
+ let rec_definition = gec_vernac "Vernac.rec_definition"
+ let red_expr = new_entry utactic "red_expr"
+ let hint_info = gec_vernac "hint_info"
+ (* Main vernac entry *)
+ let main_entry = Gram.entry_create "vernac"
+ let noedit_mode = gec_vernac "noedit_command"
+
+ let () =
+ let act_vernac = Gram.action (fun v loc -> Some (to_coqloc loc, v)) in
+ let act_eoi = Gram.action (fun _ loc -> None) in
+ let rule = [
+ ([ Symbols.stoken Tok.EOI ], act_eoi);
+ ([ Symbols.snterm (Gram.Entry.obj vernac_control) ], act_vernac );
+ ] in
+ uncurry (Gram.extend main_entry) (None, [None, None, rule])
+
+ let command_entry_ref = ref noedit_mode
+ let command_entry =
+ Gram.Entry.of_parser "command_entry"
+ (fun strm -> Gram.parse_tokens_after_filter !command_entry_ref strm)
+
+ end
+
+let main_entry = Vernac_.main_entry
+
+let set_command_entry e = Vernac_.command_entry_ref := e
+let get_command_entry () = !Vernac_.command_entry_ref
+
+let () =
+ register_grammar Stdarg.wit_red_expr (Vernac_.red_expr);
diff --git a/vernac/pvernac.mli b/vernac/pvernac.mli
new file mode 100644
index 000000000..2993a1661
--- /dev/null
+++ b/vernac/pvernac.mli
@@ -0,0 +1,36 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+open Pcoq
+open Genredexpr
+open Vernacexpr
+
+val uvernac : gram_universe
+
+module Vernac_ :
+ sig
+ val gallina : vernac_expr Gram.entry
+ val gallina_ext : vernac_expr Gram.entry
+ val command : vernac_expr Gram.entry
+ val syntax : vernac_expr Gram.entry
+ val vernac_control : vernac_control Gram.entry
+ val rec_definition : (fixpoint_expr * decl_notation list) Gram.entry
+ val noedit_mode : vernac_expr Gram.entry
+ val command_entry : vernac_expr Gram.entry
+ val red_expr : raw_red_expr Gram.entry
+ val hint_info : Hints.hint_info_expr Gram.entry
+ end
+
+(** The main entry: reads an optional vernac command *)
+val main_entry : (Loc.t * vernac_control) option Gram.entry
+
+(** Handling of the proof mode entry *)
+val get_command_entry : unit -> vernac_expr Gram.entry
+val set_command_entry : vernac_expr Gram.entry -> unit
diff --git a/vernac/record.ml b/vernac/record.ml
index 6e745b2af..940859723 100644
--- a/vernac/record.ml
+++ b/vernac/record.ml
@@ -96,13 +96,13 @@ 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 -> CAst.make ?loc:n.CAst.loc @@ CHole (None, Misctypes.IntroAnonymous, None))
+ | None -> CAst.make ?loc:n.CAst.loc @@ CHole (None, Namegen.IntroAnonymous, None))
let binders_of_decls = List.map binder_of_decl
let typecheck_params_and_fields finite def id poly pl t ps nots fs =
let env0 = Global.env () in
- let sigma, decl = Univdecls.interp_univ_decl_opt env0 pl in
+ let sigma, decl = Constrexpr_ops.interp_univ_decl_opt env0 pl in
let _ =
let error bk {CAst.loc; v=name} =
match bk, name with
@@ -114,7 +114,7 @@ let typecheck_params_and_fields finite def id poly pl t ps nots fs =
(function CLocalDef (b, _, _) -> error default_binder_kind b
| CLocalAssum (ls, bk, ce) -> List.iter (error bk) ls
| CLocalPattern {CAst.loc} ->
- Loc.raise ?loc (Stream.Error "pattern with quote not allowed in record parameters.")) ps
+ Loc.raise ?loc (Stream.Error "pattern with quote not allowed in record parameters")) ps
in
let sigma, (impls_env, ((env1,newps), imps)) = interp_context_evars env0 sigma ps in
let sigma, typ, sort, template = match t with
@@ -122,7 +122,7 @@ let typecheck_params_and_fields finite def id poly pl t ps nots fs =
let env = EConstr.push_rel_context newps env0 in
let poly =
match t with
- | { CAst.v = CSort (Misctypes.GType []) } -> true | _ -> false in
+ | { CAst.v = CSort (Glob_term.GType []) } -> true | _ -> false in
let sigma, s = interp_type_evars env sigma ~impls:empty_internalization_env t in
let sred = Reductionops.whd_allnolet env sigma s in
(match EConstr.kind sigma sred with
@@ -152,7 +152,7 @@ let typecheck_params_and_fields finite def id poly pl t ps nots fs =
interp_fields_evars env_ar sigma impls_env nots (binders_of_decls fs)
in
let sigma =
- Pretyping.solve_remaining_evars Pretyping.all_and_fail_flags env_ar sigma Evd.empty in
+ Pretyping.solve_remaining_evars Pretyping.all_and_fail_flags env_ar sigma (Evd.from_env env_ar) in
let sigma, typ =
let _, univ = compute_constructor_level sigma env_ar newfs in
if not def && (Sorts.is_prop sort ||
@@ -168,11 +168,11 @@ let typecheck_params_and_fields finite def id poly pl t ps nots fs =
EConstr.mkSort (Sorts.sort_of_univ univ)
else sigma, typ
in
- let sigma, _ = Evarutil.nf_evars_and_universes sigma in
+ let sigma = Evd.minimize_universes sigma in
let newfs = List.map (EConstr.to_rel_decl sigma) newfs in
let newps = List.map (EConstr.to_rel_decl sigma) newps in
let typ = EConstr.to_constr sigma typ in
- let ce t = Pretyping.check_evars env0 Evd.empty sigma (EConstr.of_constr t) in
+ let ce t = Pretyping.check_evars env0 (Evd.from_env env0) sigma (EConstr.of_constr t) in
let univs = Evd.check_univ_decl ~poly sigma decl in
let ubinders = Evd.universe_binders sigma in
List.iter (iter_constr ce) (List.rev newps);
@@ -316,7 +316,7 @@ let declare_projections indsp ctx ?(kind=StructureComponent) binder_name coers u
let gr = Nametab.locate (Libnames.qualid_of_ident fid) in
let kn = destConstRef gr in
Declare.definition_message fid;
- Universes.register_universe_binders gr ubinders;
+ UnivNames.register_universe_binders gr ubinders;
kn, mkProj (Projection.make kn false,mkRel 1)
else
let ccl = subst_projection fid subst ti in
@@ -352,7 +352,7 @@ let declare_projections indsp ctx ?(kind=StructureComponent) binder_name coers u
applist (mkConstU (kn,u),proj_args)
in
Declare.definition_message fid;
- Universes.register_universe_binders (ConstRef kn) ubinders;
+ UnivNames.register_universe_binders (ConstRef kn) ubinders;
kn, constr_fip
with Type_errors.TypeError (ctx,te) ->
raise (NotDefinable (BadTypedProj (fid,ctx,te)))
@@ -465,9 +465,9 @@ let declare_class finite def cum ubinders univs id idbuild paramimpls params ari
in
let cref = ConstRef cst in
Impargs.declare_manual_implicits false cref [paramimpls];
- Universes.register_universe_binders cref ubinders;
+ UnivNames.register_universe_binders cref ubinders;
Impargs.declare_manual_implicits false (ConstRef proj_cst) [List.hd fieldimpls];
- Universes.register_universe_binders (ConstRef proj_cst) ubinders;
+ UnivNames.register_universe_binders (ConstRef proj_cst) ubinders;
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)
diff --git a/vernac/record.mli b/vernac/record.mli
index 992da2aa5..b2c039f0b 100644
--- a/vernac/record.mli
+++ b/vernac/record.mli
@@ -11,7 +11,6 @@
open Names
open Vernacexpr
open Constrexpr
-open Globnames
val primitive_flag : bool ref
@@ -21,7 +20,7 @@ val declare_projections :
?kind:Decl_kinds.definition_object_kind ->
Id.t ->
bool list ->
- Universes.universe_binders ->
+ UnivNames.universe_binders ->
Impargs.manual_implicits list ->
Context.Rel.t ->
(Name.t * bool) list * Constant.t option list
@@ -30,6 +29,6 @@ val definition_structure :
inductive_kind * Decl_kinds.cumulative_inductive_flag * Decl_kinds.polymorphic *
Declarations.recursivity_kind * ident_decl with_coercion * local_binder_expr list *
(local_decl_expr with_instance with_priority with_notation) list *
- Id.t * constr_expr option -> global_reference
+ Id.t * constr_expr option -> GlobRef.t
-val declare_existing_class : global_reference -> unit
+val declare_existing_class : GlobRef.t -> unit
diff --git a/vernac/search.ml b/vernac/search.ml
index a2a4fb40f..e8ccec11c 100644
--- a/vernac/search.ml
+++ b/vernac/search.ml
@@ -22,8 +22,8 @@ open Nametab
module NamedDecl = Context.Named.Declaration
-type filter_function = global_reference -> env -> constr -> bool
-type display_function = global_reference -> env -> constr -> unit
+type filter_function = GlobRef.t -> env -> constr -> bool
+type display_function = GlobRef.t -> env -> constr -> unit
(* This option restricts the output of [SearchPattern ...],
[SearchAbout ...], etc. to the names of the symbols matching the
@@ -61,7 +61,7 @@ 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 iter_hypothesis glnum (fn : GlobRef.t -> 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
@@ -69,7 +69,7 @@ let iter_hypothesis glnum (fn : global_reference -> env -> constr -> unit) =
iter_named_context_name_type iter_hyp pfctxt
(* General search over declarations *)
-let iter_declarations (fn : global_reference -> env -> constr -> unit) =
+let iter_declarations (fn : GlobRef.t -> env -> constr -> unit) =
let env = Global.env () in
let iter_obj (sp, kn) lobj = match object_tag lobj with
| "VARIABLE" ->
@@ -117,8 +117,7 @@ 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
+ type t = GlobRef.t * Environ.env * Constr.t * priority
and priority = int
module ConstrSet = CSet.Make(Constr)
@@ -216,7 +215,7 @@ 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 (EConstr.of_constr typ)
+ Constr_matching.is_matching_appsubterm ~closed:false env (Evd.from_env env) pat (EConstr.of_constr typ)
| GlobSearchString s ->
String.string_contains ~where:(name_of_reference gr) ~what:s
@@ -227,7 +226,7 @@ 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 Evd.empty (* FIXME *) (EConstr.of_constr typ) &&
+ pattern_filter pat ref env (Evd.from_env env) (* FIXME *) (EConstr.of_constr typ) &&
blacklist_filter ref env typ
in
let iter ref env typ =
@@ -251,8 +250,8 @@ let search_rewrite 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 pat1 ref env Evd.empty (* FIXME *) (EConstr.of_constr typ) ||
- pattern_filter pat2 ref env Evd.empty (EConstr.of_constr typ)) &&
+ (pattern_filter pat1 ref env (Evd.from_env env) (* FIXME *) (EConstr.of_constr typ) ||
+ pattern_filter pat2 ref env (Evd.from_env env) (EConstr.of_constr typ)) &&
blacklist_filter ref env typ
in
let iter ref env typ =
@@ -266,7 +265,7 @@ 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 Evd.empty (* FIXME *) (EConstr.of_constr typ) &&
+ head_filter pat ref env (Evd.from_env env) (* FIXME *) (EConstr.of_constr typ) &&
blacklist_filter ref env typ
in
let iter ref env typ =
@@ -330,12 +329,12 @@ let interface_search =
toggle (Str.string_match regexp id 0) flag
in
let match_type (pat, flag) =
- toggle (Constr_matching.is_matching env Evd.empty pat (EConstr.of_constr constr)) flag
+ toggle (Constr_matching.is_matching env (Evd.from_env env) pat (EConstr.of_constr constr)) flag
in
let match_subtype (pat, flag) =
toggle
(Constr_matching.is_matching_appsubterm ~closed:false
- env Evd.empty pat (EConstr.of_constr constr)) flag
+ env (Evd.from_env env) pat (EConstr.of_constr constr)) flag
in
let match_module (mdl, flag) =
toggle (Libnames.is_dirpath_prefix_of mdl path) flag
diff --git a/vernac/search.mli b/vernac/search.mli
index a1fb7ed3e..0dc82c1c3 100644
--- a/vernac/search.mli
+++ b/vernac/search.mli
@@ -12,7 +12,6 @@ open Names
open Constr
open Environ
open Pattern
-open Globnames
(** {6 Search facilities. } *)
@@ -20,8 +19,8 @@ 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
+type filter_function = GlobRef.t -> env -> constr -> bool
+type display_function = GlobRef.t -> env -> constr -> unit
(** {6 Generic filter functions} *)
diff --git a/vernac/topfmt.ml b/vernac/topfmt.ml
index 055f6676e..609dac69a 100644
--- a/vernac/topfmt.ml
+++ b/vernac/topfmt.ml
@@ -89,12 +89,14 @@ let set_margin v =
Format.pp_set_margin Format.str_formatter v;
Format.pp_set_margin !std_ft v;
Format.pp_set_margin !deep_ft v;
+ Format.pp_set_margin !err_ft v;
(* Heuristic, based on usage: the column on the right of max_indent
column is 20% of width, capped to 30 characters *)
let m = max (64 * v / 100) (v-30) in
Format.pp_set_max_indent Format.str_formatter m;
Format.pp_set_max_indent !std_ft m;
- Format.pp_set_max_indent !deep_ft m
+ Format.pp_set_max_indent !deep_ft m;
+ Format.pp_set_max_indent !err_ft m
(** Console display of feedback *)
@@ -289,6 +291,14 @@ let init_terminal_output ~color =
let emacs_logger = gen_logger Emacs.quote_info Emacs.quote_warning
(* This is specific to the toplevel *)
+
+type execution_phase =
+ | ParsingCommandLine
+ | Initialization
+ | LoadingPrelude
+ | LoadingRcFile
+ | InteractiveLoop
+
let pr_loc loc =
let fname = loc.Loc.fname in
match fname with
@@ -301,13 +311,28 @@ let pr_loc loc =
int (loc.bp-loc.bol_pos) ++ str"-" ++ int (loc.ep-loc.bol_pos) ++
str":")
-let print_err_exn ?extra any =
+let pr_phase ?loc phase =
+ match phase, loc with
+ | LoadingRcFile, loc ->
+ (* For when all errors go through feedback:
+ str "While loading rcfile:" ++
+ Option.cata (fun loc -> fnl () ++ pr_loc loc) (mt ()) loc *)
+ Option.map pr_loc loc
+ | LoadingPrelude, loc ->
+ Some (str "While loading initial state:" ++ Option.cata (fun loc -> fnl () ++ pr_loc loc) (mt ()) loc)
+ | _, Some loc -> Some (pr_loc loc)
+ | ParsingCommandLine, _
+ | Initialization, _ -> None
+ | InteractiveLoop, _ ->
+ (* Note: interactive messages such as "foo is defined" are not located *)
+ None
+
+let print_err_exn phase any =
let (e, info) = CErrors.push any in
let loc = Loc.get_loc info in
- let msg_loc = Option.cata pr_loc (mt ()) loc in
- let pre_hdr = pr_opt_no_spc (fun x -> x) extra ++ msg_loc in
+ let pre_hdr = pr_phase ?loc phase in
let msg = CErrors.iprint (e, info) ++ fnl () in
- std_logger ~pre_hdr Feedback.Error msg
+ std_logger ?pre_hdr Feedback.Error msg
let with_output_to_file fname func input =
let channel = open_out (String.concat "." [fname; "out"]) in
diff --git a/vernac/topfmt.mli b/vernac/topfmt.mli
index 579b456a2..73dcb0064 100644
--- a/vernac/topfmt.mli
+++ b/vernac/topfmt.mli
@@ -53,8 +53,17 @@ val init_terminal_output : color:bool -> unit
(** Error printing *)
(* To be deprecated when we can fully move to feedback-based error
printing. *)
+
+type execution_phase =
+ | ParsingCommandLine
+ | Initialization
+ | LoadingPrelude
+ | LoadingRcFile
+ | InteractiveLoop
+
val pr_loc : Loc.t -> Pp.t
-val print_err_exn : ?extra:Pp.t -> exn -> unit
+val pr_phase : ?loc:Loc.t -> execution_phase -> Pp.t option
+val print_err_exn : execution_phase -> exn -> unit
(** [with_output_to_file file f x] executes [f x] with logging
redirected to a file [file] *)
diff --git a/vernac/vernac.mllib b/vernac/vernac.mllib
index f001b572a..356951b69 100644
--- a/vernac/vernac.mllib
+++ b/vernac/vernac.mllib
@@ -1,10 +1,18 @@
+Vernacexpr
+Pvernac
+G_vernac
+G_proofs
Vernacprop
-Proof_using
-Lemmas
Himsg
ExplainErr
-Class
Locality
+Egramml
+Vernacinterp
+Ppvernac
+Proof_using
+Lemmas
+Class
+Egramcoq
Metasyntax
Auto_ind_decl
Search
@@ -20,7 +28,8 @@ Classes
Record
Assumptions
Vernacstate
-Vernacinterp
Mltop
Topfmt
Vernacentries
+
+Misctypes
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
index b44c7cccb..94eb45fd3 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -29,7 +29,6 @@ open Decl_kinds
open Constrexpr
open Redexpr
open Lemmas
-open Misctypes
open Locality
open Vernacinterp
@@ -266,7 +265,7 @@ let print_namespace ns =
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 = (Global.env ()).Environ.env_globals.Environ.env_constants in
let constants_in_namespace =
Cmap_env.fold (fun c (body,_) acc ->
let kn = Constant.user c in
@@ -449,7 +448,7 @@ let start_proof_and_print k l hook =
let evi = Evarutil.nf_evar_info sigma evi in
let env = Evd.evar_filtered_env evi in
try
- let concl = EConstr.of_constr evi.Evd.evar_concl in
+ let concl = evi.Evd.evar_concl in
if not (Evarutil.is_ground_env sigma env &&
Evarutil.is_ground_term sigma concl)
then raise Exit;
@@ -518,7 +517,7 @@ 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 = Pfedit.by (Tactics.exact_proof c) in
- save_proof (Vernacexpr.(Proved(Opaque,None)));
+ save_proof (Vernacexpr.(Proved(Proof_global.Opaque,None)));
if not status then Feedback.feedback Feedback.AddedAxiom
let vernac_assumption ~atts discharge kind l nl =
@@ -637,7 +636,7 @@ let vernac_scheme l =
let vernac_combined_scheme lid l =
if Dumpglob.dump () then
(Dumpglob.dump_definition lid false "def";
- List.iter (fun {loc;v=id} -> dump_global (make ?loc @@ Misctypes.AN (make ?loc @@ Ident id))) l);
+ List.iter (fun {loc;v=id} -> dump_global (make ?loc @@ AN (make ?loc @@ Ident id))) l);
Indschemes.do_combined_scheme lid l
let vernac_universe ~atts l =
@@ -672,7 +671,7 @@ let vernac_declare_module export {loc;v=id} binders_ast mty_ast =
else (idl,ty)) binders_ast in
let mp =
Declaremods.declare_module Modintern.interp_module_ast
- id binders_ast (Enforce mty_ast) []
+ id binders_ast (Declaremods.Enforce mty_ast) []
in
Dumpglob.dump_moddef ?loc mp "mod";
Flags.if_verbose Feedback.msg_info (str "Module " ++ Id.print id ++ str " is declared");
@@ -855,7 +854,7 @@ let vernac_identity_coercion ~atts id qids qidt =
let vernac_instance ~atts abst sup inst props pri =
let global = not (make_section_locality atts.locality) in
- Dumpglob.dump_constraint inst false "inst";
+ Dumpglob.dump_constraint (fst (pi1 inst)) false "inst";
let program_mode = Flags.is_program_mode () in
ignore(Classes.new_instance ~program_mode ~abstract:abst ~global atts.polymorphic sup inst props pri)
@@ -909,7 +908,7 @@ let vernac_set_used_variables e =
if List.is_empty to_clear then (p, ())
else
let tac = Tactics.clear to_clear in
- fst (Pfedit.solve SelectAll None tac p), ()
+ fst (Pfedit.solve Goal_select.SelectAll None tac p), ()
end
(*****************************)
@@ -977,7 +976,7 @@ let vernac_remove_hints ~atts dbs ids =
let vernac_hints ~atts lb h =
let local = enforce_module_locality atts.locality in
- Hints.add_hints local lb (Hints.interp_hints atts.polymorphic h)
+ Hints.add_hints ~local lb (Hints.interp_hints atts.polymorphic h)
let vernac_syntactic_definition ~atts lid x y =
Dumpglob.dump_definition lid false "syndef";
@@ -1131,15 +1130,16 @@ let vernac_arguments ~atts reference args more_implicits nargs_for_red flags =
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 " ++ Name.print o ++
- str " renamed to " ++ Name.print n ++ str ".");
+ if !rename_flag_required && not rename_flag then begin
+ let msg =
+ match !example_renaming with
+ | None ->
+ strbrk "To rename arguments the \"rename\" flag must be specified."
+ | Some (o,n) ->
+ strbrk "Flag \"rename\" expected to rename " ++ Name.print o ++
+ strbrk " into " ++ Name.print n ++ str "."
+ in user_err ~hdr:"vernac_declare_arguments" msg
+ end;
let duplicate_names =
List.duplicates Name.equal (List.filter ((!=) Anonymous) names)
@@ -1268,7 +1268,7 @@ let vernac_reserve bl =
let vernac_generalizable ~atts =
let local = make_non_locality atts.locality in
- Implicit_quantifiers.declare_generalizable local
+ Implicit_quantifiers.declare_generalizable ~local
let _ =
declare_bool_option
@@ -1465,22 +1465,22 @@ let _ =
optkey = ["Printing";"Universes"];
optread = (fun () -> !Constrextern.print_universes);
optwrite = (fun b -> Constrextern.print_universes:=b) }
-
+
let _ =
declare_bool_option
{ optdepr = false;
optname = "dumping bytecode after compilation";
optkey = ["Dump";"Bytecode"];
- optread = Flags.get_dump_bytecode;
- optwrite = Flags.set_dump_bytecode }
+ optread = (fun () -> !Cbytegen.dump_bytecode);
+ optwrite = (:=) Cbytegen.dump_bytecode }
let _ =
declare_bool_option
{ optdepr = false;
optname = "dumping VM lambda code after compilation";
optkey = ["Dump";"Lambda"];
- optread = Flags.get_dump_lambda;
- optwrite = Flags.set_dump_lambda }
+ optread = (fun () -> !Clambda.dump_lambda);
+ optwrite = (:=) Clambda.dump_lambda }
let _ =
declare_bool_option
@@ -1611,7 +1611,7 @@ let get_current_context_of_args = function
let query_command_selector ?loc = function
| None -> None
- | Some (SelectNth n) -> Some n
+ | Some (Goal_select.SelectNth n) -> Some n
| _ -> user_err ?loc ~hdr:"query_command_selector"
(str "Query commands only support the single numbered goal selector.")
@@ -1619,17 +1619,16 @@ let vernac_check_may_eval ~atts redexp glopt rc =
let glopt = query_command_selector ?loc:atts.loc glopt in
let (sigma, env) = get_current_context_of_args glopt in
let sigma', c = interp_open_constr env sigma rc in
- let c = EConstr.Unsafe.to_constr c 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 sigma' = Evd.minimize_universes sigma' in
let uctx = Evd.universe_context_set sigma' in
let env = Environ.push_context_set uctx (Evarutil.nf_env_evar sigma' env) in
- let c = nf c in
let j =
- if Evarutil.has_undefined_evars sigma' (EConstr.of_constr c) then
- Evarutil.j_nf_evar sigma' (Retyping.get_judgment_of env sigma' (EConstr.of_constr c))
+ if Evarutil.has_undefined_evars sigma' c then
+ Evarutil.j_nf_evar sigma' (Retyping.get_judgment_of env sigma' c)
else
+ let c = EConstr.to_constr sigma' c in
(* OK to call kernel which does not support evars *)
Termops.on_judgment EConstr.of_constr (Arguments_renaming.rename_typing env c)
in
@@ -1652,7 +1651,9 @@ let vernac_check_may_eval ~atts redexp glopt rc =
let vernac_declare_reduction ~atts s r =
let local = make_locality atts.locality in
- declare_red_expr local s (snd (Hook.get f_interp_redexp (Global.env()) Evd.empty r))
+ let env = Global.env () in
+ let sigma = Evd.from_env env in
+ declare_red_expr local s (snd (Hook.get f_interp_redexp env sigma r))
(* The same but avoiding the current goal context if any *)
let vernac_global_check c =
@@ -1742,7 +1743,7 @@ let vernac_print ~atts env sigma =
else str"There may remain asynchronous universe constraints"
in
begin match dst with
- | None -> UGraph.pr_universes Universes.pr_with_global_universes univ ++ pr_remaining
+ | None -> UGraph.pr_universes UnivNames.pr_with_global_universes univ ++ pr_remaining
| Some s -> dump_universes_gen univ s
end
| PrintHint r -> Hints.pr_hint_ref env sigma (smart_global r)
@@ -1912,7 +1913,7 @@ let vernac_subproof gln =
Proof_global.simple_with_current_proof (fun _ p ->
match gln with
| None -> Proof.focus subproof_cond () 1 p
- | Some (SelectNth n) -> Proof.focus subproof_cond () n p
+ | Some (Goal_select.SelectNth n) -> Proof.focus subproof_cond () n p
| _ -> user_err ~hdr:"bracket_selector"
(str "Brackets only support the single numbered goal selector."))
@@ -1970,7 +1971,7 @@ let vernac_load interp fname =
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
+ match Pcoq.Gram.entry_parse Pvernac.main_entry po with
| Some x -> x
| None -> raise End_of_input) in
let fname =
@@ -2244,7 +2245,7 @@ let with_fail st b f =
| HasNotFailed ->
user_err ~hdr:"Fail" (str "The command has not failed!")
| HasFailed msg ->
- if not !Flags.quiet || !Flags.test_mode || !Flags.ide_slave then Feedback.msg_info
+ if not !Flags.quiet || !Flags.test_mode then Feedback.msg_info
(str "The command has indeed failed with message:" ++ fnl () ++ msg)
| _ -> assert false
end
diff --git a/vernac/vernacentries.mli b/vernac/vernacentries.mli
index f6199e820..3c88a3443 100644
--- a/vernac/vernacentries.mli
+++ b/vernac/vernacentries.mli
@@ -8,9 +8,7 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
-open Misctypes
-
-val dump_global : Libnames.reference or_by_notation -> unit
+val dump_global : Libnames.reference Constrexpr.or_by_notation -> unit
(** Vernacular entries *)
val vernac_require :
diff --git a/vernac/vernacexpr.ml b/vernac/vernacexpr.ml
new file mode 100644
index 000000000..5acac9e25
--- /dev/null
+++ b/vernac/vernacexpr.ml
@@ -0,0 +1,522 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+open Names
+open Constrexpr
+open Libnames
+
+(** Vernac expressions, produced by the parser *)
+type class_rawexpr = FunClass | SortClass | RefClass of reference or_by_notation
+
+type goal_selector = Goal_select.t =
+ | SelectAlreadyFocused
+ | SelectNth of int
+ | SelectList of (int * int) list
+ | SelectId of Id.t
+ | SelectAll
+[@@ocaml.deprecated "Use Goal_select.t"]
+
+type goal_identifier = string
+type scope_name = string
+
+type goal_reference =
+ | OpenSubgoals
+ | NthGoal of int
+ | GoalId of Id.t
+
+type univ_name_list = UnivNames.univ_name_list
+[@@ocaml.deprecated "Use [UnivNames.univ_name_list]"]
+
+type printable =
+ | PrintTables
+ | PrintFullContext
+ | PrintSectionContext of reference
+ | PrintInspect of int
+ | PrintGrammar of string
+ | PrintLoadPath of DirPath.t option
+ | PrintModules
+ | PrintModule of reference
+ | PrintModuleType of reference
+ | PrintNamespace of DirPath.t
+ | PrintMLLoadPath
+ | PrintMLModules
+ | PrintDebugGC
+ | PrintName of reference or_by_notation * UnivNames.univ_name_list option
+ | PrintGraph
+ | PrintClasses
+ | PrintTypeClasses
+ | PrintInstances of reference or_by_notation
+ | PrintCoercions
+ | PrintCoercionPaths of class_rawexpr * class_rawexpr
+ | PrintCanonicalConversions
+ | PrintUniverses of bool * string option
+ | PrintHint of reference or_by_notation
+ | PrintHintGoal
+ | PrintHintDbName of string
+ | PrintHintDb
+ | PrintScopes
+ | PrintScope of string
+ | PrintVisibility of string option
+ | PrintAbout of reference or_by_notation * UnivNames.univ_name_list option * Goal_select.t option
+ | PrintImplicit of reference or_by_notation
+ | PrintAssumptions of bool * bool * reference or_by_notation
+ | PrintStrategy of reference or_by_notation option
+
+type search_about_item =
+ | SearchSubPattern of constr_pattern_expr
+ | SearchString of string * scope_name option
+
+type searchable =
+ | SearchPattern of constr_pattern_expr
+ | SearchRewrite of constr_pattern_expr
+ | SearchHead of constr_pattern_expr
+ | SearchAbout of (bool * search_about_item) list
+
+type locatable =
+ | LocateAny of reference or_by_notation
+ | LocateTerm of reference or_by_notation
+ | LocateLibrary of reference
+ | LocateModule of reference
+ | LocateOther of string * reference
+ | LocateFile of string
+
+type showable =
+ | ShowGoal of goal_reference
+ | ShowProof
+ | ShowScript
+ | ShowExistentials
+ | ShowUniverses
+ | ShowProofNames
+ | ShowIntros of bool
+ | ShowMatch of reference
+
+type comment =
+ | CommentConstr of constr_expr
+ | CommentString of string
+ | CommentInt of int
+
+type reference_or_constr = Hints.reference_or_constr =
+ | HintsReference of reference
+ | HintsConstr of constr_expr
+[@@ocaml.deprecated "Please use [Hints.reference_or_constr]"]
+
+type hint_mode = Hints.hint_mode =
+ | ModeInput (* No evars *)
+ | ModeNoHeadEvar (* No evar at the head *)
+ | ModeOutput (* Anything *)
+[@@ocaml.deprecated "Please use [Hints.hint_mode]"]
+
+type 'a hint_info_gen = 'a Typeclasses.hint_info_gen =
+ { hint_priority : int option;
+ hint_pattern : 'a option }
+[@@ocaml.deprecated "Please use [Typeclasses.hint_info_gen]"]
+
+type hint_info_expr = Hints.hint_info_expr
+[@@ocaml.deprecated "Please use [Hints.hint_info_expr]"]
+
+type hints_expr = Hints.hints_expr =
+ | HintsResolve of (Hints.hint_info_expr * bool * Hints.reference_or_constr) list
+ | HintsResolveIFF of bool * reference list * int option
+ | HintsImmediate of Hints.reference_or_constr list
+ | HintsUnfold of reference list
+ | HintsTransparency of reference list * bool
+ | HintsMode of reference * Hints.hint_mode list
+ | HintsConstructors of reference list
+ | HintsExtern of int * constr_expr option * Genarg.raw_generic_argument
+[@@ocaml.deprecated "Please use [Hints.hints_expr]"]
+
+type search_restriction =
+ | SearchInside of reference list
+ | SearchOutside of reference list
+
+type rec_flag = bool (* true = Rec; false = NoRec *)
+type verbose_flag = bool (* true = Verbose; false = Silent *)
+type opacity_flag = Proof_global.opacity_flag = Opaque | Transparent
+ [@ocaml.deprecated "Please use [Proof_global.opacity_flag]"]
+type coercion_flag = bool (* true = AddCoercion false = NoCoercion *)
+type instance_flag = bool option
+ (* Some true = Backward instance; Some false = Forward instance, None = NoInstance *)
+type export_flag = bool (* true = Export; false = Import *)
+type inductive_flag = Declarations.recursivity_kind
+type onlyparsing_flag = Flags.compat_version option
+ (* Some v = Parse only; None = Print also.
+ If v<>Current, it contains the name of the coq version
+ which this notation is trying to be compatible with *)
+type locality_flag = bool (* true = Local *)
+
+type option_value = Goptions.option_value =
+ | BoolValue of bool
+ | IntValue of int option
+ | StringValue of string
+ | StringOptValue of string option
+
+type option_ref_value =
+ | StringRefValue of string
+ | QualidRefValue of reference
+
+(** Identifier and optional list of bound universes and constraints. *)
+
+type sort_expr = Sorts.family
+
+type definition_expr =
+ | ProveBody of local_binder_expr list * constr_expr
+ | DefineBody of local_binder_expr list * Genredexpr.raw_red_expr option * constr_expr
+ * constr_expr option
+
+type fixpoint_expr =
+ ident_decl * (lident option * recursion_order_expr) * local_binder_expr list * constr_expr * constr_expr option
+
+type cofixpoint_expr =
+ ident_decl * local_binder_expr list * constr_expr * constr_expr option
+
+type local_decl_expr =
+ | AssumExpr of lname * constr_expr
+ | DefExpr of lname * constr_expr * constr_expr option
+
+type inductive_kind = Inductive_kw | CoInductive | Variant | Record | Structure | Class of bool (* true = definitional, false = inductive *)
+type decl_notation = lstring * constr_expr * scope_name option
+type simple_binder = lident list * constr_expr
+type class_binder = lident * constr_expr list
+type 'a with_coercion = coercion_flag * 'a
+type 'a with_instance = instance_flag * 'a
+type 'a with_notation = 'a * decl_notation list
+type 'a with_priority = 'a * int option
+type constructor_expr = (lident * constr_expr) with_coercion
+type constructor_list_or_record_decl_expr =
+ | Constructors of constructor_expr list
+ | RecordDecl of lident option * local_decl_expr with_instance with_priority with_notation list
+type inductive_expr =
+ ident_decl with_coercion * local_binder_expr list * constr_expr option * inductive_kind *
+ constructor_list_or_record_decl_expr
+
+type one_inductive_expr =
+ ident_decl * local_binder_expr list * constr_expr option * constructor_expr list
+
+type typeclass_constraint = name_decl * Decl_kinds.binding_kind * constr_expr
+and typeclass_context = typeclass_constraint list
+
+type proof_expr =
+ ident_decl * (local_binder_expr list * constr_expr)
+
+type syntax_modifier =
+ | SetItemLevel of string list * Extend.production_level
+ | SetItemLevelAsBinder of string list * Notation_term.constr_as_binder_kind * Extend.production_level option
+ | SetLevel of int
+ | SetAssoc of Extend.gram_assoc
+ | SetEntryType of string * Extend.simple_constr_prod_entry_key
+ | SetOnlyParsing
+ | SetOnlyPrinting
+ | SetCompatVersion of Flags.compat_version
+ | SetFormat of string * lstring
+
+type proof_end =
+ | Admitted
+ (* name in `Save ident` when closing goal *)
+ | Proved of Proof_global.opacity_flag * lident option
+
+type scheme =
+ | InductionScheme of bool * reference or_by_notation * sort_expr
+ | CaseScheme of bool * reference or_by_notation * sort_expr
+ | EqualityScheme of reference or_by_notation
+
+type section_subset_expr =
+ | SsEmpty
+ | SsType
+ | SsSingl of lident
+ | SsCompl of section_subset_expr
+ | SsUnion of section_subset_expr * section_subset_expr
+ | SsSubstr of section_subset_expr * section_subset_expr
+ | SsFwdClose of section_subset_expr
+
+(** Extension identifiers for the VERNAC EXTEND mechanism.
+
+ {b ("Extraction", 0} indicates {b Extraction {i qualid}} command.
+
+ {b ("Extraction", 1} indicates {b Recursive Extraction {i qualid}} command.
+
+ {b ("Extraction", 2)} indicates {b Extraction "{i filename}" {i qualid{_ 1}} ... {i qualid{_ n}}} command.
+
+ {b ("ExtractionLibrary", 0)} indicates {b Extraction Library {i ident}} command.
+
+ {b ("RecursiveExtractionLibrary", 0)} indicates {b Recursive Extraction Library {i ident}} command.
+
+ {b ("SeparateExtraction", 0)} indicates {b Separate Extraction {i qualid{_ 1}} ... {i qualid{_ n}}} command.
+
+ {b ("ExtractionLanguage", 0)} indicates {b Extraction Language Ocaml} or {b Extraction Language Haskell} or {b Extraction Language Scheme} or {b Extraction Language JSON} commands.
+
+ {b ("ExtractionImplicit", 0)} indicates {b Extraction Implicit {i qualid} \[ {i ident{_1}} ... {i ident{_n} } \] } command.
+
+ {b ("ExtractionConstant", 0)} indicates {b Extract Constant {i qualid} => {i string}} command.
+
+ {b ("ExtractionInlinedConstant", 0)} indicates {b Extract Inlined Constant {i qualid} => {i string}} command.
+
+ {b ("ExtractionInductive", 0)} indicates {b Extract Inductive {i qualid} => {i string} [ {i string} ... {string} ] {i optstring}} command.
+
+ {b ("ExtractionBlacklist", 0)} indicates {b Extraction Blacklist {i ident{_1}} ... {i ident{_n}}} command.
+ *)
+type extend_name =
+ (** Name of the vernac entry where the tactic is defined, typically found
+ after the VERNAC EXTEND statement in the source. *)
+ string *
+ (** Index of the extension in the VERNAC EXTEND statement. Each parsing branch
+ is given an offset, starting from zero. *)
+ int
+
+(* This type allows registering the inlining of constants in native compiler.
+ It will be extended with primitive inductive types and operators *)
+type register_kind =
+ | RegisterInline
+
+type bullet = Proof_bullet.t
+[@@ocaml.deprecated "Alias type, please use [Proof_bullet.t]"]
+
+(** {6 Types concerning the module layer} *)
+
+(** Rigid / flexible module signature *)
+
+type 'a module_signature = 'a Declaremods.module_signature =
+ | Enforce of 'a (** ... : T *)
+ | Check of 'a list (** ... <: T1 <: T2, possibly empty *)
+[@@ocaml.deprecated "please use [Declaremods.module_signature]."]
+
+(** Which module inline annotations should we honor,
+ either None or the ones whose level is less or equal
+ to the given integer *)
+
+type inline = Declaremods.inline =
+ | NoInline
+ | DefaultInline
+ | InlineAt of int
+[@@ocaml.deprecated "please use [Declaremods.inline]."]
+
+type module_ast_inl = module_ast * Declaremods.inline
+type module_binder = bool option * lident list * module_ast_inl
+
+(** [Some b] if locally enabled/disabled according to [b], [None] if
+ we should use the global flag. *)
+type vernac_cumulative = VernacCumulative | VernacNonCumulative
+
+(** {6 The type of vernacular expressions} *)
+
+type vernac_implicit_status = Implicit | MaximallyImplicit | NotImplicit
+
+type vernac_argument_status = {
+ name : Name.t;
+ recarg_like : bool;
+ notation_scope : string CAst.t option;
+ implicit_status : vernac_implicit_status;
+}
+
+type nonrec vernac_expr =
+
+ | VernacLoad of verbose_flag * string
+ (* Syntax *)
+ | VernacSyntaxExtension of bool * (lstring * syntax_modifier list)
+ | VernacOpenCloseScope of bool * scope_name
+ | VernacDelimiters of scope_name * string option
+ | VernacBindScope of scope_name * class_rawexpr list
+ | VernacInfix of (lstring * syntax_modifier list) *
+ constr_expr * scope_name option
+ | VernacNotation of
+ constr_expr * (lstring * syntax_modifier list) *
+ scope_name option
+ | VernacNotationAddFormat of string * string * string
+
+ (* Gallina *)
+ | VernacDefinition of (Decl_kinds.discharge * Decl_kinds.definition_object_kind) * name_decl * definition_expr
+ | VernacStartTheoremProof of Decl_kinds.theorem_kind * proof_expr list
+ | VernacEndProof of proof_end
+ | VernacExactProof of constr_expr
+ | VernacAssumption of (Decl_kinds.discharge * Decl_kinds.assumption_object_kind) *
+ Declaremods.inline * (ident_decl list * constr_expr) with_coercion list
+ | VernacInductive of vernac_cumulative option * Decl_kinds.private_flag * inductive_flag * (inductive_expr * decl_notation list) list
+ | VernacFixpoint of Decl_kinds.discharge * (fixpoint_expr * decl_notation list) list
+ | VernacCoFixpoint of Decl_kinds.discharge * (cofixpoint_expr * decl_notation list) list
+ | VernacScheme of (lident option * scheme) list
+ | VernacCombinedScheme of lident * lident list
+ | VernacUniverse of lident list
+ | VernacConstraint of Glob_term.glob_constraint list
+
+ (* Gallina extensions *)
+ | VernacBeginSection of lident
+ | VernacEndSegment of lident
+ | VernacRequire of
+ reference option * export_flag option * reference list
+ | VernacImport of export_flag * reference list
+ | VernacCanonical of reference or_by_notation
+ | VernacCoercion of reference or_by_notation *
+ class_rawexpr * class_rawexpr
+ | VernacIdentityCoercion of lident * class_rawexpr * class_rawexpr
+ | VernacNameSectionHypSet of lident * section_subset_expr
+
+ (* Type classes *)
+ | VernacInstance of
+ bool * (* abstract instance *)
+ local_binder_expr list * (* super *)
+ typeclass_constraint * (* instance name, class name, params *)
+ (bool * constr_expr) option * (* props *)
+ Hints.hint_info_expr
+
+ | VernacContext of local_binder_expr list
+
+ | VernacDeclareInstances of
+ (reference * Hints.hint_info_expr) list (* instances names, priorities and patterns *)
+
+ | VernacDeclareClass of reference (* inductive or definition name *)
+
+ (* Modules and Module Types *)
+ | VernacDeclareModule of bool option * lident *
+ module_binder list * module_ast_inl
+ | VernacDefineModule of bool option * lident * module_binder list *
+ module_ast_inl Declaremods.module_signature * module_ast_inl list
+ | VernacDeclareModuleType of lident *
+ module_binder list * module_ast_inl list * module_ast_inl list
+ | VernacInclude of module_ast_inl list
+
+ (* Solving *)
+
+ | VernacSolveExistential of int * constr_expr
+
+ (* Auxiliary file and library management *)
+ | VernacAddLoadPath of rec_flag * string * DirPath.t option
+ | VernacRemoveLoadPath of string
+ | VernacAddMLPath of rec_flag * string
+ | VernacDeclareMLModule of string list
+ | VernacChdir of string option
+
+ (* State management *)
+ | VernacWriteState of string
+ | VernacRestoreState of string
+
+ (* Resetting *)
+ | VernacResetName of lident
+ | VernacResetInitial
+ | VernacBack of int
+ | VernacBackTo of int
+
+ (* Commands *)
+ | VernacCreateHintDb of string * bool
+ | VernacRemoveHints of string list * reference list
+ | VernacHints of string list * Hints.hints_expr
+ | VernacSyntacticDefinition of lident * (Id.t list * constr_expr) *
+ onlyparsing_flag
+ | VernacArguments of reference or_by_notation *
+ vernac_argument_status list (* Main arguments status list *) *
+ (Name.t * vernac_implicit_status) list list (* Extra implicit status lists *) *
+ int option (* Number of args to trigger reduction *) *
+ [ `ReductionDontExposeCase | `ReductionNeverUnfold | `Rename |
+ `ExtraScopes | `Assert | `ClearImplicits | `ClearScopes |
+ `DefaultImplicits ] list
+ | VernacReserve of simple_binder list
+ | VernacGeneralizable of (lident list) option
+ | VernacSetOpacity of (Conv_oracle.level * reference or_by_notation list)
+ | VernacSetStrategy of
+ (Conv_oracle.level * reference or_by_notation list) list
+ | VernacUnsetOption of export_flag * Goptions.option_name
+ | VernacSetOption of export_flag * Goptions.option_name * option_value
+ | VernacAddOption of Goptions.option_name * option_ref_value list
+ | VernacRemoveOption of Goptions.option_name * option_ref_value list
+ | VernacMemOption of Goptions.option_name * option_ref_value list
+ | VernacPrintOption of Goptions.option_name
+ | VernacCheckMayEval of Genredexpr.raw_red_expr option * Goal_select.t option * constr_expr
+ | VernacGlobalCheck of constr_expr
+ | VernacDeclareReduction of string * Genredexpr.raw_red_expr
+ | VernacPrint of printable
+ | VernacSearch of searchable * Goal_select.t option * search_restriction
+ | VernacLocate of locatable
+ | VernacRegister of lident * register_kind
+ | VernacComments of comment list
+
+ (* Proof management *)
+ | VernacAbort of lident option
+ | VernacAbortAll
+ | VernacRestart
+ | VernacUndo of int
+ | VernacUndoTo of int
+ | VernacFocus of int option
+ | VernacUnfocus
+ | VernacUnfocused
+ | VernacBullet of Proof_bullet.t
+ | VernacSubproof of Goal_select.t option
+ | VernacEndSubproof
+ | VernacShow of showable
+ | VernacCheckGuard
+ | VernacProof of Genarg.raw_generic_argument option * section_subset_expr option
+ | VernacProofMode of string
+
+ (* For extension *)
+ | VernacExtend of extend_name * Genarg.raw_generic_argument list
+
+type nonrec vernac_flag =
+ | VernacProgram
+ | VernacPolymorphic of bool
+ | VernacLocal of bool
+
+type vernac_control =
+ | VernacExpr of vernac_flag list * vernac_expr
+ (* boolean is true when the `-time` batch-mode command line flag was set.
+ the flag is used to print differently in `-time` vs `Time foo` *)
+ | VernacTime of bool * vernac_control CAst.t
+ | VernacRedirect of string * vernac_control CAst.t
+ | VernacTimeout of int * vernac_control
+ | VernacFail of vernac_control
+
+(* A vernac classifier provides information about the exectuion of a
+ command:
+
+ - vernac_when: encodes if the vernac may alter the parser [thus
+ forcing immediate execution], or if indeed it is pure and parsing
+ can continue without its execution.
+
+ - vernac_type: if it is starts, ends, continues a proof or
+ alters the global state or is a control command like BackTo or is
+ a query like Check.
+
+ The classification works on the assumption that we have 3 states:
+ parsing, execution (global enviroment, etc...), and proof
+ state. For example, commands that only alter the proof state are
+ considered safe to delegate to a worker.
+
+*)
+type vernac_type =
+ (* Start of a proof *)
+ | VtStartProof of vernac_start
+ (* Command altering the global state, bad for parallel
+ processing. *)
+ | VtSideff of vernac_sideff_type
+ (* End of a proof *)
+ | VtQed of vernac_qed_type
+ (* A proof step *)
+ | VtProofStep of proof_step
+ (* To be removed *)
+ | VtProofMode of string
+ (* Queries are commands assumed to be "pure", that is to say, they
+ don't modify the interpretation state. *)
+ | VtQuery
+ (* To be removed *)
+ | VtMeta
+ | VtUnknown
+and vernac_qed_type = VtKeep | VtKeepAsAxiom | VtDrop (* Qed/Admitted, Abort *)
+and vernac_start = string * opacity_guarantee * Id.t list
+and vernac_sideff_type = Id.t list
+and opacity_guarantee =
+ | GuaranteesOpacity (** Only generates opaque terms at [Qed] *)
+ | Doesn'tGuaranteeOpacity (** May generate transparent terms even with [Qed].*)
+and proof_step = { (* TODO: inline with OCaml 4.03 *)
+ parallel : [ `Yes of solving_tac * anon_abstracting_tac | `No ];
+ proof_block_detection : proof_block_name option
+}
+and solving_tac = bool (* a terminator *)
+and anon_abstracting_tac = bool (* abstracting anonymously its result *)
+and proof_block_name = string (* open type of delimiters *)
+type vernac_when =
+ | VtNow
+ | VtLater
+type vernac_classification = vernac_type * vernac_when