aboutsummaryrefslogtreecommitdiffhomepage
path: root/vernac
diff options
context:
space:
mode:
Diffstat (limited to 'vernac')
-rw-r--r--vernac/auto_ind_decl.ml2
-rw-r--r--vernac/classes.ml185
-rw-r--r--vernac/classes.mli1
-rw-r--r--vernac/comAssumption.ml182
-rw-r--r--vernac/comAssumption.mli34
-rw-r--r--vernac/comDefinition.ml132
-rw-r--r--vernac/comDefinition.mli30
-rw-r--r--vernac/comFixpoint.ml356
-rw-r--r--vernac/comFixpoint.mli93
-rw-r--r--vernac/comInductive.ml455
-rw-r--r--vernac/comInductive.mli65
-rw-r--r--vernac/comProgramFixpoint.ml342
-rw-r--r--vernac/comProgramFixpoint.mli12
-rw-r--r--vernac/command.ml1346
-rw-r--r--vernac/command.mli163
-rw-r--r--vernac/declareDef.ml3
-rw-r--r--vernac/explainErr.ml4
-rw-r--r--vernac/himsg.ml22
-rw-r--r--vernac/indschemes.ml6
-rw-r--r--vernac/lemmas.ml131
-rw-r--r--vernac/lemmas.mli9
-rw-r--r--vernac/locality.ml57
-rw-r--r--vernac/locality.mli13
-rw-r--r--vernac/metasyntax.ml366
-rw-r--r--vernac/metasyntax.mli6
-rw-r--r--vernac/mltop.ml36
-rw-r--r--vernac/mltop.mli24
-rw-r--r--vernac/obligations.ml12
-rw-r--r--vernac/obligations.mli3
-rw-r--r--vernac/record.ml147
-rw-r--r--vernac/record.mli13
-rw-r--r--vernac/topfmt.ml20
-rw-r--r--vernac/vernac.mllib6
-rw-r--r--vernac/vernacentries.ml475
-rw-r--r--vernac/vernacentries.mli4
-rw-r--r--vernac/vernacinterp.ml1
-rw-r--r--vernac/vernacinterp.mli1
-rw-r--r--vernac/vernacprop.ml48
-rw-r--r--vernac/vernacprop.mli19
-rw-r--r--vernac/vernacstate.ml30
-rw-r--r--vernac/vernacstate.mli2
41 files changed, 2553 insertions, 2303 deletions
diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml
index 51dd5cd4f..ec6b62ee2 100644
--- a/vernac/auto_ind_decl.ml
+++ b/vernac/auto_ind_decl.ml
@@ -317,7 +317,7 @@ 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 = Decl_kinds.CoFinite then
+ if mib.mind_finite = CoFinite then
raise NoDecidabilityCoInductive;
let fix = mkFix (((Array.make nb_ind 0),i),(names,types,cores)) in
create_input fix),
diff --git a/vernac/classes.ml b/vernac/classes.ml
index b80741269..4a2dba859 100644
--- a/vernac/classes.ml
+++ b/vernac/classes.ml
@@ -8,10 +8,7 @@
(*i*)
open Names
-open Term
-open Constr
-open Vars
-open Environ
+open EConstr
open Nametab
open CErrors
open Util
@@ -70,10 +67,9 @@ let existing_instance glob g info =
let c = global g in
let info = Option.default Hints.empty_hint_info info in
let instance, _ = Global.type_of_global_in_context (Global.env ()) c in
- let _, r = decompose_prod_assum instance in
+ let _, r = Term.decompose_prod_assum instance in
match class_of_constr Evd.empty (EConstr.of_constr r) with
- | Some (_, ((tc,u), _)) -> add_instance (new_instance tc info glob
- (*FIXME*) (Flags.use_polymorphic_flag ()) c)
+ | Some (_, ((tc,u), _)) -> add_instance (new_instance tc info glob c)
| None -> user_err ?loc:(loc_of_reference g)
~hdr:"declare_instance"
(Pp.str "Constant does not build instances of a declared type class.")
@@ -83,19 +79,20 @@ let mismatched_props env n m = mismatched_ctx_inst env Properties n m
(* Declare everything in the parameters as implicit, and the class instance as well *)
-let type_ctx_instance evars env ctx inst subst =
- let rec aux (subst, instctx) l = function
+let type_ctx_instance env sigma ctx inst subst =
+ let open Vars in
+ let rec aux (sigma, subst, instctx) l = function
decl :: ctx ->
let t' = substl subst (RelDecl.get_type decl) in
- let c', l =
+ let (sigma, c'), l =
match decl with
- | LocalAssum _ -> EConstr.Unsafe.to_constr (interp_casted_constr_evars env evars (List.hd l) t'), List.tl l
- | LocalDef (_,b,_) -> substl subst b, l
+ | LocalAssum _ -> interp_casted_constr_evars env sigma (List.hd l) t', List.tl l
+ | LocalDef (_,b,_) -> (sigma, substl subst b), l
in
let d = RelDecl.get_name decl, Some c', t' in
- aux (c' :: subst, d :: instctx) l ctx
- | [] -> subst
- in aux (subst, []) inst (List.rev ctx)
+ aux (sigma, c' :: subst, d :: instctx) l ctx
+ | [] -> sigma, subst
+ in aux (sigma, subst, []) inst (List.rev ctx)
let id_of_class cl =
match cl.cl_impl with
@@ -112,38 +109,38 @@ let instance_hook k info global imps ?hook cst =
Typeclasses.declare_instance (Some info) (not global) cst;
(match hook with Some h -> h cst | None -> ())
-let declare_instance_constant k info global imps ?hook id decl poly evm term termtype =
+let declare_instance_constant k info global imps ?hook id decl poly sigma term termtype =
let kind = IsDefinition Instance in
- let evm =
- let levels = Univ.LSet.union (Univops.universes_of_constr termtype)
- (Univops.universes_of_constr term) in
- Evd.restrict_universe_context evm levels
+ let sigma =
+ let env = Global.env () in
+ let levels = Univ.LSet.union (Univops.universes_of_constr env termtype)
+ (Univops.universes_of_constr env term) in
+ Evd.restrict_universe_context sigma levels
in
- let uctx = Evd.check_univ_decl ~poly evm decl in
+ let uctx = Evd.check_univ_decl ~poly sigma decl in
let entry =
Declare.definition_entry ~types:termtype ~univs:uctx term
in
let cdecl = (DefinitionEntry entry, kind) in
let kn = Declare.declare_constant id cdecl in
Declare.definition_message id;
- Universes.register_universe_binders (ConstRef kn) (Evd.universe_binders evm);
+ Declare.declare_univ_binders (ConstRef kn) (Evd.universe_binders sigma);
instance_hook k info global imps ?hook (ConstRef kn);
id
let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance)
- poly ctx (instid, bk, cl) props ?(generalize=true)
+ ~program_mode poly ctx (instid, bk, cl) props ?(generalize=true)
?(tac:unit Proofview.tactic option) ?hook pri =
let env = Global.env() in
let ((loc, instid), pl) = instid in
- let evd, decl = Univdecls.interp_univ_decl_opt env pl in
- let evars = ref evd in
+ let sigma, decl = Univdecls.interp_univ_decl_opt env pl in
let tclass, ids =
match bk with
| Decl_kinds.Implicit ->
Implicit_quantifiers.implicit_application Id.Set.empty ~allow_partial:false
(fun avoid (clname, _) ->
match clname with
- | Some (cl, b) ->
+ | Some cl ->
let t = CAst.make @@ CHole (None, Misctypes.IntroAnonymous, None) in
t, avoid
| None -> failwith ("new instance: under-applied typeclass"))
@@ -154,26 +151,26 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance)
if generalize then CAst.make @@ CGeneralization (Implicit, Some AbsPi, tclass)
else tclass
in
- let k, u, cty, ctx', ctx, len, imps, subst =
- let impls, ((env', ctx), imps) = interp_context_evars env evars ctx in
- let ctx = List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) ctx in
- let c', imps' = interp_type_evars_impls ~impls env' evars tclass in
- let c' = EConstr.Unsafe.to_constr c' in
+ let sigma, k, u, cty, ctx', ctx, len, imps, subst =
+ let sigma, (impls, ((env', ctx), imps)) = interp_context_evars env sigma ctx in
+ let sigma, (c', imps') = interp_type_evars_impls ~impls env' sigma tclass in
let len = List.length ctx in
let imps = imps @ Impargs.lift_implicits len imps' in
- let ctx', c = decompose_prod_assum c' in
+ let ctx', c = decompose_prod_assum sigma c' in
let ctx'' = ctx' @ ctx in
- let (k, u), args = Typeclasses.dest_class_app (push_rel_context ctx'' env) !evars (EConstr.of_constr c) in
- let u = EConstr.EInstance.kind !evars u in
- let cl = Typeclasses.typeclass_univ_instance (k, u) in
- let _, args =
+ let (k, u), args = Typeclasses.dest_class_app (push_rel_context ctx'' env) sigma c in
+ let u_s = EInstance.kind sigma u in
+ let cl = Typeclasses.typeclass_univ_instance (k, u_s) in
+ let args = List.map of_constr args in
+ let cl_context = List.map (Termops.map_rel_decl of_constr) (snd cl.cl_context) in
+ let _, args =
List.fold_right (fun decl (args, args') ->
match decl with
| LocalAssum _ -> (List.tl args, List.hd args :: args')
- | LocalDef (_,b,_) -> (args, substl args' b :: args'))
- (snd cl.cl_context) (args, [])
+ | LocalDef (_,b,_) -> (args, Vars.substl args' b :: args'))
+ cl_context (args, [])
in
- cl, u, c', ctx', ctx, len, imps, args
+ sigma, cl, u, c', ctx', ctx, len, imps, args
in
let id =
match instid with
@@ -187,9 +184,8 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance)
Namegen.next_global_ident_away i (Termops.vars_of_env env)
in
let env' = push_rel_context ctx env in
- evars := Evarutil.nf_evar_map !evars;
- evars := resolve_typeclasses ~filter:Typeclasses.all_evars ~fail:true env !evars;
- let subst = List.map (fun c -> EConstr.Unsafe.to_constr (Evarutil.nf_evar !evars (EConstr.of_constr c))) subst in
+ let sigma = Evarutil.nf_evar_map sigma in
+ let sigma = resolve_typeclasses ~filter:Typeclasses.all_evars ~fail:true env sigma in
if abstract then
begin
let subst = List.fold_left2
@@ -197,18 +193,16 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance)
[] subst (snd k.cl_context)
in
let (_, ty_constr) = instance_constructor (k,u) subst in
- let nf, subst = Evarutil.e_nf_evars_and_universes evars in
- let termtype =
- let t = it_mkProd_or_LetIn ty_constr (ctx' @ ctx) in
- nf t
- in
- Pretyping.check_evars env Evd.empty !evars (EConstr.of_constr termtype);
- let univs = Evd.check_univ_decl ~poly !evars decl 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 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
(ParameterEntry
(None,(termtype,univs),None), Decl_kinds.IsAssumption Decl_kinds.Logical)
in
- Universes.register_universe_binders (ConstRef cst) (Evd.universe_binders !evars);
+ Declare.declare_univ_binders (ConstRef cst) (Evd.universe_binders sigma);
instance_hook k pri global imps ?hook (ConstRef cst); id
end
else (
@@ -220,16 +214,16 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance)
Some (Inl fs)
| Some (_, t) -> Some (Inr t)
| None ->
- if Flags.is_program_mode () then Some (Inl [])
+ if program_mode then Some (Inl [])
else None
in
- let subst =
+ let subst, sigma =
match props with
- | None -> if List.is_empty k.cl_props then Some (Inl subst) else None
+ | None ->
+ (if List.is_empty k.cl_props then Some (Inl subst) else None), sigma
| Some (Inr term) ->
- let c = interp_casted_constr_evars env' evars term cty in
- let c = EConstr.Unsafe.to_constr c in
- Some (Inr (c, subst))
+ let sigma, c = interp_casted_constr_evars env' sigma term cty in
+ Some (Inr (c, subst)), sigma
| Some (Inl props) ->
let get_id =
function
@@ -266,9 +260,10 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance)
| (n, _) :: _ ->
unbound_method env' k.cl_impl (get_id n)
| _ ->
- Some (Inl (type_ctx_instance evars (push_rel_context ctx' env')
- k.cl_props props subst))
- in
+ let kcl_props = List.map (Termops.map_rel_decl of_constr) k.cl_props in
+ let sigma, res = type_ctx_instance (push_rel_context ctx' env') sigma kcl_props props subst in
+ Some (Inl res), sigma
+ in
let term, termtype =
match subst with
| None -> let termtype = it_mkProd_or_LetIn cty ctx in
@@ -280,34 +275,30 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance)
in
let (app, ty_constr) = instance_constructor (k,u) subst in
let termtype = it_mkProd_or_LetIn ty_constr (ctx' @ ctx) in
- let term = Termops.it_mkLambda_or_LetIn (Option.get app) (ctx' @ ctx) in
+ let term = it_mkLambda_or_LetIn (Option.get app) (ctx' @ ctx) in
Some term, termtype
| Some (Inr (def, subst)) ->
let termtype = it_mkProd_or_LetIn cty ctx in
- let term = Termops.it_mkLambda_or_LetIn def ctx in
+ let term = it_mkLambda_or_LetIn def ctx in
Some term, termtype
in
- let _ =
- evars := Evarutil.nf_evar_map !evars;
- evars := Typeclasses.resolve_typeclasses ~filter:Typeclasses.no_goals_or_obligations ~fail:true
- env !evars;
- (* Try resolving fields that are typeclasses automatically. *)
- evars := Typeclasses.resolve_typeclasses ~filter:Typeclasses.all_evars ~fail:false
- env !evars
- in
- let _ = evars := Evarutil.nf_evar_map_undefined !evars in
- let evm, nf = Evarutil.nf_evar_map_universes !evars in
- let termtype = nf termtype in
- let _ = (* Check that the type is free of evars now. *)
- Pretyping.check_evars env Evd.empty evm (EConstr.of_constr termtype)
- in
- let term = Option.map nf term in
- if not (Evd.has_undefined evm) && not (Option.is_empty term) then
+ let sigma = Evarutil.nf_evar_map sigma in
+ let sigma = Typeclasses.resolve_typeclasses ~filter:Typeclasses.no_goals_or_obligations ~fail:true env sigma in
+ (* Try resolving fields that are typeclasses automatically. *)
+ 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
+ (* Check that the type is free of evars now. *)
+ Pretyping.check_evars env Evd.empty sigma termtype;
+ let termtype = to_constr sigma termtype in
+ let term = Option.map (to_constr 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 evm (Option.get term) termtype
- else if Flags.is_program_mode () || refine || Option.is_empty term then begin
+ poly sigma (Option.get term) termtype
+ else if program_mode || refine || Option.is_empty term then begin
let kind = Decl_kinds.Global, poly, Decl_kinds.DefinitionBody Decl_kinds.Instance in
- if Flags.is_program_mode () then
+ 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];
@@ -317,12 +308,12 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance)
match term with
| Some t ->
let obls, _, constr, typ =
- Obligations.eterm_obligations env id evm 0 t termtype
+ Obligations.eterm_obligations env id sigma 0 t termtype
in obls, Some constr, typ
| None -> [||], None, termtype
in
let hook = Lemmas.mk_hook hook in
- let ctx = Evd.evar_universe_context evm in
+ let ctx = Evd.evar_universe_context sigma in
ignore (Obligations.add_definition id ?term:constr
~univdecl:decl typ ctx ~kind:(Global,poly,Instance) ~hook obls);
id
@@ -333,16 +324,16 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance)
the pretyping after the proof has opened. As a
consequence, we use the low-level primitives to code
the refinement manually.*)
- let gls = List.rev (Evd.future_goals evm) in
- let evm = Evd.reset_future_goals evm in
- Lemmas.start_proof id ~pl:decl kind evm (EConstr.of_constr termtype)
+ let gls = List.rev (Evd.future_goals sigma) in
+ let sigma = Evd.reset_future_goals sigma in
+ Lemmas.start_proof id ~pl:decl kind sigma (EConstr.of_constr termtype)
(Lemmas.mk_hook
(fun _ -> instance_hook k pri global imps ?hook));
(* spiwack: I don't know what to do with the status here. *)
if not (Option.is_empty term) then
let init_refine =
Tacticals.New.tclTHENLIST [
- Refine.refine ~typecheck:false (fun evm -> (evm,EConstr.of_constr (Option.get term)));
+ Refine.refine ~typecheck:false (fun sigma -> (sigma,EConstr.of_constr (Option.get term)));
Proofview.Unsafe.tclNEWGOALS gls;
Tactics.New.reduce_after_refine;
]
@@ -356,6 +347,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance)
else CErrors.user_err Pp.(str "Unsolved obligations remaining."))
let named_of_rel_context l =
+ let open Vars in
let acc, ctx =
List.fold_right
(fun decl (subst, ctx) ->
@@ -369,20 +361,20 @@ let named_of_rel_context l =
let context poly l =
let env = Global.env() in
- let evars = ref (Evd.from_env env) in
- let _, ((env', fullctx), impls) = interp_context_evars env evars l in
- let fullctx = List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) fullctx in
- let subst = Evarutil.evd_comb0 Evarutil.nf_evars_and_universes evars in
- let fullctx = Context.Rel.map subst fullctx in
- let ce t = Pretyping.check_evars env Evd.empty !evars (EConstr.of_constr t) in
+ 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 () = List.iter (fun decl -> Context.Rel.Declaration.iter_constr ce decl) fullctx in
let ctx =
try named_of_rel_context fullctx
with e when CErrors.noncritical e ->
user_err Pp.(str "Anonymous variables not allowed in contexts.")
in
- let uctx = ref (Evd.universe_context_set !evars) in
+ let uctx = ref (Evd.universe_context_set sigma) in
let fn status (id, b, t) =
+ let b, t = Option.map (to_constr sigma) b, to_constr sigma t in
if Lib.is_modtype () && not (Lib.sections_are_opened ()) then
(* Declare the universe context once *)
let univs = if poly
@@ -398,10 +390,9 @@ let context poly l =
(DefinitionEntry entry, IsAssumption Logical)
in
let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id decl in
- match class_of_constr !evars (EConstr.of_constr t) with
+ match class_of_constr sigma (of_constr t) with
| Some (rels, ((tc,_), args) as _cl) ->
- add_instance (Typeclasses.new_instance tc Hints.empty_hint_info false (*FIXME*)
- poly (ConstRef cst));
+ add_instance (Typeclasses.new_instance tc Hints.empty_hint_info false (ConstRef cst));
status
(* declare_subclasses (ConstRef cst) cl *)
| None -> status
@@ -418,7 +409,7 @@ let context poly l =
in
let nstatus = match b with
| None ->
- pi3 (Command.declare_assumption false decl (t, univs) Universes.empty_binders [] impl
+ pi3 (ComAssumption.declare_assumption false decl (t, univs) Universes.empty_binders [] impl
Vernacexpr.NoInline (Loc.tag id))
| Some b ->
let decl = (Discharge, poly, Definition) in
diff --git a/vernac/classes.mli b/vernac/classes.mli
index c0f03227c..d47c6a6f8 100644
--- a/vernac/classes.mli
+++ b/vernac/classes.mli
@@ -41,6 +41,7 @@ val new_instance :
?abstract:bool -> (** Not abstract by default. *)
?global:bool -> (** Not global by default. *)
?refine:bool -> (** Allow refinement *)
+ program_mode:bool ->
Decl_kinds.polymorphic ->
local_binder_expr list ->
Vernacexpr.typeclass_constraint ->
diff --git a/vernac/comAssumption.ml b/vernac/comAssumption.ml
new file mode 100644
index 000000000..5d7adb24a
--- /dev/null
+++ b/vernac/comAssumption.ml
@@ -0,0 +1,182 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Pp
+open CErrors
+open Util
+open Vars
+open Environ
+open Declare
+open Names
+open Globnames
+open Constrexpr_ops
+open Constrintern
+open Impargs
+open Decl_kinds
+open Pretyping
+open Vernacexpr
+open Entries
+
+(* 2| Variable/Hypothesis/Parameter/Axiom declarations *)
+
+let axiom_into_instance = ref false
+
+let _ =
+ let open Goptions in
+ declare_bool_option
+ { optdepr = false;
+ optname = "automatically declare axioms whose type is a typeclass as instances";
+ optkey = ["Typeclasses";"Axioms";"Are";"Instances"];
+ optread = (fun _ -> !axiom_into_instance);
+ optwrite = (:=) axiom_into_instance; }
+
+let should_axiom_into_instance = function
+ | Discharge ->
+ (* The typeclass behaviour of Variable and Context doesn't depend
+ on section status *)
+ true
+ | Global | Local -> !axiom_into_instance
+
+let declare_assumption is_coe (local,p,kind) (c,ctx) pl imps impl nl (_,ident) =
+match local with
+| Discharge when Lib.sections_are_opened () ->
+ let ctx = match ctx with
+ | Monomorphic_const_entry ctx -> ctx
+ | Polymorphic_const_entry ctx -> Univ.ContextSet.of_context ctx
+ in
+ let decl = (Lib.cwd(), SectionLocalAssum ((c,ctx),p,impl), IsAssumption kind) in
+ let _ = declare_variable ident decl in
+ let () = assumption_message ident in
+ let () =
+ if not !Flags.quiet && Proof_global.there_are_pending_proofs () then
+ Feedback.msg_info (str"Variable" ++ spc () ++ Id.print ident ++
+ strbrk " is not visible from current goals")
+ in
+ let r = VarRef ident in
+ let () = Typeclasses.declare_instance None true r in
+ let () = if is_coe then Class.try_add_new_coercion r ~local:true false in
+ (r,Univ.Instance.empty,true)
+
+| Global | Local | Discharge ->
+ let do_instance = should_axiom_into_instance local in
+ let local = DeclareDef.get_locality ident ~kind:"axiom" local in
+ let inl = match nl with
+ | NoInline -> None
+ | DefaultInline -> Some (Flags.get_inline_level())
+ | InlineAt i -> Some i
+ in
+ let decl = (ParameterEntry (None,(c,ctx),inl), IsAssumption kind) in
+ let kn = declare_constant ident ~local decl in
+ let gr = ConstRef kn in
+ let () = maybe_declare_manual_implicits false gr imps in
+ let () = Declare.declare_univ_binders gr pl in
+ let () = assumption_message ident in
+ let () = if do_instance then Typeclasses.declare_instance None false gr in
+ let () = if is_coe then Class.try_add_new_coercion gr ~local p in
+ let inst = match ctx with
+ | Polymorphic_const_entry ctx -> Univ.UContext.instance ctx
+ | Monomorphic_const_entry _ -> Univ.Instance.empty
+ in
+ (gr,inst,Lib.is_modtype_strict ())
+
+let interp_assumption sigma env impls bl c =
+ let c = mkCProdN ?loc:(local_binders_loc bl) bl c in
+ let sigma, (ty, impls) = interp_type_evars_impls env sigma ~impls c in
+ let ty = EConstr.Unsafe.to_constr ty in
+ sigma, (ty, impls)
+
+(* When monomorphic the universe constraints are declared with the first declaration only. *)
+let next_uctx =
+ let empty_uctx = Monomorphic_const_entry Univ.ContextSet.empty in
+ function
+ | Polymorphic_const_entry _ as uctx -> uctx
+ | Monomorphic_const_entry _ -> empty_uctx
+
+let declare_assumptions idl is_coe k (c,uctx) pl imps nl =
+ let refs, status, _ =
+ List.fold_left (fun (refs,status,uctx) id ->
+ let ref',u',status' =
+ declare_assumption is_coe k (c,uctx) pl imps false nl id in
+ (ref',u')::refs, status' && status, next_uctx uctx)
+ ([],true,uctx) idl
+ in
+ List.rev refs, status
+
+
+let maybe_error_many_udecls = function
+ | ((loc,id), Some _) ->
+ user_err ?loc ~hdr:"many_universe_declarations"
+ Pp.(str "When declaring multiple axioms in one command, " ++
+ str "only the first is allowed a universe binder " ++
+ str "(which will be shared by the whole block).")
+ | (_, None) -> ()
+
+let process_assumptions_udecls kind l =
+ let udecl, first_id = match l with
+ | (coe, ((id, udecl)::rest, c))::rest' ->
+ List.iter maybe_error_many_udecls rest;
+ List.iter (fun (coe, (idl, c)) -> List.iter maybe_error_many_udecls idl) rest';
+ udecl, id
+ | (_, ([], _))::_ | [] -> assert false
+ in
+ let () = match kind, udecl with
+ | (Discharge, _, _), Some _ when Lib.sections_are_opened () ->
+ let loc = fst first_id in
+ let msg = Pp.str "Section variables cannot be polymorphic." in
+ user_err ?loc msg
+ | _ -> ()
+ in
+ udecl, List.map (fun (coe, (idl, c)) -> coe, (List.map fst idl, c)) l
+
+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 l =
+ if pi2 kind (* poly *) then
+ (* Separate declarations so that A B : Type puts A and B in different levels. *)
+ List.fold_right (fun (is_coe,(idl,c)) acc ->
+ List.fold_right (fun id acc ->
+ (is_coe, ([id], c)) :: acc) idl acc)
+ l []
+ else l
+ in
+ (* We intepret all declarations in the same evar_map, i.e. as a telescope. *)
+ let (sigma,_,_),l = List.fold_left_map (fun (sigma,env,ienv) (is_coe,(idl,c)) ->
+ let sigma,(t,imps) = interp_assumption sigma env ienv [] c in
+ let env =
+ push_named_context (List.map (fun (_,id) -> LocalAssum (id,t)) idl) env in
+ let ienv = List.fold_right (fun (_,id) ienv ->
+ let impls = compute_internalization_data env Variable t imps in
+ Id.Map.add id impls ienv) idl ienv in
+ ((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
+ (* The universe constraints come from the whole telescope. *)
+ let sigma = Evd.nf_constraints sigma in
+ let nf_evar c = EConstr.to_constr sigma (EConstr.of_constr c) in
+ let uvars, l = List.fold_left_map (fun uvars (coe,t,imps) ->
+ let t = nf_evar t in
+ let uvars = Univ.LSet.union uvars (Univops.universes_of_constr env t) in
+ uvars, (coe,t,imps))
+ Univ.LSet.empty l
+ in
+ let sigma = Evd.restrict_universe_context sigma uvars in
+ let uctx = Evd.check_univ_decl ~poly:(pi2 kind) sigma udecl in
+ let ubinders = Evd.universe_binders sigma in
+ pi2 (List.fold_left (fun (subst,status,uctx) ((is_coe,idl),t,imps) ->
+ 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 (_,id) (c,u) -> (id, Universes.constr_of_global_univ (c,u)))
+ idl refs
+ in
+ subst'@subst, status' && status, next_uctx uctx)
+ ([], true, uctx) l)
diff --git a/vernac/comAssumption.mli b/vernac/comAssumption.mli
new file mode 100644
index 000000000..2fa156abd
--- /dev/null
+++ b/vernac/comAssumption.mli
@@ -0,0 +1,34 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open Constr
+open Entries
+open Globnames
+open Vernacexpr
+open Constrexpr
+open Decl_kinds
+
+(** {6 Parameters/Assumptions} *)
+
+val do_assumptions : locality * polymorphic * assumption_object_kind ->
+ Vernacexpr.inline -> (Vernacexpr.ident_decl list * constr_expr) with_coercion list -> bool
+
+(************************************************************************)
+(** Internal API *)
+(************************************************************************)
+
+(** Exported for Classes *)
+
+(** returns [false] if the assumption is neither local to a section,
+ nor in a module type and meant to be instantiated. *)
+val declare_assumption : coercion_flag -> assumption_kind ->
+ types in_constant_universes_entry ->
+ Universes.universe_binders -> Impargs.manual_implicits ->
+ bool (** implicit *) -> Vernacexpr.inline -> variable Loc.located ->
+ global_reference * Univ.Instance.t * bool
diff --git a/vernac/comDefinition.ml b/vernac/comDefinition.ml
new file mode 100644
index 000000000..d376696f7
--- /dev/null
+++ b/vernac/comDefinition.ml
@@ -0,0 +1,132 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2018 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Pp
+open Util
+open Constr
+open Environ
+open Entries
+open Redexpr
+open Declare
+open Constrintern
+open Pretyping
+
+open Context.Rel.Declaration
+
+(* Commands of the interface: Constant definitions *)
+
+let rec under_binders env sigma f n c =
+ if Int.equal n 0 then f env sigma (EConstr.of_constr c) else
+ match Constr.kind c with
+ | Lambda (x,t,c) ->
+ mkLambda (x,t,under_binders (push_rel (LocalAssum (x,t)) env) sigma f (n-1) c)
+ | LetIn (x,b,t,c) ->
+ mkLetIn (x,b,t,under_binders (push_rel (LocalDef (x,b,t)) env) sigma f (n-1) c)
+ | _ -> assert false
+
+let red_constant_entry n ce sigma = function
+ | None -> ce
+ | Some red ->
+ let proof_out = ce.const_entry_body in
+ let env = Global.env () in
+ let (redfun, _) = reduction_of_red_expr env red in
+ let redfun env sigma c =
+ let (_, c) = redfun env sigma c in
+ EConstr.Unsafe.to_constr c
+ in
+ { ce with const_entry_body = Future.chain proof_out
+ (fun ((body,ctx),eff) -> (under_binders env sigma redfun n body,ctx),eff) }
+
+let warn_implicits_in_term =
+ CWarnings.create ~name:"implicits-in-term" ~category:"implicits"
+ (fun () ->
+ strbrk "Implicit arguments declaration relies on type." ++ spc () ++
+ strbrk "The term declares more implicits than the type here.")
+
+let check_imps ~impsty ~impsbody =
+ let b =
+ try
+ List.for_all (fun (key, (va:bool*bool*bool)) ->
+ (* Pervasives.(=) is OK for this type *)
+ Pervasives.(=) (List.assoc_f Impargs.explicitation_eq key impsty) va)
+ impsbody
+ with Not_found -> false
+ in
+ if not b then warn_implicits_in_term ()
+
+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
+ (* Build the parameters *)
+ let evd, (impls, ((env_bl, ctx), imps1)) = interp_context_evars env evd bl in
+ (* Build the type *)
+ let evd, tyopt = Option.fold_left_map
+ (interp_type_evars_impls ~impls env_bl)
+ evd ctypopt
+ in
+ (* Build the body, and merge implicits from parameters and from type/body *)
+ let evd, c, imps, tyopt =
+ match tyopt with
+ | None ->
+ let evd, (c, impsbody) = interp_constr_evars_impls ~impls env_bl evd c in
+ evd, c, imps1@Impargs.lift_implicits (Context.Rel.nhyps ctx) impsbody, None
+ | Some (ty, impsty) ->
+ let evd, (c, impsbody) = interp_casted_constr_evars_impls ~impls env_bl evd c ty in
+ check_imps ~impsty ~impsbody;
+ evd, c, imps1@Impargs.lift_implicits (Context.Rel.nhyps ctx) impsty, Some ty
+ in
+ (* universe minimization *)
+ let evd = Evd.nf_constraints 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
+ (* Keep only useful universes. *)
+ let uvars_fold uvars c =
+ Univ.LSet.union uvars (universes_of_constr env evd (of_constr c))
+ in
+ let uvars = List.fold_left uvars_fold Univ.LSet.empty (Option.List.cons tyopt [c]) in
+ let evd = Evd.restrict_universe_context evd uvars in
+ (* Check we conform to declared universes *)
+ let uctx = Evd.check_univ_decl ~poly evd decl in
+ (* We're done! *)
+ let ce = definition_entry ?types:tyopt ~univs:uctx c in
+ (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;
+ ce
+
+let do_definition ~program_mode ident k univdecl bl red_option c ctypopt hook =
+ let (ce, evd, univdecl, imps as def) =
+ interp_definition univdecl bl (pi2 k) red_option c ctypopt
+ in
+ if program_mode then
+ let env = Global.env () in
+ let (c,ctx), sideff = Future.force ce.const_entry_body in
+ assert(Safe_typing.empty_private_constants = sideff);
+ assert(Univ.ContextSet.is_empty ctx);
+ let typ = match ce.const_entry_type with
+ | Some t -> t
+ | None -> EConstr.to_constr evd (Retyping.get_type_of env evd (EConstr.of_constr c))
+ in
+ Obligations.check_evars env evd;
+ let obls, _, c, cty =
+ Obligations.eterm_obligations env ident evd 0 c typ
+ in
+ let ctx = Evd.evar_universe_context evd in
+ let hook = Lemmas.mk_hook (fun l r _ -> Lemmas.call_hook (fun exn -> exn) hook l r) in
+ ignore(Obligations.add_definition
+ ident ~term:c cty ctx ~univdecl ~implicits:imps ~kind:k ~hook obls)
+ else let ce = check_definition def in
+ ignore(DeclareDef.declare_definition ident k ce (Evd.universe_binders evd) imps
+ (Lemmas.mk_hook
+ (fun l r -> Lemmas.call_hook (fun exn -> exn) hook l r;r)))
diff --git a/vernac/comDefinition.mli b/vernac/comDefinition.mli
new file mode 100644
index 000000000..4a65c1e91
--- /dev/null
+++ b/vernac/comDefinition.mli
@@ -0,0 +1,30 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2018 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open Entries
+open Decl_kinds
+open Redexpr
+open Constrexpr
+
+(** {6 Definitions/Let} *)
+
+val do_definition : program_mode:bool ->
+ Id.t -> definition_kind -> Vernacexpr.universe_decl_expr option ->
+ local_binder_expr list -> red_expr option -> constr_expr ->
+ constr_expr option -> unit Lemmas.declaration_hook -> unit
+
+(************************************************************************)
+(** Internal API *)
+(************************************************************************)
+
+(** Not used anywhere. *)
+val interp_definition :
+ Vernacexpr.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
diff --git a/vernac/comFixpoint.ml b/vernac/comFixpoint.ml
new file mode 100644
index 000000000..d648e293a
--- /dev/null
+++ b/vernac/comFixpoint.ml
@@ -0,0 +1,356 @@
+open Pp
+open CErrors
+open Util
+open Constr
+open Vars
+open Termops
+open Declare
+open Names
+open Constrexpr
+open Constrexpr_ops
+open Constrintern
+open Decl_kinds
+open Pretyping
+open Evarutil
+open Evarconv
+open Misctypes
+open Vernacexpr
+
+module RelDecl = Context.Rel.Declaration
+
+(* 3c| Fixpoints and co-fixpoints *)
+
+(* An (unoptimized) function that maps preorders to partial orders...
+
+ Input: a list of associations (x,[y1;...;yn]), all yi distincts
+ and different of x, meaning x<=y1, ..., x<=yn
+
+ Output: a list of associations (x,Inr [y1;...;yn]), collecting all
+ distincts yi greater than x, _or_, (x, Inl y) meaning that
+ x is in the same class as y (in which case, x occurs
+ nowhere else in the association map)
+
+ partial_order : ('a * 'a list) list -> ('a * ('a,'a list) union) list
+*)
+
+let rec partial_order cmp = function
+ | [] -> []
+ | (x,xge)::rest ->
+ let rec browse res xge' = function
+ | [] ->
+ let res = List.map (function
+ | (z, Inr zge) when List.mem_f cmp x zge ->
+ (z, Inr (List.union cmp zge xge'))
+ | r -> r) res in
+ (x,Inr xge')::res
+ | y::xge ->
+ let rec link y =
+ try match List.assoc_f cmp y res with
+ | Inl z -> link z
+ | Inr yge ->
+ if List.mem_f cmp x yge then
+ let res = List.remove_assoc_f cmp y res in
+ let res = List.map (function
+ | (z, Inl t) ->
+ if cmp t y then (z, Inl x) else (z, Inl t)
+ | (z, Inr zge) ->
+ if List.mem_f cmp y zge then
+ (z, Inr (List.add_set cmp x (List.remove cmp y zge)))
+ else
+ (z, Inr zge)) res in
+ browse ((y,Inl x)::res) xge' (List.union cmp xge (List.remove cmp x yge))
+ else
+ browse res (List.add_set cmp y (List.union cmp xge' yge)) xge
+ with Not_found -> browse res (List.add_set cmp y xge') xge
+ in link y
+ in browse (partial_order cmp rest) [] xge
+
+let non_full_mutual_message x xge y yge isfix rest =
+ let reason =
+ if Id.List.mem x yge then
+ Id.print y ++ str " depends on " ++ Id.print x ++ strbrk " but not conversely"
+ else if Id.List.mem y xge then
+ Id.print x ++ str " depends on " ++ Id.print y ++ strbrk " but not conversely"
+ else
+ Id.print y ++ str " and " ++ Id.print x ++ strbrk " are not mutually dependent" in
+ let e = if List.is_empty rest then reason else strbrk "e.g., " ++ reason in
+ let k = if isfix then "fixpoint" else "cofixpoint" in
+ let w =
+ if isfix
+ then strbrk "Well-foundedness check may fail unexpectedly." ++ fnl()
+ else mt () in
+ strbrk "Not a fully mutually defined " ++ str k ++ fnl () ++
+ str "(" ++ e ++ str ")." ++ fnl () ++ w
+
+let warn_non_full_mutual =
+ CWarnings.create ~name:"non-full-mutual" ~category:"fixpoints"
+ (fun (x,xge,y,yge,isfix,rest) ->
+ non_full_mutual_message x xge y yge isfix rest)
+
+let check_mutuality env evd isfix fixl =
+ let names = List.map fst fixl in
+ let preorder =
+ List.map (fun (id,def) ->
+ (id, List.filter (fun id' -> not (Id.equal id id') && occur_var env evd id' (EConstr.of_constr def)) names))
+ fixl in
+ let po = partial_order Id.equal preorder in
+ match List.filter (function (_,Inr _) -> true | _ -> false) po with
+ | (x,Inr xge)::(y,Inr yge)::rest ->
+ warn_non_full_mutual (x,xge,y,yge,isfix,rest)
+ | _ -> ()
+
+type structured_fixpoint_expr = {
+ fix_name : Id.t;
+ fix_univs : universe_decl_expr option;
+ fix_annot : Id.t Loc.located option;
+ fix_binders : local_binder_expr list;
+ fix_body : constr_expr option;
+ fix_type : constr_expr
+}
+
+let interp_fix_context ~cofix env sigma fix =
+ let before, after = if not cofix then split_at_annot fix.fix_binders fix.fix_annot else [], fix.fix_binders in
+ let sigma, (impl_env, ((env', ctx), imps)) = interp_context_evars env sigma before in
+ let sigma, (impl_env', ((env'', ctx'), imps')) = interp_context_evars ~impl_env ~shift:(Context.Rel.nhyps ctx) env' sigma after in
+ let annot = Option.map (fun _ -> List.length (assums_of_rel_context ctx)) fix.fix_annot in
+ sigma, ((env'', ctx' @ ctx), (impl_env',imps @ imps'), annot)
+
+let interp_fix_ccl sigma impls (env,_) fix =
+ interp_type_evars_impls ~impls env sigma fix.fix_type
+
+let interp_fix_body env_rec sigma impls (_,ctx) fix ccl =
+ let open EConstr in
+ Option.cata (fun body ->
+ let env = push_rel_context ctx env_rec in
+ let sigma, body = interp_casted_constr_evars env sigma ~impls body ccl in
+ sigma, Some (it_mkLambda_or_LetIn body ctx)) (sigma, None) fix.fix_body
+
+let build_fix_type (_,ctx) ccl = EConstr.it_mkProd_or_LetIn ccl ctx
+
+let prepare_recursive_declaration fixnames fixtypes fixdefs =
+ let defs = List.map (subst_vars (List.rev fixnames)) fixdefs in
+ let names = List.map (fun id -> Name id) fixnames in
+ (Array.of_list names, Array.of_list fixtypes, Array.of_list defs)
+
+(* Jump over let-bindings. *)
+
+let compute_possible_guardness_evidences (ctx,_,recindex) =
+ (* A recursive index is characterized by the number of lambdas to
+ skip before finding the relevant inductive argument *)
+ match recindex with
+ | Some i -> [i]
+ | None ->
+ (* If recursive argument was not given by user, we try all args.
+ An earlier approach was to look only for inductive arguments,
+ but doing it properly involves delta-reduction, and it finally
+ doesn't seem to worth the effort (except for huge mutual
+ fixpoints ?) *)
+ List.interval 0 (Context.Rel.nhyps ctx - 1)
+
+type recursive_preentry =
+ Id.t list * constr option list * types list
+
+(* Wellfounded definition *)
+
+let contrib_name = "Program"
+let subtac_dir = [contrib_name]
+let tactics_module = subtac_dir @ ["Tactics"]
+
+let init_constant dir s sigma =
+ Evarutil.new_global sigma (Coqlib.coq_reference "Command" dir s)
+
+let fix_proto = init_constant tactics_module "fix_proto"
+
+let interp_recursive ~program_mode ~cofix fixl notations =
+ let open Context.Named.Declaration in
+ let open EConstr in
+ let env = Global.env() in
+ let fixnames = List.map (fun fix -> fix.fix_name) fixl in
+
+ (* Interp arities allowing for unresolved types *)
+ let all_universes =
+ List.fold_right (fun sfe acc ->
+ match sfe.fix_univs , acc with
+ | None , acc -> acc
+ | x , None -> x
+ | Some ls , Some us ->
+ let lsu = ls.univdecl_instance and usu = us.univdecl_instance in
+ if not (CList.for_all2eq (fun x y -> Id.equal (snd x) (snd y)) 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, (fixctxs, fiximppairs, fixannots) =
+ on_snd List.split3 @@
+ List.fold_left_map (fun sigma -> interp_fix_context env sigma ~cofix) sigma fixl in
+ let fixctximpenvs, fixctximps = List.split fiximppairs in
+ let sigma, (fixccls,fixcclimps) =
+ on_snd List.split @@
+ List.fold_left3_map interp_fix_ccl sigma fixctximpenvs fixctxs fixl in
+ let fixtypes = List.map2 build_fix_type fixctxs fixccls in
+ let fixtypes = List.map (fun c -> nf_evar sigma c) fixtypes in
+ let fiximps = List.map3
+ (fun ctximps cclimps (_,ctx) -> ctximps@(Impargs.lift_implicits (Context.Rel.nhyps ctx) cclimps))
+ fixctximps fixcclimps fixctxs in
+ let sigma, rec_sign =
+ List.fold_left2
+ (fun (sigma, env') id t ->
+ if program_mode then
+ let sigma, sort = Typing.type_of ~refresh:true env sigma t in
+ let sigma, fixprot =
+ 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
+ with e when CErrors.noncritical e -> sigma, t
+ in
+ sigma, LocalAssum (id,fixprot) :: env'
+ else sigma, LocalAssum (id,t) :: env')
+ (sigma,[]) fixnames fixtypes
+ in
+ let env_rec = push_named_context rec_sign env in
+
+ (* Get interpretation metadatas *)
+ let fixtypes = List.map EConstr.Unsafe.to_constr fixtypes in
+ let impls = compute_internalization_env env Recursive fixnames fixtypes fiximps in
+
+ (* Interp bodies with rollback because temp use of notations/implicit *)
+ let sigma, fixdefs =
+ Metasyntax.with_syntax_protection (fun () ->
+ List.iter (Metasyntax.set_notation_for_interpretation env_rec impls) notations;
+ List.fold_left4_map
+ (fun sigma fixctximpenv -> interp_fix_body env_rec sigma (Id.Map.fold Id.Map.add fixctximpenv impls))
+ sigma fixctximpenvs fixctxs fixl fixccls)
+ () in
+
+ (* Instantiate evars and check all are resolved *)
+ let sigma = solve_unif_constraints_with_heuristics env_rec sigma in
+ let sigma, nf = nf_evars_and_universes sigma in
+ let fixdefs = List.map (fun c -> Option.map EConstr.Unsafe.to_constr c) fixdefs in
+ let fixdefs = List.map (Option.map nf) fixdefs in
+ let fixtypes = List.map nf fixtypes in
+ let fixctxs = List.map (fun (_,ctx) -> ctx) fixctxs in
+
+ (* Build the fix declaration block *)
+ (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;
+ 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)
+ end
+
+let interp_fixpoint ~cofix l ntns =
+ let (env,_,pl,evd),fix,info = interp_recursive ~program_mode:false ~cofix l ntns in
+ check_recursive true env evd fix;
+ (fix,pl,Evd.evar_universe_context evd,info)
+
+let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) indexes ntns =
+ if List.exists Option.is_empty fixdefs then
+ (* Some bodies to define by proof *)
+ let thms =
+ List.map3 (fun id t (ctx,imps,_) -> (id,(EConstr.of_constr t,(List.map RelDecl.get_name ctx,imps))))
+ fixnames fixtypes fiximps in
+ let init_tac =
+ 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)
+ evd pl (Some(false,indexes,init_tac)) thms None (Lemmas.mk_hook (fun _ _ -> ()))
+ else begin
+ (* We shortcut the proof process *)
+ let fixdefs = List.map Option.get fixdefs in
+ let fixdecls = prepare_recursive_declaration fixnames fixtypes fixdefs in
+ let env = Global.env() in
+ let indexes = search_guard env indexes fixdecls in
+ let fiximps = List.map (fun (n,r,p) -> r) fiximps in
+ let vars = Univops.universes_of_constr env (mkFix ((indexes,0),fixdecls)) in
+ let fixdecls =
+ List.map_i (fun i _ -> mkFix ((indexes,i),fixdecls)) 0 fixnames in
+ let evd = Evd.from_ctx ctx in
+ let evd = Evd.restrict_universe_context evd vars in
+ let ctx = Evd.check_univ_decl ~poly evd pl in
+ let pl = Evd.universe_binders evd in
+ let fixdecls = List.map Safe_typing.mk_pure_proof fixdecls in
+ ignore (List.map4 (DeclareDef.declare_fix (local, poly, Fixpoint) pl ctx)
+ fixnames fixdecls fixtypes fiximps);
+ (* Declare the recursive definitions *)
+ fixpoint_message (Some indexes) fixnames;
+ end;
+ (* Declare notations *)
+ List.iter (Metasyntax.add_notation_interpretation (Global.env())) ntns
+
+let declare_cofixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ntns =
+ if List.exists Option.is_empty fixdefs then
+ (* Some bodies to define by proof *)
+ let thms =
+ List.map3 (fun id t (ctx,imps,_) -> (id,(EConstr.of_constr t,(List.map RelDecl.get_name ctx,imps))))
+ fixnames fixtypes fiximps in
+ let init_tac =
+ 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 CoFixpoint)
+ evd pl (Some(true,[],init_tac)) thms None (Lemmas.mk_hook (fun _ _ -> ()))
+ else begin
+ (* We shortcut the proof process *)
+ let fixdefs = List.map Option.get fixdefs in
+ let fixdecls = prepare_recursive_declaration fixnames fixtypes fixdefs in
+ let fixdecls = List.map_i (fun i _ -> mkCoFix (i,fixdecls)) 0 fixnames in
+ let env = Global.env () in
+ let vars = Univops.universes_of_constr env (List.hd fixdecls) in
+ let fixdecls = List.map Safe_typing.mk_pure_proof fixdecls in
+ let fiximps = List.map (fun (len,imps,idx) -> imps) fiximps in
+ let evd = Evd.from_ctx ctx in
+ let evd = Evd.restrict_universe_context evd vars in
+ let ctx = Evd.check_univ_decl ~poly evd pl in
+ let pl = Evd.universe_binders evd in
+ ignore (List.map4 (DeclareDef.declare_fix (local, poly, CoFixpoint) pl ctx)
+ fixnames fixdecls fixtypes fiximps);
+ (* Declare the recursive definitions *)
+ cofixpoint_message fixnames
+ end;
+ (* Declare notations *)
+ List.iter (Metasyntax.add_notation_interpretation (Global.env())) ntns
+
+let extract_decreasing_argument limit = function
+ | (na,CStructRec) -> na
+ | (na,_) when not limit -> na
+ | _ -> user_err Pp.(str
+ "Only structural decreasing is supported for a non-Program Fixpoint")
+
+let extract_fixpoint_components limit l =
+ let fixl, ntnl = List.split l in
+ let fixl = List.map (fun (((_,id),pl),ann,bl,typ,def) ->
+ let ann = extract_decreasing_argument limit ann in
+ {fix_name = id; fix_annot = ann; fix_univs = pl;
+ fix_binders = bl; fix_body = def; fix_type = typ}) fixl in
+ fixl, List.flatten ntnl
+
+let extract_cofixpoint_components l =
+ let fixl, ntnl = List.split l in
+ List.map (fun (((_,id),pl),bl,typ,def) ->
+ {fix_name = id; fix_annot = None; fix_univs = pl;
+ fix_binders = bl; fix_body = def; fix_type = typ}) fixl,
+ List.flatten ntnl
+
+let check_safe () =
+ let open Declarations in
+ let flags = Environ.typing_flags (Global.env ()) in
+ flags.check_universes && flags.check_guarded
+
+let do_fixpoint local poly l =
+ let fixl, ntns = extract_fixpoint_components true l in
+ let (_, _, _, info as fix) = interp_fixpoint ~cofix:false fixl ntns in
+ let possible_indexes =
+ List.map compute_possible_guardness_evidences info in
+ declare_fixpoint local poly fix possible_indexes ntns;
+ if not (check_safe ()) then Feedback.feedback Feedback.AddedAxiom else ()
+
+let do_cofixpoint local poly l =
+ let fixl,ntns = extract_cofixpoint_components l in
+ let cofix = interp_fixpoint ~cofix:true fixl ntns in
+ declare_cofixpoint local poly cofix ntns;
+ if not (check_safe ()) then Feedback.feedback Feedback.AddedAxiom else ()
diff --git a/vernac/comFixpoint.mli b/vernac/comFixpoint.mli
new file mode 100644
index 000000000..2c84bd84d
--- /dev/null
+++ b/vernac/comFixpoint.mli
@@ -0,0 +1,93 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2018 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open Constr
+open Decl_kinds
+open Constrexpr
+open Vernacexpr
+
+(** {6 Fixpoints and cofixpoints} *)
+
+(** Entry points for the vernacular commands Fixpoint and CoFixpoint *)
+
+val do_fixpoint :
+ (* When [false], assume guarded. *)
+ locality -> polymorphic -> (fixpoint_expr * decl_notation list) list -> unit
+
+val do_cofixpoint :
+ (* When [false], assume guarded. *)
+ locality -> polymorphic -> (cofixpoint_expr * decl_notation list) list -> unit
+
+(************************************************************************)
+(** Internal API *)
+(************************************************************************)
+
+type structured_fixpoint_expr = {
+ fix_name : Id.t;
+ fix_univs : Vernacexpr.universe_decl_expr option;
+ fix_annot : Id.t Loc.located option;
+ fix_binders : local_binder_expr list;
+ fix_body : constr_expr option;
+ fix_type : constr_expr
+}
+
+(** Typing global fixpoints and cofixpoint_expr *)
+
+(** Exported for Program *)
+val interp_recursive :
+ (* Misc arguments *)
+ program_mode:bool -> cofix:bool ->
+ (* Notations of the fixpoint / should that be folded in the previous argument? *)
+ structured_fixpoint_expr list -> decl_notation list ->
+
+ (* env / signature / univs / evar_map *)
+ (Environ.env * EConstr.named_context * Univdecls.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 *)
+ (EConstr.rel_context * Impargs.manual_explicitation list * int option) list
+
+(** Exported for Funind *)
+
+(** Extracting the semantical components out of the raw syntax of
+ (co)fixpoints declarations *)
+
+val extract_fixpoint_components : bool ->
+ (fixpoint_expr * decl_notation list) list ->
+ structured_fixpoint_expr list * decl_notation list
+
+val extract_cofixpoint_components :
+ (cofixpoint_expr * decl_notation list) list ->
+ structured_fixpoint_expr list * decl_notation list
+
+type recursive_preentry =
+ Id.t list * constr option list * types list
+
+val interp_fixpoint :
+ cofix:bool ->
+ structured_fixpoint_expr list -> decl_notation list ->
+ recursive_preentry * Univdecls.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 *
+ (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 *
+ (Context.Rel.t * Impargs.manual_implicits * int option) list ->
+ decl_notation list -> unit
+
+(** Very private function, do not use *)
+val compute_possible_guardness_evidences :
+ ('a, 'b) Context.Rel.pt * 'c * int option -> int list
diff --git a/vernac/comInductive.ml b/vernac/comInductive.ml
new file mode 100644
index 000000000..41474fc63
--- /dev/null
+++ b/vernac/comInductive.ml
@@ -0,0 +1,455 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Pp
+open CErrors
+open Sorts
+open Util
+open Constr
+open Termops
+open Environ
+open Declare
+open Names
+open Libnames
+open Globnames
+open Nameops
+open Constrexpr
+open Constrexpr_ops
+open Constrintern
+open Nametab
+open Impargs
+open Reductionops
+open Indtypes
+open Pretyping
+open Evarutil
+open Indschemes
+open Misctypes
+open Context.Rel.Declaration
+open Entries
+
+module RelDecl = Context.Rel.Declaration
+
+(* 3b| Mutual inductive definitions *)
+
+let rec complete_conclusion a cs = CAst.map_with_loc (fun ?loc -> function
+ | CProdN (bl,c) -> CProdN (bl,complete_conclusion a cs c)
+ | CLetIn (na,b,t,c) -> CLetIn (na,b,t,complete_conclusion a cs c)
+ | CHole (k, _, _) ->
+ let (has_no_args,name,params) = a in
+ if not has_no_args then
+ user_err ?loc
+ (strbrk"Cannot infer the non constant arguments of the conclusion of "
+ ++ Id.print cs ++ str ".");
+ let args = List.map (fun id -> CAst.make ?loc @@ CRef(Ident(loc,id),None)) params in
+ CAppExpl ((None,Ident(loc,name),None),List.rev args)
+ | c -> c
+ )
+
+let push_types env idl tl =
+ List.fold_left2 (fun env id t -> Environ.push_rel (LocalAssum (Name id,t)) env)
+ env idl tl
+
+type structured_one_inductive_expr = {
+ ind_name : Id.t;
+ ind_univs : Vernacexpr.universe_decl_expr option;
+ ind_arity : constr_expr;
+ ind_lc : (Id.t * constr_expr) list
+}
+
+type structured_inductive_expr =
+ local_binder_expr list * structured_one_inductive_expr list
+
+let minductive_message warn = function
+ | [] -> user_err Pp.(str "No inductive definition.")
+ | [x] -> (Id.print x ++ str " is defined" ++
+ if warn then str " as a non-primitive record" else mt())
+ | l -> hov 0 (prlist_with_sep pr_comma Id.print l ++
+ spc () ++ str "are defined")
+
+let check_all_names_different indl =
+ let ind_names = List.map (fun ind -> ind.ind_name) indl in
+ let cstr_names = List.map_append (fun ind -> List.map fst ind.ind_lc) indl in
+ let l = List.duplicates Id.equal ind_names in
+ let () = match l with
+ | [] -> ()
+ | t :: _ -> raise (InductiveError (SameNamesTypes t))
+ in
+ let l = List.duplicates Id.equal cstr_names in
+ let () = match l with
+ | [] -> ()
+ | c :: _ -> raise (InductiveError (SameNamesConstructors (List.hd l)))
+ in
+ let l = List.intersect Id.equal ind_names cstr_names in
+ match l with
+ | [] -> ()
+ | _ -> raise (InductiveError (SameNamesOverlap l))
+
+let mk_mltype_data sigma env assums arity indname =
+ let is_ml_type = is_sort env sigma (EConstr.of_constr arity) in
+ (is_ml_type,indname,assums)
+
+let prepare_param = function
+ | LocalAssum (na,t) -> Name.get_id na, LocalAssumEntry t
+ | LocalDef (na,b,_) -> Name.get_id na, LocalDefEntry b
+
+(** Make the arity conclusion flexible to avoid generating an upper bound universe now,
+ only if the universe does not appear anywhere else.
+ This is really a hack to stay compatible with the semantics of template polymorphic
+ inductives which are recognized when a "Type" appears at the end of the conlusion in
+ the source syntax. *)
+
+let rec check_anonymous_type ind =
+ let open Glob_term in
+ match DAst.get ind with
+ | GSort (GType []) -> true
+ | GProd ( _, _, _, e)
+ | GLetIn (_, _, _, e)
+ | GLambda (_, _, _, e)
+ | GApp (e, _)
+ | GCast (e, _) -> check_anonymous_type e
+ | _ -> false
+
+let make_conclusion_flexible sigma ty poly =
+ if poly && Term.isArity ty then
+ let _, concl = Term.destArity ty in
+ match concl with
+ | Type u ->
+ (match Univ.universe_level u with
+ | Some u ->
+ Evd.make_flexible_variable sigma ~algebraic:true u
+ | None -> sigma)
+ | _ -> sigma
+ else sigma
+
+let is_impredicative env u =
+ u = Prop Null || (is_impredicative_set env && u = Prop Pos)
+
+let interp_ind_arity env sigma ind =
+ let c = intern_gen IsType env ind.ind_arity in
+ let impls = Implicit_quantifiers.implicits_of_glob_constr ~with_products:true c in
+ let sigma,t = understand_tcc env sigma ~expected_type:IsType c in
+ let pseudo_poly = check_anonymous_type c in
+ let () = if not (Reductionops.is_arity env sigma t) then
+ user_err ?loc:(constr_loc ind.ind_arity) (str "Not an arity")
+ in
+ let t = EConstr.Unsafe.to_constr t in
+ sigma, (t, pseudo_poly, impls)
+
+let interp_cstrs env sigma impls mldata arity ind =
+ let cnames,ctyps = List.split ind.ind_lc in
+ (* Complete conclusions of constructor types if given in ML-style syntax *)
+ let ctyps' = List.map2 (complete_conclusion mldata) cnames ctyps in
+ (* Interpret the constructor types *)
+ let sigma, (ctyps'', cimpls) =
+ on_snd List.split @@
+ List.fold_left_map (fun sigma l ->
+ on_snd (on_fst EConstr.Unsafe.to_constr) @@
+ interp_type_evars_impls env sigma ~impls l) sigma ctyps' in
+ sigma, (cnames, ctyps'', cimpls)
+
+let sign_level env evd sign =
+ fst (List.fold_right
+ (fun d (lev,env) ->
+ match d with
+ | LocalDef _ -> lev, push_rel d env
+ | LocalAssum _ ->
+ let s = destSort (Reduction.whd_all env
+ (EConstr.Unsafe.to_constr (nf_evar 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))
+ sign (Univ.type0m_univ,env))
+
+let sup_list min = List.fold_left Univ.sup min
+
+let extract_level env evd min tys =
+ let sorts = List.map (fun ty ->
+ let ctx, concl = Reduction.dest_prod_assum env ty in
+ sign_level env evd (LocalAssum (Anonymous, concl) :: ctx)) tys
+ in sup_list min sorts
+
+let is_flexible_sort evd u =
+ match Univ.Universe.level u with
+ | Some l -> Evd.is_flexible_level evd l
+ | None -> false
+
+let inductive_levels env 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)) ->
+ if a = Prop Null then None
+ else Some (univ_of_sort a)) destarities
+ in
+ let cstrs_levels, min_levels, sizes =
+ CList.split3
+ (List.map2 (fun (_,tys,_) (arity,(ctx,du)) ->
+ let len = List.length tys in
+ let minlev = Sorts.univ_of_sort du in
+ let minlev =
+ if len > 1 && not (is_impredicative env du) then
+ Univ.sup minlev Univ.type0_univ
+ else minlev
+ in
+ let minlev =
+ (** Indices contribute. *)
+ if Indtypes.is_indices_matter () && List.length ctx > 0 then (
+ let ilev = sign_level env evd ctx in
+ Univ.sup ilev minlev)
+ else minlev
+ in
+ let clev = extract_level env evd minlev tys in
+ (clev, minlev, len)) inds destarities)
+ in
+ (* Take the transitive closure of the system of constructors *)
+ (* level constraints and remove the recursive dependencies *)
+ let levels' = Universes.solve_constraints_system (Array.of_list levels)
+ (Array.of_list cstrs_levels) (Array.of_list min_levels)
+ in
+ let evd, arities =
+ CList.fold_left3 (fun (evd, arities) cu (arity,(ctx,du)) len ->
+ if is_impredicative env du then
+ (** Any product is allowed here. *)
+ evd, arity :: arities
+ else (** If in a predicative sort, or asked to infer the type,
+ we take the max of:
+ - indices (if in indices-matter mode)
+ - constructors
+ - Type(1) if there is more than 1 constructor
+ *)
+ (** Constructors contribute. *)
+ let evd =
+ if Sorts.is_set du then
+ if not (Evd.check_leq evd cu Univ.type0_univ) then
+ raise (Indtypes.InductiveError Indtypes.LargeNonPropInductiveNotInType)
+ else evd
+ else evd
+ (* Evd.set_leq_sort env evd (Type cu) du *)
+ in
+ let evd =
+ if len >= 2 && Univ.is_type0m_univ cu then
+ (** "Polymorphic" type constraint and more than one constructor,
+ should not land in Prop. Add constraint only if it would
+ land in Prop directly (no informative arguments as well). *)
+ Evd.set_leq_sort env evd (Prop Pos) du
+ else evd
+ in
+ let duu = Sorts.univ_of_sort du in
+ let evd =
+ if not (Univ.is_small_univ duu) && Univ.Universe.equal cu duu then
+ if is_flexible_sort evd duu && not (Evd.check_leq evd Univ.type0_univ duu) then
+ Evd.set_eq_sort env evd (Prop Null) du
+ else evd
+ else Evd.set_eq_sort env evd (Type cu) du
+ in
+ (evd, arity :: arities))
+ (evd,[]) (Array.to_list levels') destarities sizes
+ in evd, List.rev arities
+
+let check_named (loc, na) = match na with
+| Name _ -> ()
+| Anonymous ->
+ let msg = str "Parameters must be named." in
+ user_err ?loc msg
+
+
+let check_param = function
+| CLocalDef (na, _, _) -> check_named na
+| CLocalAssum (nas, Default _, _) -> List.iter check_named nas
+| CLocalAssum (nas, Generalized _, _) -> ()
+| CLocalPattern (loc,_) ->
+ 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, (impls, ((env_params, ctx_params), userimpls)) =
+ interp_context_evars env0 sigma paramsl
+ in
+ let ctx_params = List.map (fun d -> map_rel_decl EConstr.Unsafe.to_constr d) ctx_params in
+ let indnames = List.map (fun ind -> ind.ind_name) indl in
+
+ (* Names of parameters as arguments of the inductive type (defs removed) *)
+ let assums = List.filter is_local_assum ctx_params in
+ let params = List.map (RelDecl.get_name %> Name.get_id) assums in
+
+ (* Interpret the arities *)
+ let sigma, arities = List.fold_left_map (fun sigma -> interp_ind_arity env_params sigma) sigma indl in
+
+ let fullarities = List.map (fun (c, _, _) -> Term.it_mkProd_or_LetIn c ctx_params) arities in
+ let env_ar = push_types env0 indnames fullarities in
+ let env_ar_params = push_rel_context ctx_params env_ar in
+
+ (* Compute interpretation metadatas *)
+ let indimpls = List.map (fun (_, _, impls) -> userimpls @
+ lift_implicits (Context.Rel.nhyps ctx_params) impls) arities in
+ let arities = List.map pi1 arities and aritypoly = List.map pi2 arities in
+ let impls = compute_internalization_env env0 ~impls (Inductive (params,true)) indnames fullarities indimpls in
+ let ntn_impls = compute_internalization_env env0 (Inductive (params,true)) indnames fullarities indimpls in
+ let mldatas = List.map2 (mk_mltype_data sigma env_params params) arities indnames in
+
+ let sigma, constructors =
+ Metasyntax.with_syntax_protection (fun () ->
+ (* Temporary declaration of notations and scopes *)
+ List.iter (Metasyntax.set_notation_for_interpretation env_params ntn_impls) notations;
+ (* Interpret the constructor types *)
+ List.fold_left3_map (fun sigma -> interp_cstrs env_ar_params sigma impls) sigma mldatas arities indl)
+ () in
+
+ (* Try further to solve evars, and instantiate them *)
+ let sigma = solve_remaining_evars all_and_fail_flags env_params sigma Evd.empty in
+ (* Compute renewed arities *)
+ 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 = 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 nf x = nf' (nf x) in
+ let arities = List.map nf' arities in
+ let constructors = List.map (fun (idl,cl,impsl) -> (idl,List.map nf' cl,impsl)) constructors in
+ let ctx_params = Context.Rel.map nf ctx_params in
+ let 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 (_,ctyps,_) ->
+ List.iter (fun c -> check_evars env_ar_params Evd.empty sigma (EConstr.of_constr c)) ctyps)
+ constructors;
+
+ (* Build the inductive entries *)
+ let entries = List.map4 (fun ind arity template (cnames,ctypes,cimpls) -> {
+ mind_entry_typename = ind.ind_name;
+ mind_entry_arity = arity;
+ mind_entry_template = template;
+ mind_entry_consnames = cnames;
+ mind_entry_lc = ctypes
+ }) indl arities aritypoly constructors in
+ let impls =
+ let len = Context.Rel.nhyps ctx_params in
+ List.map2 (fun indimpls (_,_,cimpls) ->
+ indimpls, List.map (fun impls ->
+ userimpls @ (lift_implicits len impls)) cimpls) indimpls constructors
+ in
+ let univs =
+ match uctx with
+ | Polymorphic_const_entry uctx ->
+ if cum then
+ Cumulative_ind_entry (Univ.CumulativityInfo.from_universe_context uctx)
+ else Polymorphic_ind_entry uctx
+ | Monomorphic_const_entry uctx ->
+ Monomorphic_ind_entry uctx
+ in
+ (* Build the mutual inductive entry *)
+ let mind_ent =
+ { mind_entry_params = List.map prepare_param ctx_params;
+ mind_entry_record = None;
+ mind_entry_finite = finite;
+ mind_entry_inds = entries;
+ mind_entry_private = if prv then Some false else None;
+ mind_entry_universes = univs;
+ }
+ in
+ (if poly && cum then
+ InferCumulativity.infer_inductive env_ar mind_ent
+ else mind_ent), Evd.universe_binders sigma, impls
+
+(* Very syntactical equality *)
+let eq_local_binders bl1 bl2 =
+ List.equal local_binder_eq bl1 bl2
+
+let extract_coercions indl =
+ let mkqid (_,((_,id),_)) = qualid_of_ident id in
+ let extract lc = List.filter (fun (iscoe,_) -> iscoe) lc in
+ List.map mkqid (List.flatten(List.map (fun (_,_,_,lc) -> extract lc) indl))
+
+let extract_params indl =
+ let paramsl = List.map (fun (_,params,_,_) -> params) indl in
+ match paramsl with
+ | [] -> anomaly (Pp.str "empty list of inductive types.")
+ | params::paramsl ->
+ if not (List.for_all (eq_local_binders params) paramsl) then user_err Pp.(str
+ "Parameters should be syntactically the same for each inductive type.");
+ params
+
+let extract_inductive indl =
+ List.map (fun (((_,indname),pl),_,ar,lc) -> {
+ ind_name = indname; ind_univs = pl;
+ ind_arity = Option.cata (fun x -> x) (CAst.make @@ CSort (GType [])) ar;
+ ind_lc = List.map (fun (_,((_,id),t)) -> (id,t)) lc
+ }) indl
+
+let extract_mutual_inductive_declaration_components indl =
+ let indl,ntnl = List.split indl in
+ let params = extract_params indl in
+ let coes = extract_coercions indl in
+ let indl = extract_inductive indl in
+ (params,indl), coes, List.flatten ntnl
+
+let is_recursive mie =
+ let rec is_recursive_constructor lift typ =
+ match Constr.kind typ with
+ | Prod (_,arg,rest) ->
+ not (EConstr.Vars.noccurn Evd.empty (** FIXME *) lift (EConstr.of_constr arg)) ||
+ is_recursive_constructor (lift+1) rest
+ | LetIn (na,b,t,rest) -> is_recursive_constructor (lift+1) rest
+ | _ -> false
+ in
+ match mie.mind_entry_inds with
+ | [ind] ->
+ let nparams = List.length mie.mind_entry_params in
+ List.exists (fun t -> is_recursive_constructor (nparams+1) t) ind.mind_entry_lc
+ | _ -> false
+
+let declare_mutual_inductive_with_eliminations mie pl impls =
+ (* spiwack: raises an error if the structure is supposed to be non-recursive,
+ but isn't *)
+ begin match mie.mind_entry_finite with
+ | Declarations.BiFinite when is_recursive mie ->
+ if Option.has_some mie.mind_entry_record then
+ user_err Pp.(str "Records declared with the keywords Record or Structure cannot be recursive. You can, however, define recursive records using the Inductive or CoInductive command.")
+ else
+ user_err Pp.(str ("Types declared with the keyword Variant cannot be recursive. Recursive types are defined with the Inductive and CoInductive command."))
+ | _ -> ()
+ end;
+ let names = List.map (fun e -> e.mind_entry_typename) mie.mind_entry_inds in
+ let (_, kn), prim = declare_mind mie in
+ let mind = Global.mind_of_delta_kn kn in
+ List.iteri (fun i (indimpls, constrimpls) ->
+ let ind = (mind,i) in
+ let gr = IndRef ind in
+ maybe_declare_manual_implicits false gr indimpls;
+ Declare.declare_univ_binders gr pl;
+ List.iteri
+ (fun j impls ->
+ maybe_declare_manual_implicits false
+ (ConstructRef (ind, succ j)) impls)
+ constrimpls)
+ impls;
+ let warn_prim = match mie.mind_entry_record with Some (Some _) -> not prim | _ -> false in
+ Flags.if_verbose Feedback.msg_info (minductive_message warn_prim names);
+ if mie.mind_entry_private == None
+ then declare_default_schemes mind;
+ mind
+
+type one_inductive_impls =
+ Impargs.manual_explicitation list (* for inds *)*
+ Impargs.manual_explicitation list list (* for constrs *)
+
+let do_mutual_inductive indl cum poly prv finite =
+ let indl,coes,ntns = extract_mutual_inductive_declaration_components indl in
+ (* Interpret the types *)
+ let mie,pl,impls = interp_mutual_inductive indl ntns cum poly prv finite in
+ (* Declare the mutual inductive block with its associated schemes *)
+ ignore (declare_mutual_inductive_with_eliminations mie pl impls);
+ (* Declare the possible notations of inductive types *)
+ List.iter (Metasyntax.add_notation_interpretation (Global.env ())) ntns;
+ (* Declare the coercions *)
+ List.iter (fun qid -> Class.try_add_new_coercion (locate qid) ~local:false poly) coes;
+ (* If positivity is assumed declares itself as unsafe. *)
+ if Environ.deactivated_guard (Global.env ()) then Feedback.feedback Feedback.AddedAxiom else ()
diff --git a/vernac/comInductive.mli b/vernac/comInductive.mli
new file mode 100644
index 000000000..82ea131e1
--- /dev/null
+++ b/vernac/comInductive.mli
@@ -0,0 +1,65 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open Entries
+open Libnames
+open Vernacexpr
+open Constrexpr
+open Decl_kinds
+
+(** {6 Inductive and coinductive types} *)
+
+(** Entry points for the vernacular commands Inductive and CoInductive *)
+
+val do_mutual_inductive :
+ (one_inductive_expr * decl_notation list) list -> cumulative_inductive_flag ->
+ polymorphic -> private_flag -> Declarations.recursivity_kind -> unit
+
+(************************************************************************)
+(** Internal API *)
+(************************************************************************)
+
+(** Exported for Record and Funind *)
+
+(** Registering a mutual inductive definition together with its
+ associated schemes *)
+
+type one_inductive_impls =
+ Impargs.manual_implicits (** for inds *)*
+ Impargs.manual_implicits list (** for constrs *)
+
+val declare_mutual_inductive_with_eliminations :
+ mutual_inductive_entry -> Universes.universe_binders -> one_inductive_impls list ->
+ MutInd.t
+
+(** Exported for Funind *)
+
+(** Extracting the semantical components out of the raw syntax of mutual
+ inductive declarations *)
+
+type structured_one_inductive_expr = {
+ ind_name : Id.t;
+ ind_univs : Vernacexpr.universe_decl_expr option;
+ ind_arity : constr_expr;
+ ind_lc : (Id.t * constr_expr) list
+}
+
+type structured_inductive_expr =
+ local_binder_expr list * structured_one_inductive_expr list
+
+val extract_mutual_inductive_declaration_components :
+ (one_inductive_expr * decl_notation list) list ->
+ structured_inductive_expr * (*coercions:*) qualid list * decl_notation list
+
+(** Typing mutual inductive definitions *)
+
+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
diff --git a/vernac/comProgramFixpoint.ml b/vernac/comProgramFixpoint.ml
new file mode 100644
index 000000000..a9a91e304
--- /dev/null
+++ b/vernac/comProgramFixpoint.ml
@@ -0,0 +1,342 @@
+open Pp
+open CErrors
+open Util
+open Constr
+open Entries
+open Vars
+open Declare
+open Names
+open Libnames
+open Globnames
+open Nameops
+open Constrexpr
+open Constrexpr_ops
+open Constrintern
+open Decl_kinds
+open Evarutil
+open Context.Rel.Declaration
+open ComFixpoint
+
+module RelDecl = Context.Rel.Declaration
+
+(* Wellfounded definition *)
+
+open Coqlib
+
+let contrib_name = "Program"
+let subtac_dir = [contrib_name]
+let fixsub_module = subtac_dir @ ["Wf"]
+(* let tactics_module = subtac_dir @ ["Tactics"] *)
+
+let init_reference dir s () = Coqlib.coq_reference "Command" dir s
+let init_constant dir s sigma =
+ Evarutil.new_global sigma (Coqlib.coq_reference "Command" dir s)
+
+let make_ref l s = init_reference l s
+(* let fix_proto = init_constant tactics_module "fix_proto" *)
+let fix_sub_ref = make_ref fixsub_module "Fix_sub"
+let measure_on_R_ref = make_ref fixsub_module "MR"
+let well_founded = init_constant ["Init"; "Wf"] "well_founded"
+let mkSubset sigma name typ prop =
+ let open EConstr in
+ let sigma, app_h = Evarutil.new_global sigma (delayed_force build_sigma).typ in
+ sigma, mkApp (app_h, [| typ; mkLambda (name, typ, prop) |])
+
+let sigT = Lazy.from_fun build_sigma_type
+
+let make_qref s = Qualid (Loc.tag @@ qualid_of_string s)
+let lt_ref = make_qref "Init.Peano.lt"
+
+let rec telescope sigma l =
+ let open EConstr in
+ let open Vars in
+ match l with
+ | [] -> assert false
+ | [LocalAssum (n, t)] ->
+ sigma, t, [LocalDef (n, mkRel 1, t)], mkRel 1
+ | LocalAssum (n, t) :: tl ->
+ let sigma, ty, tys, (k, constr) =
+ List.fold_left
+ (fun (sigma, ty, tys, (k, constr)) decl ->
+ let t = RelDecl.get_type decl in
+ let pred = mkLambda (RelDecl.get_name decl, t, ty) in
+ let sigma, ty = Evarutil.new_global sigma (Lazy.force sigT).typ in
+ let sigma, intro = Evarutil.new_global sigma (Lazy.force sigT).intro in
+ let sigty = mkApp (ty, [|t; pred|]) in
+ let intro = mkApp (intro, [|lift k t; lift k pred; mkRel k; constr|]) in
+ (sigma, sigty, pred :: tys, (succ k, intro)))
+ (sigma, t, [], (2, mkRel 1)) tl
+ in
+ let sigma, last, subst = List.fold_right2
+ (fun pred decl (sigma, prev, subst) ->
+ let t = RelDecl.get_type decl in
+ let sigma, p1 = Evarutil.new_global sigma (Lazy.force sigT).proj1 in
+ let sigma, p2 = Evarutil.new_global sigma (Lazy.force sigT).proj2 in
+ let proj1 = applist (p1, [t; pred; prev]) in
+ let proj2 = applist (p2, [t; pred; prev]) in
+ (sigma, lift 1 proj2, LocalDef (get_name decl, proj1, t) :: subst))
+ (List.rev tys) tl (sigma, mkRel 1, [])
+ in sigma, ty, (LocalDef (n, last, t) :: subst), constr
+
+ | LocalDef (n, b, t) :: tl ->
+ let sigma, ty, subst, term = telescope sigma tl in
+ sigma, ty, (LocalDef (n, b, t) :: subst), lift 1 term
+
+let nf_evar_context sigma ctx =
+ List.map (map_constr (fun c -> Evarutil.nf_evar sigma c)) ctx
+
+let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
+ let open EConstr in
+ let open Vars in
+ 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, (_, ((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
+ let sigma, top_arity = interp_type_evars top_env sigma arityc in
+ let full_arity = it_mkProd_or_LetIn top_arity binders_rel in
+ let sigma, argtyp, letbinders, make = telescope sigma binders_rel in
+ let argname = Id.of_string "recarg" in
+ let arg = LocalAssum (Name argname, argtyp) in
+ let binders = letbinders @ [arg] in
+ let binders_env = push_rel_context binders_rel env in
+ let sigma, (rel, _) = interp_constr_evars_impls env sigma r in
+ let relty = Typing.unsafe_type_of env sigma rel in
+ let relargty =
+ let error () =
+ user_err ?loc:(constr_loc r)
+ ~hdr:"Command.build_wellfounded"
+ (Printer.pr_econstr_env env sigma rel ++ str " is not an homogeneous binary relation.")
+ in
+ try
+ let ctx, ar = Reductionops.splay_prod_n env sigma 2 relty in
+ match ctx, EConstr.kind sigma ar with
+ | [LocalAssum (_,t); LocalAssum (_,u)], Sort s
+ when Sorts.is_prop (ESorts.kind sigma s) && Reductionops.is_conv env sigma t u -> t
+ | _, _ -> error ()
+ with e when CErrors.noncritical e -> error ()
+ in
+ let sigma, measure = interp_casted_constr_evars binders_env sigma measure relargty in
+ let sigma, wf_rel, wf_rel_fun, measure_fn =
+ let measure_body, measure =
+ it_mkLambda_or_LetIn measure letbinders,
+ it_mkLambda_or_LetIn measure binders
+ in
+ let sigma, comb = Evarutil.new_global sigma (delayed_force measure_on_R_ref) in
+ let wf_rel = mkApp (comb, [| argtyp; relargty; rel; measure |]) in
+ let wf_rel_fun x y =
+ mkApp (rel, [| subst1 x measure_body;
+ subst1 y measure_body |])
+ in sigma, wf_rel, wf_rel_fun, measure
+ in
+ let sigma, wf_term = well_founded sigma in
+ let wf_proof = mkApp (wf_term, [| argtyp ; wf_rel |]) in
+ let argid' = Id.of_string (Id.to_string argname ^ "'") in
+ let wfarg sigma len =
+ let sigma, ss_term = mkSubset sigma (Name argid') argtyp (wf_rel_fun (mkRel 1) (mkRel (len + 1))) in
+ sigma, LocalAssum (Name argid', ss_term)
+ in
+ let sigma, intern_bl =
+ let sigma, wfa = wfarg sigma 1 in
+ sigma, wfa :: [arg]
+ in
+ let _intern_env = push_rel_context intern_bl env in
+ let sigma, proj = Evarutil.new_global sigma (delayed_force build_sigma).Coqlib.proj1 in
+ let wfargpred = mkLambda (Name argid', argtyp, wf_rel_fun (mkRel 1) (mkRel 3)) in
+ let projection = (* in wfarg :: arg :: before *)
+ mkApp (proj, [| argtyp ; wfargpred ; mkRel 1 |])
+ in
+ let top_arity_let = it_mkLambda_or_LetIn top_arity letbinders in
+ let intern_arity = substl [projection] top_arity_let in
+ (* substitute the projection of wfarg for something,
+ now intern_arity is in wfarg :: arg *)
+ let sigma, wfa = wfarg sigma 1 in
+ let intern_fun_arity_prod = it_mkProd_or_LetIn intern_arity [wfa] in
+ let intern_fun_binder = LocalAssum (Name (add_suffix recname "'"), intern_fun_arity_prod) in
+ let sigma, curry_fun =
+ let wfpred = mkLambda (Name argid', argtyp, wf_rel_fun (mkRel 1) (mkRel (2 * len + 4))) in
+ let sigma, intro = Evarutil.new_global sigma (delayed_force build_sigma).Coqlib.intro in
+ let arg = mkApp (intro, [| argtyp; wfpred; lift 1 make; mkRel 1 |]) in
+ let app = mkApp (mkRel (2 * len + 2 (* recproof + orig binders + current binders *)), [| arg |]) in
+ let rcurry = mkApp (rel, [| measure; lift len measure |]) in
+ let lam = LocalAssum (Name (Id.of_string "recproof"), rcurry) in
+ let body = it_mkLambda_or_LetIn app (lam :: binders_rel) in
+ let ty = it_mkProd_or_LetIn (lift 1 top_arity) (lam :: binders_rel) in
+ sigma, LocalDef (Name recname, body, ty)
+ in
+ let fun_bl = intern_fun_binder :: [arg] in
+ let lift_lets = lift_rel_context 1 letbinders in
+ let sigma, intern_body =
+ let ctx = LocalAssum (Name recname, get_type curry_fun) :: binders_rel in
+ let (r, l, impls, scopes) =
+ Constrintern.compute_internalization_data env
+ Constrintern.Recursive (EConstr.Unsafe.to_constr full_arity) impls
+ in
+ let newimpls = Id.Map.singleton recname
+ (r, l, impls @ [(Some (Id.of_string "recproof", Impargs.Manual, (true, false)))],
+ scopes @ [None]) in
+ interp_casted_constr_evars (push_rel_context ctx env) sigma
+ ~impls:newimpls body (lift 1 top_arity)
+ in
+ let intern_body_lam = it_mkLambda_or_LetIn intern_body (curry_fun :: lift_lets @ fun_bl) in
+ let prop = mkLambda (Name argname, argtyp, top_arity_let) in
+ (* XXX: Previous code did parallel evdref update, so possible old
+ weak ordering semantics may bite here. *)
+ let sigma, def =
+ let sigma, h_a_term = Evarutil.new_global sigma (delayed_force fix_sub_ref) in
+ let sigma, h_e_term = Evarutil.new_evar env sigma
+ ~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 = Evarutil.nf_evar_map sigma in
+ let def = mkApp (def, [|intern_body_lam|]) in
+ let binders_rel = nf_evar_context sigma binders_rel in
+ let binders = nf_evar_context sigma binders in
+ let top_arity = Evarutil.nf_evar sigma top_arity in
+ let hook, recname, typ =
+ if List.length binders_rel > 1 then
+ let name = add_suffix recname "_func" in
+ (* XXX: Mutating the evar_map in the hook! *)
+ (* XXX: Likely the sigma is out of date when the hook is called .... *)
+ let hook sigma l gr _ =
+ let sigma, h_body = Evarutil.new_global sigma gr in
+ let body = it_mkLambda_or_LetIn (mkApp (h_body, [|make|])) binders_rel in
+ let ty = it_mkProd_or_LetIn top_arity binders_rel in
+ let ty = EConstr.Unsafe.to_constr ty in
+ let univs = Evd.check_univ_decl ~poly sigma decl in
+ (*FIXME poly? *)
+ let ce = definition_entry ~types:ty ~univs (EConstr.to_constr sigma body) in
+ (** 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
+ if Impargs.is_implicit_args () || not (List.is_empty impls) then
+ Impargs.declare_manual_implicits false gr [impls]
+ in
+ let typ = it_mkProd_or_LetIn top_arity binders in
+ hook, name, typ
+ else
+ let typ = it_mkProd_or_LetIn top_arity binders_rel in
+ let hook sigma l gr _ =
+ if Impargs.is_implicit_args () || not (List.is_empty impls) then
+ Impargs.declare_manual_implicits false gr [impls]
+ in hook, recname, typ
+ in
+ (* XXX: Capturing sigma here... bad bad *)
+ let hook = Lemmas.mk_hook (hook sigma) in
+ let fullcoqc = EConstr.to_constr sigma def in
+ let fullctyp = EConstr.to_constr sigma typ in
+ Obligations.check_evars env sigma;
+ let evars, _, evars_def, evars_typ =
+ Obligations.eterm_obligations env recname sigma 0 fullcoqc fullctyp
+ in
+ let ctx = Evd.evar_universe_context sigma in
+ ignore(Obligations.add_definition recname ~term:evars_def ~univdecl:decl
+ evars_typ ctx evars ~hook)
+
+let out_def = function
+ | Some def -> def
+ | None -> user_err Pp.(str "Program Fixpoint needs defined bodies.")
+
+let collect_evars_of_term evd c ty =
+ let evars = Evar.Set.union (Evd.evars_of_term c) (Evd.evars_of_term ty) in
+ Evar.Set.fold (fun ev acc -> Evd.add acc ev (Evd.find_undefined evd ev))
+ evars (Evd.from_ctx (Evd.evar_universe_context evd))
+
+let do_program_recursive local poly fixkind fixl ntns =
+ let cofix = fixkind = Obligations.IsCoFixpoint in
+ let (env, rec_sign, pl, evd), fix, info =
+ interp_recursive ~cofix ~program_mode:true fixl ntns
+ in
+ (* Program-specific code *)
+ (* Get the interesting evars, those that were not instanciated *)
+ let evd = Typeclasses.resolve_typeclasses ~filter:Typeclasses.no_goals ~fail:true env evd in
+ (* Solve remaining evars *)
+ let evd = nf_evar_map_undefined evd in
+ let collect_evars id def typ imps =
+ (* Generalize by the recursive prototypes *)
+ let def =
+ EConstr.to_constr 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)
+ in
+ let evm = collect_evars_of_term evd def typ in
+ let evars, _, def, typ =
+ Obligations.eterm_obligations env id evm
+ (List.length rec_sign) def typ
+ in (id, def, typ, imps, evars)
+ in
+ let (fixnames,fixdefs,fixtypes) = fix in
+ let fiximps = List.map pi2 info in
+ let fixdefs = List.map out_def fixdefs in
+ let defs = List.map4 collect_evars fixnames fixdefs fixtypes fiximps in
+ let () = if not cofix then begin
+ let possible_indexes = List.map ComFixpoint.compute_possible_guardness_evidences info in
+ let fixdecls = Array.of_list (List.map (fun x -> Name x) fixnames),
+ Array.of_list fixtypes,
+ Array.of_list (List.map (subst_vars (List.rev fixnames)) fixdefs)
+ in
+ let indexes =
+ Pretyping.search_guard (Global.env ()) possible_indexes fixdecls in
+ List.iteri (fun i _ ->
+ Inductive.check_fix env
+ ((indexes,i),fixdecls))
+ fixl
+ end in
+ let ctx = Evd.evar_universe_context evd in
+ let kind = match fixkind with
+ | Obligations.IsFixpoint _ -> (local, poly, Fixpoint)
+ | Obligations.IsCoFixpoint -> (local, poly, CoFixpoint)
+ in
+ Obligations.add_mutual_definitions defs ~kind ~univdecl:pl ctx ntns fixkind
+
+let do_program_fixpoint local poly l =
+ let g = List.map (fun ((_,wf,_,_,_),_) -> wf) l in
+ match g, l with
+ | [(n, CWfRec r)], [((((_,id),pl),_,bl,typ,def),ntn)] ->
+ let recarg =
+ match n with
+ | Some n -> mkIdentC (snd n)
+ | None ->
+ user_err ~hdr:"do_program_fixpoint"
+ (str "Recursive argument required for well-founded fixpoints")
+ in build_wellfounded (id, pl, n, bl, typ, out_def def) poly r recarg ntn
+
+ | [(n, CMeasureRec (m, r))], [((((_,id),pl),_,bl,typ,def),ntn)] ->
+ build_wellfounded (id, pl, n, bl, typ, out_def def) poly
+ (Option.default (CAst.make @@ CRef (lt_ref,None)) r) m ntn
+
+ | _, _ when List.for_all (fun (n, ro) -> ro == CStructRec) g ->
+ let fixl,ntns = extract_fixpoint_components true l in
+ let fixkind = Obligations.IsFixpoint g in
+ do_program_recursive local poly fixkind fixl ntns
+
+ | _, _ ->
+ user_err ~hdr:"do_program_fixpoint"
+ (str "Well-founded fixpoints not allowed in mutually recursive blocks")
+
+let extract_cofixpoint_components l =
+ let fixl, ntnl = List.split l in
+ List.map (fun (((_,id),pl),bl,typ,def) ->
+ {fix_name = id; fix_annot = None; fix_univs = pl;
+ fix_binders = bl; fix_body = def; fix_type = typ}) fixl,
+ List.flatten ntnl
+
+let check_safe () =
+ let open Declarations in
+ let flags = Environ.typing_flags (Global.env ()) in
+ flags.check_universes && flags.check_guarded
+
+let do_fixpoint local poly l =
+ do_program_fixpoint local poly l;
+ if not (check_safe ()) then Feedback.feedback Feedback.AddedAxiom else ()
+
+let do_cofixpoint local poly l =
+ let fixl,ntns = extract_cofixpoint_components l in
+ do_program_recursive local poly Obligations.IsCoFixpoint fixl ntns;
+ if not (check_safe ()) then Feedback.feedback Feedback.AddedAxiom else ()
diff --git a/vernac/comProgramFixpoint.mli b/vernac/comProgramFixpoint.mli
new file mode 100644
index 000000000..943cb8efe
--- /dev/null
+++ b/vernac/comProgramFixpoint.mli
@@ -0,0 +1,12 @@
+open Decl_kinds
+open Vernacexpr
+
+(** Special Fixpoint handling when command is activated. *)
+
+val do_fixpoint :
+ (* When [false], assume guarded. *)
+ locality -> polymorphic -> (fixpoint_expr * decl_notation list) list -> unit
+
+val do_cofixpoint :
+ (* When [false], assume guarded. *)
+ locality -> polymorphic -> (cofixpoint_expr * decl_notation list) list -> unit
diff --git a/vernac/command.ml b/vernac/command.ml
deleted file mode 100644
index 373e5a1be..000000000
--- a/vernac/command.ml
+++ /dev/null
@@ -1,1346 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-open Pp
-open CErrors
-open Sorts
-open Util
-open Constr
-open Vars
-open Termops
-open Environ
-open Redexpr
-open Declare
-open Names
-open Libnames
-open Globnames
-open Nameops
-open Constrexpr
-open Constrexpr_ops
-open Constrintern
-open Nametab
-open Impargs
-open Reductionops
-open Indtypes
-open Decl_kinds
-open Pretyping
-open Evarutil
-open Evarconv
-open Indschemes
-open Misctypes
-open Vernacexpr
-open Context.Rel.Declaration
-open Entries
-
-module RelDecl = Context.Rel.Declaration
-
-let do_universe poly l = Declare.do_universe poly l
-let do_constraint poly l = Declare.do_constraint poly l
-
-let rec under_binders env sigma f n c =
- if Int.equal n 0 then f env sigma (EConstr.of_constr c) else
- match Constr.kind c with
- | Lambda (x,t,c) ->
- mkLambda (x,t,under_binders (push_rel (LocalAssum (x,t)) env) sigma f (n-1) c)
- | LetIn (x,b,t,c) ->
- mkLetIn (x,b,t,under_binders (push_rel (LocalDef (x,b,t)) env) sigma f (n-1) c)
- | _ -> assert false
-
-let rec complete_conclusion a cs = CAst.map_with_loc (fun ?loc -> function
- | CProdN (bl,c) -> CProdN (bl,complete_conclusion a cs c)
- | CLetIn (na,b,t,c) -> CLetIn (na,b,t,complete_conclusion a cs c)
- | CHole (k, _, _) ->
- let (has_no_args,name,params) = a in
- if not has_no_args then
- user_err ?loc
- (strbrk"Cannot infer the non constant arguments of the conclusion of "
- ++ Id.print cs ++ str ".");
- let args = List.map (fun id -> CAst.make ?loc @@ CRef(Ident(loc,id),None)) params in
- CAppExpl ((None,Ident(loc,name),None),List.rev args)
- | c -> c
- )
-
-(* Commands of the interface *)
-
-(* 1| Constant definitions *)
-
-let red_constant_entry n ce sigma = function
- | None -> ce
- | Some red ->
- let proof_out = ce.const_entry_body in
- let env = Global.env () in
- let (redfun, _) = reduction_of_red_expr env red in
- let redfun env sigma c =
- let (_, c) = redfun env sigma c in
- EConstr.Unsafe.to_constr c
- in
- { ce with const_entry_body = Future.chain proof_out
- (fun ((body,ctx),eff) -> (under_binders env sigma redfun n body,ctx),eff) }
-
-let warn_implicits_in_term =
- CWarnings.create ~name:"implicits-in-term" ~category:"implicits"
- (fun () ->
- strbrk "Implicit arguments declaration relies on type." ++ spc () ++
- strbrk "The term declares more implicits than the type here.")
-
-let interp_definition pl bl poly red_option c ctypopt =
- let env = Global.env() in
- let evd, decl = Univdecls.interp_univ_decl_opt env pl in
- let evdref = ref evd in
- let impls, ((env_bl, ctx), imps1) = interp_context_evars env evdref bl in
- let ctx = List.map (fun d -> map_rel_decl EConstr.Unsafe.to_constr d) ctx in
- let nb_args = Context.Rel.nhyps ctx in
- let imps,pl,ce =
- match ctypopt with
- None ->
- let subst = evd_comb0 Evd.nf_univ_variables evdref in
- let ctx = Context.Rel.map (Vars.subst_univs_constr subst) ctx in
- let env_bl = push_rel_context ctx env in
- let c, imps2 = interp_constr_evars_impls ~impls env_bl evdref c in
- let c = EConstr.Unsafe.to_constr c in
- let nf,subst = Evarutil.e_nf_evars_and_universes evdref in
- let body = nf (it_mkLambda_or_LetIn c ctx) in
- let vars = Univops.universes_of_constr body in
- let evd = Evd.restrict_universe_context !evdref vars in
- let uctx = Evd.check_univ_decl ~poly evd decl in
- let binders = Evd.universe_binders evd in
- imps1@(Impargs.lift_implicits nb_args imps2), binders,
- definition_entry ~univs:uctx body
- | Some ctyp ->
- let ty, impsty = interp_type_evars_impls ~impls env_bl evdref ctyp in
- let subst = evd_comb0 Evd.nf_univ_variables evdref in
- let ctx = Context.Rel.map (Vars.subst_univs_constr subst) ctx in
- let env_bl = push_rel_context ctx env in
- let c, imps2 = interp_casted_constr_evars_impls ~impls env_bl evdref c ty in
- let c = EConstr.Unsafe.to_constr c in
- let nf, subst = Evarutil.e_nf_evars_and_universes evdref in
- let body = nf (it_mkLambda_or_LetIn c ctx) in
- let ty = EConstr.Unsafe.to_constr ty in
- let typ = nf (Term.it_mkProd_or_LetIn ty ctx) in
- let beq b1 b2 = if b1 then b2 else not b2 in
- let impl_eq (x,y,z) (x',y',z') = beq x x' && beq y y' && beq z z' in
- (* Check that all implicit arguments inferable from the term
- are inferable from the type *)
- let chk (key,va) =
- impl_eq (List.assoc_f Pervasives.(=) key impsty) va (* FIXME *)
- in
- if not (try List.for_all chk imps2 with Not_found -> false)
- then warn_implicits_in_term ();
- let vars = Univ.LSet.union (Univops.universes_of_constr body)
- (Univops.universes_of_constr typ) in
- let ctx = Evd.restrict_universe_context !evdref vars in
- let uctx = Evd.check_univ_decl ~poly ctx decl in
- let binders = Evd.universe_binders evd in
- imps1@(Impargs.lift_implicits nb_args impsty), binders,
- definition_entry ~types:typ
- ~univs:uctx body
- in
- red_constant_entry (Context.Rel.length ctx) ce !evdref red_option, !evdref, decl, pl, imps
-
-let check_definition (ce, evd, _, _, imps) =
- check_evars_are_solved (Global.env ()) evd Evd.empty;
- ce
-
-let do_definition ident k univdecl bl red_option c ctypopt hook =
- let (ce, evd, univdecl, pl', imps as def) =
- interp_definition univdecl bl (pi2 k) red_option c ctypopt
- in
- if Flags.is_program_mode () then
- let env = Global.env () in
- let (c,ctx), sideff = Future.force ce.const_entry_body in
- assert(Safe_typing.empty_private_constants = sideff);
- assert(Univ.ContextSet.is_empty ctx);
- let typ = match ce.const_entry_type with
- | Some t -> t
- | None -> EConstr.to_constr evd (Retyping.get_type_of env evd (EConstr.of_constr c))
- in
- Obligations.check_evars env evd;
- let obls, _, c, cty =
- Obligations.eterm_obligations env ident evd 0 c typ
- in
- let ctx = Evd.evar_universe_context evd in
- let hook = Lemmas.mk_hook (fun l r _ -> Lemmas.call_hook (fun exn -> exn) hook l r) in
- ignore(Obligations.add_definition
- ident ~term:c cty ctx ~univdecl ~implicits:imps ~kind:k ~hook obls)
- else let ce = check_definition def in
- ignore(DeclareDef.declare_definition ident k ce pl' imps
- (Lemmas.mk_hook
- (fun l r -> Lemmas.call_hook (fun exn -> exn) hook l r;r)))
-
-(* 2| Variable/Hypothesis/Parameter/Axiom declarations *)
-
-let declare_assumption is_coe (local,p,kind) (c,ctx) pl imps impl nl (_,ident) =
-match local with
-| Discharge when Lib.sections_are_opened () ->
- let ctx = match ctx with
- | Monomorphic_const_entry ctx -> ctx
- | Polymorphic_const_entry ctx -> Univ.ContextSet.of_context ctx
- in
- let decl = (Lib.cwd(), SectionLocalAssum ((c,ctx),p,impl), IsAssumption kind) in
- let _ = declare_variable ident decl in
- let () = assumption_message ident in
- let () =
- if not !Flags.quiet && Proof_global.there_are_pending_proofs () then
- Feedback.msg_info (str"Variable" ++ spc () ++ Id.print ident ++
- strbrk " is not visible from current goals")
- in
- let r = VarRef ident in
- let () = Typeclasses.declare_instance None true r in
- let () = if is_coe then Class.try_add_new_coercion r ~local:true false in
- (r,Univ.Instance.empty,true)
-
-| Global | Local | Discharge ->
- let local = DeclareDef.get_locality ident ~kind:"axiom" local in
- let inl = match nl with
- | NoInline -> None
- | DefaultInline -> Some (Flags.get_inline_level())
- | InlineAt i -> Some i
- in
- let decl = (ParameterEntry (None,(c,ctx),inl), IsAssumption kind) in
- let kn = declare_constant ident ~local decl in
- let gr = ConstRef kn in
- let () = maybe_declare_manual_implicits false gr imps in
- let () = Universes.register_universe_binders gr pl in
- let () = assumption_message ident in
- let () = Typeclasses.declare_instance None false gr in
- let () = if is_coe then Class.try_add_new_coercion gr ~local p in
- let inst = match ctx with
- | Polymorphic_const_entry ctx -> Univ.UContext.instance ctx
- | Monomorphic_const_entry _ -> Univ.Instance.empty
- in
- (gr,inst,Lib.is_modtype_strict ())
-
-let interp_assumption evdref env impls bl c =
- let c = mkCProdN ?loc:(local_binders_loc bl) bl c in
- let ty, impls = interp_type_evars_impls env evdref ~impls c in
- let ty = EConstr.Unsafe.to_constr ty in
- (ty, impls)
-
-(* When monomorphic the universe constraints are declared with the first declaration only. *)
-let next_uctx =
- let empty_uctx = Monomorphic_const_entry Univ.ContextSet.empty in
- function
- | Polymorphic_const_entry _ as uctx -> uctx
- | Monomorphic_const_entry _ -> empty_uctx
-
-let declare_assumptions idl is_coe k (c,uctx) pl imps nl =
- let refs, status, _ =
- List.fold_left (fun (refs,status,uctx) id ->
- let ref',u',status' =
- declare_assumption is_coe k (c,uctx) pl imps false nl id in
- (ref',u')::refs, status' && status, next_uctx uctx)
- ([],true,uctx) idl
- in
- List.rev refs, status
-
-
-let maybe_error_many_udecls = function
- | ((loc,id), Some _) ->
- user_err ?loc ~hdr:"many_universe_declarations"
- Pp.(str "When declaring multiple axioms in one command, " ++
- str "only the first is allowed a universe binder " ++
- str "(which will be shared by the whole block).")
- | (_, None) -> ()
-
-let process_assumptions_udecls kind l =
- let udecl, first_id = match l with
- | (coe, ((id, udecl)::rest, c))::rest' ->
- List.iter maybe_error_many_udecls rest;
- List.iter (fun (coe, (idl, c)) -> List.iter maybe_error_many_udecls idl) rest';
- udecl, id
- | (_, ([], _))::_ | [] -> assert false
- in
- let () = match kind, udecl with
- | (Discharge, _, _), Some _ when Lib.sections_are_opened () ->
- let loc = fst first_id in
- let msg = Pp.str "Section variables cannot be polymorphic." in
- user_err ?loc msg
- | _ -> ()
- in
- udecl, List.map (fun (coe, (idl, c)) -> coe, (List.map fst idl, c)) l
-
-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 evdref, udecl =
- let evd, udecl = Univdecls.interp_univ_decl_opt env udecl in
- ref evd, udecl
- in
- let l =
- if pi2 kind (* poly *) then
- (* Separate declarations so that A B : Type puts A and B in different levels. *)
- List.fold_right (fun (is_coe,(idl,c)) acc ->
- List.fold_right (fun id acc ->
- (is_coe, ([id], c)) :: acc) idl acc)
- l []
- else l
- in
- (* We intepret all declarations in the same evar_map, i.e. as a telescope. *)
- let _,l = List.fold_left_map (fun (env,ienv) (is_coe,(idl,c)) ->
- let t,imps = interp_assumption evdref env ienv [] c in
- let env =
- push_named_context (List.map (fun (_,id) -> LocalAssum (id,t)) idl) env in
- let ienv = List.fold_right (fun (_,id) ienv ->
- let impls = compute_internalization_data env Variable t imps in
- Id.Map.add id impls ienv) idl ienv in
- ((env,ienv),((is_coe,idl),t,imps)))
- (env,empty_internalization_env) l
- in
- let evd = solve_remaining_evars all_and_fail_flags env !evdref Evd.empty in
- (* The universe constraints come from the whole telescope. *)
- let evd = Evd.nf_constraints evd in
- let nf_evar c = EConstr.to_constr evd (EConstr.of_constr c) in
- let uvars, l = List.fold_left_map (fun uvars (coe,t,imps) ->
- let t = nf_evar t in
- let uvars = Univ.LSet.union uvars (Univops.universes_of_constr t) in
- uvars, (coe,t,imps))
- Univ.LSet.empty l
- in
- let evd = Evd.restrict_universe_context evd uvars in
- let uctx = Evd.check_univ_decl ~poly:(pi2 kind) evd udecl in
- let ubinders = Evd.universe_binders evd in
- pi2 (List.fold_left (fun (subst,status,uctx) ((is_coe,idl),t,imps) ->
- 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 (_,id) (c,u) -> (id, Universes.constr_of_global_univ (c,u)))
- idl refs
- in
- subst'@subst, status' && status, next_uctx uctx)
- ([], true, uctx) l)
-
-(* 3a| Elimination schemes for mutual inductive definitions *)
-
-(* 3b| Mutual inductive definitions *)
-
-let push_types env idl tl =
- List.fold_left2 (fun env id t -> Environ.push_rel (LocalAssum (Name id,t)) env)
- env idl tl
-
-type structured_one_inductive_expr = {
- ind_name : Id.t;
- ind_univs : Vernacexpr.universe_decl_expr option;
- ind_arity : constr_expr;
- ind_lc : (Id.t * constr_expr) list
-}
-
-type structured_inductive_expr =
- local_binder_expr list * structured_one_inductive_expr list
-
-let minductive_message warn = function
- | [] -> user_err Pp.(str "No inductive definition.")
- | [x] -> (Id.print x ++ str " is defined" ++
- if warn then str " as a non-primitive record" else mt())
- | l -> hov 0 (prlist_with_sep pr_comma Id.print l ++
- spc () ++ str "are defined")
-
-let check_all_names_different indl =
- let ind_names = List.map (fun ind -> ind.ind_name) indl in
- let cstr_names = List.map_append (fun ind -> List.map fst ind.ind_lc) indl in
- let l = List.duplicates Id.equal ind_names in
- let () = match l with
- | [] -> ()
- | t :: _ -> raise (InductiveError (SameNamesTypes t))
- in
- let l = List.duplicates Id.equal cstr_names in
- let () = match l with
- | [] -> ()
- | c :: _ -> raise (InductiveError (SameNamesConstructors (List.hd l)))
- in
- let l = List.intersect Id.equal ind_names cstr_names in
- match l with
- | [] -> ()
- | _ -> raise (InductiveError (SameNamesOverlap l))
-
-let mk_mltype_data evdref env assums arity indname =
- let is_ml_type = is_sort env !evdref (EConstr.of_constr arity) in
- (is_ml_type,indname,assums)
-
-let prepare_param = function
- | LocalAssum (na,t) -> Name.get_id na, LocalAssumEntry t
- | LocalDef (na,b,_) -> Name.get_id na, LocalDefEntry b
-
-(** Make the arity conclusion flexible to avoid generating an upper bound universe now,
- only if the universe does not appear anywhere else.
- This is really a hack to stay compatible with the semantics of template polymorphic
- inductives which are recognized when a "Type" appears at the end of the conlusion in
- the source syntax. *)
-
-let rec check_anonymous_type ind =
- let open Glob_term in
- match DAst.get ind with
- | GSort (GType []) -> true
- | GProd ( _, _, _, e)
- | GLetIn (_, _, _, e)
- | GLambda (_, _, _, e)
- | GApp (e, _)
- | GCast (e, _) -> check_anonymous_type e
- | _ -> false
-
-let make_conclusion_flexible evdref ty poly =
- if poly && Term.isArity ty then
- let _, concl = Term.destArity ty in
- match concl with
- | Type u ->
- (match Univ.universe_level u with
- | Some u ->
- evdref := Evd.make_flexible_variable !evdref ~algebraic:true u
- | None -> ())
- | _ -> ()
- else ()
-
-let is_impredicative env u =
- u = Prop Null || (is_impredicative_set env && u = Prop Pos)
-
-let interp_ind_arity env evdref ind =
- let c = intern_gen IsType env ind.ind_arity in
- let impls = Implicit_quantifiers.implicits_of_glob_constr ~with_products:true c in
- let (evd,t) = understand_tcc env !evdref ~expected_type:IsType c in
- evdref := evd;
- let pseudo_poly = check_anonymous_type c in
- let () = if not (Reductionops.is_arity env !evdref t) then
- user_err ?loc:(constr_loc ind.ind_arity) (str "Not an arity")
- in
- let t = EConstr.Unsafe.to_constr t in
- t, pseudo_poly, impls
-
-let interp_cstrs evdref env impls mldata arity ind =
- let cnames,ctyps = List.split ind.ind_lc in
- (* Complete conclusions of constructor types if given in ML-style syntax *)
- let ctyps' = List.map2 (complete_conclusion mldata) cnames ctyps in
- (* Interpret the constructor types *)
- let ctyps'', cimpls = List.split (List.map (interp_type_evars_impls evdref env ~impls %> on_fst EConstr.Unsafe.to_constr) ctyps') in
- (cnames, ctyps'', cimpls)
-
-let sign_level env evd sign =
- fst (List.fold_right
- (fun d (lev,env) ->
- match d with
- | LocalDef _ -> lev, push_rel d env
- | LocalAssum _ ->
- let s = destSort (Reduction.whd_all env
- (EConstr.Unsafe.to_constr (nf_evar 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))
- sign (Univ.type0m_univ,env))
-
-let sup_list min = List.fold_left Univ.sup min
-
-let extract_level env evd min tys =
- let sorts = List.map (fun ty ->
- let ctx, concl = Reduction.dest_prod_assum env ty in
- sign_level env evd (LocalAssum (Anonymous, concl) :: ctx)) tys
- in sup_list min sorts
-
-let is_flexible_sort evd u =
- match Univ.Universe.level u with
- | Some l -> Evd.is_flexible_level evd l
- | None -> false
-
-let inductive_levels env evdref poly arities inds =
- let destarities = List.map (fun x -> x, Reduction.dest_arity env x) arities in
- let levels = List.map (fun (x,(ctx,a)) ->
- if a = Prop Null then None
- else Some (univ_of_sort a)) destarities
- in
- let cstrs_levels, min_levels, sizes =
- CList.split3
- (List.map2 (fun (_,tys,_) (arity,(ctx,du)) ->
- let len = List.length tys in
- let minlev = Sorts.univ_of_sort du in
- let minlev =
- if len > 1 && not (is_impredicative env du) then
- Univ.sup minlev Univ.type0_univ
- else minlev
- in
- let minlev =
- (** Indices contribute. *)
- if Indtypes.is_indices_matter () && List.length ctx > 0 then (
- let ilev = sign_level env !evdref ctx in
- Univ.sup ilev minlev)
- else minlev
- in
- let clev = extract_level env !evdref minlev tys in
- (clev, minlev, len)) inds destarities)
- in
- (* Take the transitive closure of the system of constructors *)
- (* level constraints and remove the recursive dependencies *)
- let levels' = Universes.solve_constraints_system (Array.of_list levels)
- (Array.of_list cstrs_levels) (Array.of_list min_levels)
- in
- let evd, arities =
- CList.fold_left3 (fun (evd, arities) cu (arity,(ctx,du)) len ->
- if is_impredicative env du then
- (** Any product is allowed here. *)
- evd, arity :: arities
- else (** If in a predicative sort, or asked to infer the type,
- we take the max of:
- - indices (if in indices-matter mode)
- - constructors
- - Type(1) if there is more than 1 constructor
- *)
- (** Constructors contribute. *)
- let evd =
- if Sorts.is_set du then
- if not (Evd.check_leq evd cu Univ.type0_univ) then
- raise (Indtypes.InductiveError Indtypes.LargeNonPropInductiveNotInType)
- else evd
- else evd
- (* Evd.set_leq_sort env evd (Type cu) du *)
- in
- let evd =
- if len >= 2 && Univ.is_type0m_univ cu then
- (** "Polymorphic" type constraint and more than one constructor,
- should not land in Prop. Add constraint only if it would
- land in Prop directly (no informative arguments as well). *)
- Evd.set_leq_sort env evd (Prop Pos) du
- else evd
- in
- let duu = Sorts.univ_of_sort du in
- let evd =
- if not (Univ.is_small_univ duu) && Univ.Universe.equal cu duu then
- if is_flexible_sort evd duu && not (Evd.check_leq evd Univ.type0_univ duu) then
- Evd.set_eq_sort env evd (Prop Null) du
- else evd
- else Evd.set_eq_sort env evd (Type cu) du
- in
- (evd, arity :: arities))
- (!evdref,[]) (Array.to_list levels') destarities sizes
- in evdref := evd; List.rev arities
-
-let check_named (loc, na) = match na with
-| Name _ -> ()
-| Anonymous ->
- let msg = str "Parameters must be named." in
- user_err ?loc msg
-
-
-let check_param = function
-| CLocalDef (na, _, _) -> check_named na
-| CLocalAssum (nas, Default _, _) -> List.iter check_named nas
-| CLocalAssum (nas, Generalized _, _) -> ()
-| CLocalPattern (loc,_) ->
- 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 evd, decl = Univdecls.interp_univ_decl_opt env0 pl in
- let evdref = ref evd in
- let impls, ((env_params, ctx_params), userimpls) =
- interp_context_evars env0 evdref paramsl
- in
- let ctx_params = List.map (fun d -> map_rel_decl EConstr.Unsafe.to_constr d) ctx_params in
- let indnames = List.map (fun ind -> ind.ind_name) indl in
-
- (* Names of parameters as arguments of the inductive type (defs removed) *)
- let assums = List.filter is_local_assum ctx_params in
- let params = List.map (RelDecl.get_name %> Name.get_id) assums in
-
- (* Interpret the arities *)
- let arities = List.map (interp_ind_arity env_params evdref) indl in
-
- let fullarities = List.map (fun (c, _, _) -> Term.it_mkProd_or_LetIn c ctx_params) arities in
- let env_ar = push_types env0 indnames fullarities in
- let env_ar_params = push_rel_context ctx_params env_ar in
-
- (* Compute interpretation metadatas *)
- let indimpls = List.map (fun (_, _, impls) -> userimpls @
- lift_implicits (Context.Rel.nhyps ctx_params) impls) arities in
- let arities = List.map pi1 arities and aritypoly = List.map pi2 arities in
- let impls = compute_internalization_env env0 ~impls (Inductive (params,true)) indnames fullarities indimpls in
- let ntn_impls = compute_internalization_env env0 (Inductive (params,true)) indnames fullarities indimpls in
- let mldatas = List.map2 (mk_mltype_data evdref env_params params) arities indnames in
-
- let constructors =
- Metasyntax.with_syntax_protection (fun () ->
- (* Temporary declaration of notations and scopes *)
- List.iter (Metasyntax.set_notation_for_interpretation env_params ntn_impls) notations;
- (* Interpret the constructor types *)
- List.map3 (interp_cstrs env_ar_params evdref impls) mldatas arities indl)
- () in
-
- (* Try further to solve evars, and instantiate them *)
- let sigma = solve_remaining_evars all_and_fail_flags env_params !evdref Evd.empty in
- evdref := sigma;
- (* Compute renewed arities *)
- let nf,_ = e_nf_evars_and_universes evdref in
- let arities = List.map nf arities in
- let constructors = List.map (fun (idl,cl,impsl) -> (idl,List.map nf cl,impsl)) constructors in
- let _ = List.iter2 (fun ty poly -> make_conclusion_flexible evdref ty poly) arities aritypoly in
- let arities = inductive_levels env_ar_params evdref poly arities constructors in
- let nf',_ = e_nf_evars_and_universes evdref in
- let nf x = nf' (nf x) in
- let arities = List.map nf' arities in
- let constructors = List.map (fun (idl,cl,impsl) -> (idl,List.map nf' cl,impsl)) constructors in
- let ctx_params = Context.Rel.map nf ctx_params in
- let evd = !evdref in
- let uctx = Evd.check_univ_decl ~poly evd decl in
- List.iter (fun c -> check_evars env_params Evd.empty evd (EConstr.of_constr c)) arities;
- Context.Rel.iter (fun c -> check_evars env0 Evd.empty evd (EConstr.of_constr c)) ctx_params;
- List.iter (fun (_,ctyps,_) ->
- List.iter (fun c -> check_evars env_ar_params Evd.empty evd (EConstr.of_constr c)) ctyps)
- constructors;
-
- (* Build the inductive entries *)
- let entries = List.map4 (fun ind arity template (cnames,ctypes,cimpls) -> {
- mind_entry_typename = ind.ind_name;
- mind_entry_arity = arity;
- mind_entry_template = template;
- mind_entry_consnames = cnames;
- mind_entry_lc = ctypes
- }) indl arities aritypoly constructors in
- let impls =
- let len = Context.Rel.nhyps ctx_params in
- List.map2 (fun indimpls (_,_,cimpls) ->
- indimpls, List.map (fun impls ->
- userimpls @ (lift_implicits len impls)) cimpls) indimpls constructors
- in
- let univs =
- match uctx with
- | Polymorphic_const_entry uctx ->
- if cum then
- Cumulative_ind_entry (Universes.univ_inf_ind_from_universe_context uctx)
- else Polymorphic_ind_entry uctx
- | Monomorphic_const_entry uctx ->
- Monomorphic_ind_entry uctx
- in
- (* Build the mutual inductive entry *)
- let mind_ent =
- { mind_entry_params = List.map prepare_param ctx_params;
- mind_entry_record = None;
- mind_entry_finite = finite;
- mind_entry_inds = entries;
- mind_entry_private = if prv then Some false else None;
- mind_entry_universes = univs;
- }
- in
- (if poly && cum then
- Inductiveops.infer_inductive_subtyping env_ar evd mind_ent
- else mind_ent), Evd.universe_binders evd, impls
-
-(* Very syntactical equality *)
-let eq_local_binders bl1 bl2 =
- List.equal local_binder_eq bl1 bl2
-
-let extract_coercions indl =
- let mkqid (_,((_,id),_)) = qualid_of_ident id in
- let extract lc = List.filter (fun (iscoe,_) -> iscoe) lc in
- List.map mkqid (List.flatten(List.map (fun (_,_,_,lc) -> extract lc) indl))
-
-let extract_params indl =
- let paramsl = List.map (fun (_,params,_,_) -> params) indl in
- match paramsl with
- | [] -> anomaly (Pp.str "empty list of inductive types.")
- | params::paramsl ->
- if not (List.for_all (eq_local_binders params) paramsl) then user_err Pp.(str
- "Parameters should be syntactically the same for each inductive type.");
- params
-
-let extract_inductive indl =
- List.map (fun (((_,indname),pl),_,ar,lc) -> {
- ind_name = indname; ind_univs = pl;
- ind_arity = Option.cata (fun x -> x) (CAst.make @@ CSort (GType [])) ar;
- ind_lc = List.map (fun (_,((_,id),t)) -> (id,t)) lc
- }) indl
-
-let extract_mutual_inductive_declaration_components indl =
- let indl,ntnl = List.split indl in
- let params = extract_params indl in
- let coes = extract_coercions indl in
- let indl = extract_inductive indl in
- (params,indl), coes, List.flatten ntnl
-
-let is_recursive mie =
- let rec is_recursive_constructor lift typ =
- match Constr.kind typ with
- | Prod (_,arg,rest) ->
- not (EConstr.Vars.noccurn Evd.empty (** FIXME *) lift (EConstr.of_constr arg)) ||
- is_recursive_constructor (lift+1) rest
- | LetIn (na,b,t,rest) -> is_recursive_constructor (lift+1) rest
- | _ -> false
- in
- match mie.mind_entry_inds with
- | [ind] ->
- let nparams = List.length mie.mind_entry_params in
- List.exists (fun t -> is_recursive_constructor (nparams+1) t) ind.mind_entry_lc
- | _ -> false
-
-let declare_mutual_inductive_with_eliminations mie pl impls =
- (* spiwack: raises an error if the structure is supposed to be non-recursive,
- but isn't *)
- begin match mie.mind_entry_finite with
- | BiFinite when is_recursive mie ->
- if Option.has_some mie.mind_entry_record then
- user_err Pp.(str "Records declared with the keywords Record or Structure cannot be recursive. You can, however, define recursive records using the Inductive or CoInductive command.")
- else
- user_err Pp.(str ("Types declared with the keyword Variant cannot be recursive. Recursive types are defined with the Inductive and CoInductive command."))
- | _ -> ()
- end;
- let names = List.map (fun e -> e.mind_entry_typename) mie.mind_entry_inds in
- let (_, kn), prim = declare_mind mie in
- let mind = Global.mind_of_delta_kn kn in
- List.iteri (fun i (indimpls, constrimpls) ->
- let ind = (mind,i) in
- let gr = IndRef ind in
- maybe_declare_manual_implicits false gr indimpls;
- Universes.register_universe_binders gr pl;
- List.iteri
- (fun j impls ->
- maybe_declare_manual_implicits false
- (ConstructRef (ind, succ j)) impls)
- constrimpls)
- impls;
- let warn_prim = match mie.mind_entry_record with Some (Some _) -> not prim | _ -> false in
- Flags.if_verbose Feedback.msg_info (minductive_message warn_prim names);
- if mie.mind_entry_private == None
- then declare_default_schemes mind;
- mind
-
-type one_inductive_impls =
- Impargs.manual_explicitation list (* for inds *)*
- Impargs.manual_explicitation list list (* for constrs *)
-
-let do_mutual_inductive indl cum poly prv finite =
- let indl,coes,ntns = extract_mutual_inductive_declaration_components indl in
- (* Interpret the types *)
- let mie,pl,impls = interp_mutual_inductive indl ntns cum poly prv finite in
- (* Declare the mutual inductive block with its associated schemes *)
- ignore (declare_mutual_inductive_with_eliminations mie pl impls);
- (* Declare the possible notations of inductive types *)
- List.iter (Metasyntax.add_notation_interpretation (Global.env ())) ntns;
- (* Declare the coercions *)
- List.iter (fun qid -> Class.try_add_new_coercion (locate qid) ~local:false poly) coes;
- (* If positivity is assumed declares itself as unsafe. *)
- if Environ.deactivated_guard (Global.env ()) then Feedback.feedback Feedback.AddedAxiom else ()
-
-(* 3c| Fixpoints and co-fixpoints *)
-
-(* An (unoptimized) function that maps preorders to partial orders...
-
- Input: a list of associations (x,[y1;...;yn]), all yi distincts
- and different of x, meaning x<=y1, ..., x<=yn
-
- Output: a list of associations (x,Inr [y1;...;yn]), collecting all
- distincts yi greater than x, _or_, (x, Inl y) meaning that
- x is in the same class as y (in which case, x occurs
- nowhere else in the association map)
-
- partial_order : ('a * 'a list) list -> ('a * ('a,'a list) union) list
-*)
-
-let rec partial_order cmp = function
- | [] -> []
- | (x,xge)::rest ->
- let rec browse res xge' = function
- | [] ->
- let res = List.map (function
- | (z, Inr zge) when List.mem_f cmp x zge ->
- (z, Inr (List.union cmp zge xge'))
- | r -> r) res in
- (x,Inr xge')::res
- | y::xge ->
- let rec link y =
- try match List.assoc_f cmp y res with
- | Inl z -> link z
- | Inr yge ->
- if List.mem_f cmp x yge then
- let res = List.remove_assoc_f cmp y res in
- let res = List.map (function
- | (z, Inl t) ->
- if cmp t y then (z, Inl x) else (z, Inl t)
- | (z, Inr zge) ->
- if List.mem_f cmp y zge then
- (z, Inr (List.add_set cmp x (List.remove cmp y zge)))
- else
- (z, Inr zge)) res in
- browse ((y,Inl x)::res) xge' (List.union cmp xge (List.remove cmp x yge))
- else
- browse res (List.add_set cmp y (List.union cmp xge' yge)) xge
- with Not_found -> browse res (List.add_set cmp y xge') xge
- in link y
- in browse (partial_order cmp rest) [] xge
-
-let non_full_mutual_message x xge y yge isfix rest =
- let reason =
- if Id.List.mem x yge then
- Id.print y ++ str " depends on " ++ Id.print x ++ strbrk " but not conversely"
- else if Id.List.mem y xge then
- Id.print x ++ str " depends on " ++ Id.print y ++ strbrk " but not conversely"
- else
- Id.print y ++ str " and " ++ Id.print x ++ strbrk " are not mutually dependent" in
- let e = if List.is_empty rest then reason else strbrk "e.g., " ++ reason in
- let k = if isfix then "fixpoint" else "cofixpoint" in
- let w =
- if isfix
- then strbrk "Well-foundedness check may fail unexpectedly." ++ fnl()
- else mt () in
- strbrk "Not a fully mutually defined " ++ str k ++ fnl () ++
- str "(" ++ e ++ str ")." ++ fnl () ++ w
-
-let warn_non_full_mutual =
- CWarnings.create ~name:"non-full-mutual" ~category:"fixpoints"
- (fun (x,xge,y,yge,isfix,rest) ->
- non_full_mutual_message x xge y yge isfix rest)
-
-let check_mutuality env evd isfix fixl =
- let names = List.map fst fixl in
- let preorder =
- List.map (fun (id,def) ->
- (id, List.filter (fun id' -> not (Id.equal id id') && occur_var env evd id' (EConstr.of_constr def)) names))
- fixl in
- let po = partial_order Id.equal preorder in
- match List.filter (function (_,Inr _) -> true | _ -> false) po with
- | (x,Inr xge)::(y,Inr yge)::rest ->
- warn_non_full_mutual (x,xge,y,yge,isfix,rest)
- | _ -> ()
-
-type structured_fixpoint_expr = {
- fix_name : Id.t;
- fix_univs : universe_decl_expr option;
- fix_annot : Id.t Loc.located option;
- fix_binders : local_binder_expr list;
- fix_body : constr_expr option;
- fix_type : constr_expr
-}
-
-let interp_fix_context env evdref isfix fix =
- let before, after = if isfix then split_at_annot fix.fix_binders fix.fix_annot else [], fix.fix_binders in
- let impl_env, ((env', ctx), imps) = interp_context_evars env evdref before in
- let impl_env', ((env'', ctx'), imps') = interp_context_evars ~impl_env ~shift:(Context.Rel.nhyps ctx) env' evdref after in
- let annot = Option.map (fun _ -> List.length (assums_of_rel_context ctx)) fix.fix_annot in
- ((env'', ctx' @ ctx), (impl_env',imps @ imps'), annot)
-
-let interp_fix_ccl evdref impls (env,_) fix =
- let (c, impl) = interp_type_evars_impls ~impls env evdref fix.fix_type in
- (c, impl)
-
-let interp_fix_body env_rec evdref impls (_,ctx) fix ccl =
- let open EConstr in
- Option.map (fun body ->
- let env = push_rel_context ctx env_rec in
- let body = interp_casted_constr_evars env evdref ~impls body ccl in
- it_mkLambda_or_LetIn body ctx) fix.fix_body
-
-let build_fix_type (_,ctx) ccl = EConstr.it_mkProd_or_LetIn ccl ctx
-
-let prepare_recursive_declaration fixnames fixtypes fixdefs =
- let defs = List.map (subst_vars (List.rev fixnames)) fixdefs in
- let names = List.map (fun id -> Name id) fixnames in
- (Array.of_list names, Array.of_list fixtypes, Array.of_list defs)
-
-(* Jump over let-bindings. *)
-
-let compute_possible_guardness_evidences (ctx,_,recindex) =
- (* A recursive index is characterized by the number of lambdas to
- skip before finding the relevant inductive argument *)
- match recindex with
- | Some i -> [i]
- | None ->
- (* If recursive argument was not given by user, we try all args.
- An earlier approach was to look only for inductive arguments,
- but doing it properly involves delta-reduction, and it finally
- doesn't seem to worth the effort (except for huge mutual
- fixpoints ?) *)
- List.interval 0 (Context.Rel.nhyps ctx - 1)
-
-type recursive_preentry =
- Id.t list * constr option list * types list
-
-(* Wellfounded definition *)
-
-open Coqlib
-
-let contrib_name = "Program"
-let subtac_dir = [contrib_name]
-let fixsub_module = subtac_dir @ ["Wf"]
-let tactics_module = subtac_dir @ ["Tactics"]
-
-let init_reference dir s () = Coqlib.coq_reference "Command" dir s
-let init_constant dir s evdref =
- let (sigma, c) = Evarutil.new_global !evdref (Coqlib.coq_reference "Command" dir s)
- in evdref := sigma; c
-
-let make_ref l s = init_reference l s
-let fix_proto = init_constant tactics_module "fix_proto"
-let fix_sub_ref = make_ref fixsub_module "Fix_sub"
-let measure_on_R_ref = make_ref fixsub_module "MR"
-let well_founded = init_constant ["Init"; "Wf"] "well_founded"
-let mkSubset evdref name typ prop =
- let open EConstr in
- mkApp (Evarutil.e_new_global evdref (delayed_force build_sigma).typ,
- [| typ; mkLambda (name, typ, prop) |])
-let sigT = Lazy.from_fun build_sigma_type
-
-let make_qref s = Qualid (Loc.tag @@ qualid_of_string s)
-let lt_ref = make_qref "Init.Peano.lt"
-
-let rec telescope evdref l =
- let open EConstr in
- let open Vars in
- match l with
- | [] -> assert false
- | [LocalAssum (n, t)] -> t, [LocalDef (n, mkRel 1, t)], mkRel 1
- | LocalAssum (n, t) :: tl ->
- let ty, tys, (k, constr) =
- List.fold_left
- (fun (ty, tys, (k, constr)) decl ->
- let t = RelDecl.get_type decl in
- let pred = mkLambda (RelDecl.get_name decl, t, ty) in
- let ty = Evarutil.e_new_global evdref (Lazy.force sigT).typ in
- let intro = Evarutil.e_new_global evdref (Lazy.force sigT).intro in
- let sigty = mkApp (ty, [|t; pred|]) in
- let intro = mkApp (intro, [|lift k t; lift k pred; mkRel k; constr|]) in
- (sigty, pred :: tys, (succ k, intro)))
- (t, [], (2, mkRel 1)) tl
- in
- let (last, subst) = List.fold_right2
- (fun pred decl (prev, subst) ->
- let t = RelDecl.get_type decl in
- let p1 = Evarutil.e_new_global evdref (Lazy.force sigT).proj1 in
- let p2 = Evarutil.e_new_global evdref (Lazy.force sigT).proj2 in
- let proj1 = applist (p1, [t; pred; prev]) in
- let proj2 = applist (p2, [t; pred; prev]) in
- (lift 1 proj2, LocalDef (get_name decl, proj1, t) :: subst))
- (List.rev tys) tl (mkRel 1, [])
- in ty, (LocalDef (n, last, t) :: subst), constr
-
- | LocalDef (n, b, t) :: tl -> let ty, subst, term = telescope evdref tl in
- ty, (LocalDef (n, b, t) :: subst), lift 1 term
-
-let nf_evar_context sigma ctx =
- List.map (map_constr (fun c -> Evarutil.nf_evar sigma c)) ctx
-
-let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
- let open EConstr in
- let open Vars in
- 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 evd, decl = Univdecls.interp_univ_decl_opt env pl in
- let evdref = ref evd in
- let _, ((env', binders_rel), impls) = interp_context_evars env evdref bl in
- let len = List.length binders_rel in
- let top_env = push_rel_context binders_rel env in
- let top_arity = interp_type_evars top_env evdref arityc in
- let full_arity = it_mkProd_or_LetIn top_arity binders_rel in
- let argtyp, letbinders, make = telescope evdref binders_rel in
- let argname = Id.of_string "recarg" in
- let arg = LocalAssum (Name argname, argtyp) in
- let binders = letbinders @ [arg] in
- let binders_env = push_rel_context binders_rel env in
- let rel, _ = interp_constr_evars_impls env evdref r in
- let relty = Typing.unsafe_type_of env !evdref rel in
- let relargty =
- let error () =
- user_err ?loc:(constr_loc r)
- ~hdr:"Command.build_wellfounded"
- (Printer.pr_econstr_env env !evdref rel ++ str " is not an homogeneous binary relation.")
- in
- try
- let ctx, ar = Reductionops.splay_prod_n env !evdref 2 relty in
- match ctx, EConstr.kind !evdref ar with
- | [LocalAssum (_,t); LocalAssum (_,u)], Sort s
- when Sorts.is_prop (ESorts.kind !evdref s) && Reductionops.is_conv env !evdref t u -> t
- | _, _ -> error ()
- with e when CErrors.noncritical e -> error ()
- in
- let relargty = EConstr.Unsafe.to_constr relargty in
- let measure = interp_casted_constr_evars binders_env evdref measure relargty in
- let wf_rel, wf_rel_fun, measure_fn =
- let measure_body, measure =
- it_mkLambda_or_LetIn measure letbinders,
- it_mkLambda_or_LetIn measure binders
- in
- let comb = Evarutil.e_new_global evdref (delayed_force measure_on_R_ref) in
- let relargty = EConstr.of_constr relargty in
- let wf_rel = mkApp (comb, [| argtyp; relargty; rel; measure |]) in
- let wf_rel_fun x y =
- mkApp (rel, [| subst1 x measure_body;
- subst1 y measure_body |])
- in wf_rel, wf_rel_fun, measure
- in
- let wf_proof = mkApp (well_founded evdref, [| argtyp ; wf_rel |]) in
- let argid' = Id.of_string (Id.to_string argname ^ "'") in
- let wfarg len = LocalAssum (Name argid',
- mkSubset evdref (Name argid') argtyp
- (wf_rel_fun (mkRel 1) (mkRel (len + 1))))
- in
- let intern_bl = wfarg 1 :: [arg] in
- let _intern_env = push_rel_context intern_bl env in
- let proj = Evarutil.e_new_global evdref (delayed_force build_sigma).Coqlib.proj1 in
- let wfargpred = mkLambda (Name argid', argtyp, wf_rel_fun (mkRel 1) (mkRel 3)) in
- let projection = (* in wfarg :: arg :: before *)
- mkApp (proj, [| argtyp ; wfargpred ; mkRel 1 |])
- in
- let top_arity_let = it_mkLambda_or_LetIn top_arity letbinders in
- let intern_arity = substl [projection] top_arity_let in
- (* substitute the projection of wfarg for something,
- now intern_arity is in wfarg :: arg *)
- let intern_fun_arity_prod = it_mkProd_or_LetIn intern_arity [wfarg 1] in
- let intern_fun_binder = LocalAssum (Name (add_suffix recname "'"), intern_fun_arity_prod) in
- let curry_fun =
- let wfpred = mkLambda (Name argid', argtyp, wf_rel_fun (mkRel 1) (mkRel (2 * len + 4))) in
- let intro = Evarutil.e_new_global evdref (delayed_force build_sigma).Coqlib.intro in
- let arg = mkApp (intro, [| argtyp; wfpred; lift 1 make; mkRel 1 |]) in
- let app = mkApp (mkRel (2 * len + 2 (* recproof + orig binders + current binders *)), [| arg |]) in
- let rcurry = mkApp (rel, [| measure; lift len measure |]) in
- let lam = LocalAssum (Name (Id.of_string "recproof"), rcurry) in
- let body = it_mkLambda_or_LetIn app (lam :: binders_rel) in
- let ty = it_mkProd_or_LetIn (lift 1 top_arity) (lam :: binders_rel) in
- LocalDef (Name recname, body, ty)
- in
- let fun_bl = intern_fun_binder :: [arg] in
- let lift_lets = lift_rel_context 1 letbinders in
- let intern_body =
- let ctx = LocalAssum (Name recname, get_type curry_fun) :: binders_rel in
- let (r, l, impls, scopes) =
- Constrintern.compute_internalization_data env
- Constrintern.Recursive (EConstr.Unsafe.to_constr full_arity) impls
- in
- let newimpls = Id.Map.singleton recname
- (r, l, impls @ [(Some (Id.of_string "recproof", Impargs.Manual, (true, false)))],
- scopes @ [None]) in
- interp_casted_constr_evars (push_rel_context ctx env) evdref
- ~impls:newimpls body (EConstr.Unsafe.to_constr (lift 1 top_arity))
- in
- let intern_body_lam = it_mkLambda_or_LetIn intern_body (curry_fun :: lift_lets @ fun_bl) in
- let prop = mkLambda (Name argname, argtyp, top_arity_let) in
- let def =
- mkApp (Evarutil.e_new_global evdref (delayed_force fix_sub_ref),
- [| argtyp ; wf_rel ;
- Evarutil.e_new_evar env evdref
- ~src:(Loc.tag @@ Evar_kinds.QuestionMark (Evar_kinds.Define false,Anonymous)) wf_proof;
- prop |])
- in
- let def = Typing.e_solve_evars env evdref def in
- let _ = evdref := Evarutil.nf_evar_map !evdref in
- let def = mkApp (def, [|intern_body_lam|]) in
- let binders_rel = nf_evar_context !evdref binders_rel in
- let binders = nf_evar_context !evdref binders in
- let top_arity = Evarutil.nf_evar !evdref top_arity in
- let hook, recname, typ =
- if List.length binders_rel > 1 then
- let name = add_suffix recname "_func" in
- let hook l gr _ =
- let body = it_mkLambda_or_LetIn (mkApp (Evarutil.e_new_global evdref gr, [|make|])) binders_rel in
- let ty = it_mkProd_or_LetIn top_arity binders_rel in
- let ty = EConstr.Unsafe.to_constr ty in
- let univs = Evd.check_univ_decl ~poly !evdref decl in
- (*FIXME poly? *)
- let ce = definition_entry ~types:ty ~univs (EConstr.to_constr !evdref body) in
- (** 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 !evdref) in
- if Impargs.is_implicit_args () || not (List.is_empty impls) then
- Impargs.declare_manual_implicits false gr [impls]
- in
- let typ = it_mkProd_or_LetIn top_arity binders in
- hook, name, typ
- else
- let typ = it_mkProd_or_LetIn top_arity binders_rel in
- let hook l gr _ =
- if Impargs.is_implicit_args () || not (List.is_empty impls) then
- Impargs.declare_manual_implicits false gr [impls]
- in hook, recname, typ
- in
- let hook = Lemmas.mk_hook hook in
- let fullcoqc = EConstr.to_constr !evdref def in
- let fullctyp = EConstr.to_constr !evdref typ in
- Obligations.check_evars env !evdref;
- let evars, _, evars_def, evars_typ =
- Obligations.eterm_obligations env recname !evdref 0 fullcoqc fullctyp
- in
- let ctx = Evd.evar_universe_context !evdref in
- ignore(Obligations.add_definition recname ~term:evars_def ~univdecl:decl
- evars_typ ctx evars ~hook)
-
-let interp_recursive isfix fixl notations =
- let open Context.Named.Declaration in
- let open EConstr in
- let env = Global.env() in
- let fixnames = List.map (fun fix -> fix.fix_name) fixl in
-
- (* Interp arities allowing for unresolved types *)
- let all_universes =
- List.fold_right (fun sfe acc ->
- match sfe.fix_univs , acc with
- | None , acc -> acc
- | x , None -> x
- | Some ls , Some us ->
- let lsu = ls.univdecl_instance and usu = us.univdecl_instance in
- if not (CList.for_all2eq (fun x y -> Id.equal (snd x) (snd y)) lsu usu) then
- user_err Pp.(str "(co)-recursive definitions should all have the same universe binders");
- Some us) fixl None in
- let evd, decl = Univdecls.interp_univ_decl_opt env all_universes in
- let evdref = ref evd in
- let fixctxs, fiximppairs, fixannots =
- List.split3 (List.map (interp_fix_context env evdref isfix) fixl) in
- let fixctximpenvs, fixctximps = List.split fiximppairs in
- let fixccls,fixcclimps = List.split (List.map3 (interp_fix_ccl evdref) fixctximpenvs fixctxs fixl) in
- let fixtypes = List.map2 build_fix_type fixctxs fixccls in
- let fixtypes = List.map (fun c -> nf_evar !evdref c) fixtypes in
- let fiximps = List.map3
- (fun ctximps cclimps (_,ctx) -> ctximps@(Impargs.lift_implicits (Context.Rel.nhyps ctx) cclimps))
- fixctximps fixcclimps fixctxs in
- let rec_sign =
- List.fold_left2
- (fun env' id t ->
- if Flags.is_program_mode () then
- let sort = Evarutil.evd_comb1 (Typing.type_of ~refresh:true env) evdref t in
- let fixprot =
- try
- let app = mkApp (fix_proto evdref, [|sort; t|]) in
- Typing.e_solve_evars env evdref app
- with e when CErrors.noncritical e -> t
- in
- LocalAssum (id,fixprot) :: env'
- else LocalAssum (id,t) :: env')
- [] fixnames fixtypes
- in
- let env_rec = push_named_context rec_sign env in
-
- (* Get interpretation metadatas *)
- let fixtypes = List.map EConstr.Unsafe.to_constr fixtypes in
- let fixccls = List.map EConstr.Unsafe.to_constr fixccls in
- let impls = compute_internalization_env env Recursive fixnames fixtypes fiximps in
-
- (* Interp bodies with rollback because temp use of notations/implicit *)
- let fixdefs =
- Metasyntax.with_syntax_protection (fun () ->
- List.iter (Metasyntax.set_notation_for_interpretation env_rec impls) notations;
- List.map4
- (fun fixctximpenv -> interp_fix_body env_rec evdref (Id.Map.fold Id.Map.add fixctximpenv impls))
- fixctximpenvs fixctxs fixl fixccls)
- () in
-
- (* Instantiate evars and check all are resolved *)
- let evd = solve_unif_constraints_with_heuristics env_rec !evdref in
- let evd, nf = nf_evars_and_universes evd in
- let fixdefs = List.map (fun c -> Option.map EConstr.Unsafe.to_constr c) fixdefs in
- let fixdefs = List.map (Option.map nf) fixdefs in
- let fixtypes = List.map nf fixtypes in
- let fixctxs = List.map (fun (_,ctx) -> ctx) fixctxs in
-
- (* Build the fix declaration block *)
- (env,rec_sign,decl,evd), (fixnames,fixdefs,fixtypes), List.combine3 fixctxs fiximps fixannots
-
-let check_recursive isfix env evd (fixnames,fixdefs,_) =
- check_evars_are_solved env evd Evd.empty;
- 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)
- end
-
-let interp_fixpoint l ntns =
- let (env,_,pl,evd),fix,info = interp_recursive true l ntns in
- check_recursive true env evd fix;
- (fix,pl,Evd.evar_universe_context evd,info)
-
-let interp_cofixpoint l ntns =
- let (env,_,pl,evd),fix,info = interp_recursive false l ntns in
- check_recursive false env evd fix;
- (fix,pl,Evd.evar_universe_context evd,info)
-
-let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) indexes ntns =
- if List.exists Option.is_empty fixdefs then
- (* Some bodies to define by proof *)
- let thms =
- List.map3 (fun id t (ctx,imps,_) -> (id,(t,(List.map RelDecl.get_name ctx,imps))))
- fixnames fixtypes fiximps in
- let init_tac =
- 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)
- evd pl (Some(false,indexes,init_tac)) thms None (Lemmas.mk_hook (fun _ _ -> ()))
- else begin
- (* We shortcut the proof process *)
- let fixdefs = List.map Option.get fixdefs in
- let fixdecls = prepare_recursive_declaration fixnames fixtypes fixdefs in
- let env = Global.env() in
- let indexes = search_guard env indexes fixdecls in
- let fiximps = List.map (fun (n,r,p) -> r) fiximps in
- let vars = Univops.universes_of_constr (mkFix ((indexes,0),fixdecls)) in
- let fixdecls =
- List.map_i (fun i _ -> mkFix ((indexes,i),fixdecls)) 0 fixnames in
- let evd = Evd.from_ctx ctx in
- let evd = Evd.restrict_universe_context evd vars in
- let ctx = Evd.check_univ_decl ~poly evd pl in
- let pl = Evd.universe_binders evd in
- let fixdecls = List.map Safe_typing.mk_pure_proof fixdecls in
- ignore (List.map4 (DeclareDef.declare_fix (local, poly, Fixpoint) pl ctx)
- fixnames fixdecls fixtypes fiximps);
- (* Declare the recursive definitions *)
- fixpoint_message (Some indexes) fixnames;
- end;
- (* Declare notations *)
- List.iter (Metasyntax.add_notation_interpretation (Global.env())) ntns
-
-let declare_cofixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ntns =
- if List.exists Option.is_empty fixdefs then
- (* Some bodies to define by proof *)
- let thms =
- List.map3 (fun id t (ctx,imps,_) -> (id,(t,(List.map RelDecl.get_name ctx,imps))))
- fixnames fixtypes fiximps in
- let init_tac =
- 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 CoFixpoint)
- evd pl (Some(true,[],init_tac)) thms None (Lemmas.mk_hook (fun _ _ -> ()))
- else begin
- (* We shortcut the proof process *)
- let fixdefs = List.map Option.get fixdefs in
- let fixdecls = prepare_recursive_declaration fixnames fixtypes fixdefs in
- let fixdecls = List.map_i (fun i _ -> mkCoFix (i,fixdecls)) 0 fixnames in
- let vars = Univops.universes_of_constr (List.hd fixdecls) in
- let fixdecls = List.map Safe_typing.mk_pure_proof fixdecls in
- let fiximps = List.map (fun (len,imps,idx) -> imps) fiximps in
- let evd = Evd.from_ctx ctx in
- let evd = Evd.restrict_universe_context evd vars in
- let ctx = Evd.check_univ_decl ~poly evd pl in
- let pl = Evd.universe_binders evd in
- ignore (List.map4 (DeclareDef.declare_fix (local, poly, CoFixpoint) pl ctx)
- fixnames fixdecls fixtypes fiximps);
- (* Declare the recursive definitions *)
- cofixpoint_message fixnames
- end;
- (* Declare notations *)
- List.iter (Metasyntax.add_notation_interpretation (Global.env())) ntns
-
-let extract_decreasing_argument limit = function
- | (na,CStructRec) -> na
- | (na,_) when not limit -> na
- | _ -> user_err Pp.(str
- "Only structural decreasing is supported for a non-Program Fixpoint")
-
-let extract_fixpoint_components limit l =
- let fixl, ntnl = List.split l in
- let fixl = List.map (fun (((_,id),pl),ann,bl,typ,def) ->
- let ann = extract_decreasing_argument limit ann in
- {fix_name = id; fix_annot = ann; fix_univs = pl;
- fix_binders = bl; fix_body = def; fix_type = typ}) fixl in
- fixl, List.flatten ntnl
-
-let extract_cofixpoint_components l =
- let fixl, ntnl = List.split l in
- List.map (fun (((_,id),pl),bl,typ,def) ->
- {fix_name = id; fix_annot = None; fix_univs = pl;
- fix_binders = bl; fix_body = def; fix_type = typ}) fixl,
- List.flatten ntnl
-
-let out_def = function
- | Some def -> def
- | None -> user_err Pp.(str "Program Fixpoint needs defined bodies.")
-
-let collect_evars_of_term evd c ty =
- let evars = Evar.Set.union (Evd.evars_of_term c) (Evd.evars_of_term ty) in
- Evar.Set.fold (fun ev acc -> Evd.add acc ev (Evd.find_undefined evd ev))
- evars (Evd.from_ctx (Evd.evar_universe_context evd))
-
-let do_program_recursive local p fixkind fixl ntns =
- let isfix = fixkind != Obligations.IsCoFixpoint in
- let (env, rec_sign, pl, evd), fix, info =
- interp_recursive isfix fixl ntns
- in
- (* Program-specific code *)
- (* Get the interesting evars, those that were not instanciated *)
- let evd = Typeclasses.resolve_typeclasses ~filter:Typeclasses.no_goals ~fail:true env evd in
- (* Solve remaining evars *)
- let evd = nf_evar_map_undefined evd in
- let collect_evars id def typ imps =
- (* Generalize by the recursive prototypes *)
- let def =
- EConstr.to_constr 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)
- in
- let evm = collect_evars_of_term evd def typ in
- let evars, _, def, typ =
- Obligations.eterm_obligations env id evm
- (List.length rec_sign) def typ
- in (id, def, typ, imps, evars)
- in
- let (fixnames,fixdefs,fixtypes) = fix in
- let fiximps = List.map pi2 info in
- let fixdefs = List.map out_def fixdefs in
- let defs = List.map4 collect_evars fixnames fixdefs fixtypes fiximps in
- let () = if isfix then begin
- let possible_indexes = List.map compute_possible_guardness_evidences info in
- let fixdecls = Array.of_list (List.map (fun x -> Name x) fixnames),
- Array.of_list fixtypes,
- Array.of_list (List.map (subst_vars (List.rev fixnames)) fixdefs)
- in
- let indexes =
- Pretyping.search_guard (Global.env ()) possible_indexes fixdecls in
- List.iteri (fun i _ ->
- Inductive.check_fix env
- ((indexes,i),fixdecls))
- fixl
- end in
- let ctx = Evd.evar_universe_context evd in
- let kind = match fixkind with
- | Obligations.IsFixpoint _ -> (local, p, Fixpoint)
- | Obligations.IsCoFixpoint -> (local, p, CoFixpoint)
- in
- Obligations.add_mutual_definitions defs ~kind ~univdecl:pl ctx ntns fixkind
-
-let do_program_fixpoint local poly l =
- let g = List.map (fun ((_,wf,_,_,_),_) -> wf) l in
- match g, l with
- | [(n, CWfRec r)], [((((_,id),pl),_,bl,typ,def),ntn)] ->
- let recarg =
- match n with
- | Some n -> mkIdentC (snd n)
- | None ->
- user_err ~hdr:"do_program_fixpoint"
- (str "Recursive argument required for well-founded fixpoints")
- in build_wellfounded (id, pl, n, bl, typ, out_def def) poly r recarg ntn
-
- | [(n, CMeasureRec (m, r))], [((((_,id),pl),_,bl,typ,def),ntn)] ->
- build_wellfounded (id, pl, n, bl, typ, out_def def) poly
- (Option.default (CAst.make @@ CRef (lt_ref,None)) r) m ntn
-
- | _, _ when List.for_all (fun (n, ro) -> ro == CStructRec) g ->
- let fixl,ntns = extract_fixpoint_components true l in
- let fixkind = Obligations.IsFixpoint g in
- do_program_recursive local poly fixkind fixl ntns
-
- | _, _ ->
- user_err ~hdr:"do_program_fixpoint"
- (str "Well-founded fixpoints not allowed in mutually recursive blocks")
-
-let check_safe () =
- let open Declarations in
- let flags = Environ.typing_flags (Global.env ()) in
- flags.check_universes && flags.check_guarded
-
-let do_fixpoint local poly l =
- if Flags.is_program_mode () then do_program_fixpoint local poly l
- else
- let fixl, ntns = extract_fixpoint_components true l in
- let (_, _, _, info as fix) = interp_fixpoint fixl ntns in
- let possible_indexes =
- List.map compute_possible_guardness_evidences info in
- declare_fixpoint local poly fix possible_indexes ntns;
- if not (check_safe ()) then Feedback.feedback Feedback.AddedAxiom else ()
-
-let do_cofixpoint local poly l =
- let fixl,ntns = extract_cofixpoint_components l in
- if Flags.is_program_mode () then
- do_program_recursive local poly Obligations.IsCoFixpoint fixl ntns
- else
- let cofix = interp_cofixpoint fixl ntns in
- declare_cofixpoint local poly cofix ntns;
- if not (check_safe ()) then Feedback.feedback Feedback.AddedAxiom else ()
diff --git a/vernac/command.mli b/vernac/command.mli
deleted file mode 100644
index 070f3e112..000000000
--- a/vernac/command.mli
+++ /dev/null
@@ -1,163 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-open Names
-open Constr
-open Entries
-open Libnames
-open Globnames
-open Vernacexpr
-open Constrexpr
-open Decl_kinds
-open Redexpr
-
-(** This file is about the interpretation of raw commands into typed
- ones and top-level declaration of the main Gallina objects *)
-
-val do_universe : polymorphic -> Id.t Loc.located list -> unit
-val do_constraint : polymorphic ->
- (Misctypes.glob_level * Univ.constraint_type * Misctypes.glob_level) list -> unit
-
-(** {6 Definitions/Let} *)
-
-val interp_definition :
- Vernacexpr.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 * Universes.universe_binders * Impargs.manual_implicits
-
-val do_definition : Id.t -> definition_kind -> Vernacexpr.universe_decl_expr option ->
- local_binder_expr list -> red_expr option -> constr_expr ->
- constr_expr option -> unit Lemmas.declaration_hook -> unit
-
-(** {6 Parameters/Assumptions} *)
-
-(* val interp_assumption : env -> evar_map ref -> *)
-(* local_binder_expr list -> constr_expr -> *)
-(* types Univ.in_universe_context_set * Impargs.manual_implicits *)
-
-(** returns [false] if the assumption is neither local to a section,
- nor in a module type and meant to be instantiated. *)
-val declare_assumption : coercion_flag -> assumption_kind ->
- types in_constant_universes_entry ->
- Universes.universe_binders -> Impargs.manual_implicits ->
- bool (** implicit *) -> Vernacexpr.inline -> variable Loc.located ->
- global_reference * Univ.Instance.t * bool
-
-val do_assumptions : locality * polymorphic * assumption_object_kind ->
- Vernacexpr.inline -> (Vernacexpr.ident_decl list * constr_expr) with_coercion list -> bool
-
-(* val declare_assumptions : variable Loc.located list -> *)
-(* coercion_flag -> assumption_kind -> types Univ.in_universe_context_set -> *)
-(* Impargs.manual_implicits -> bool -> Vernacexpr.inline -> bool *)
-
-(** {6 Inductive and coinductive types} *)
-
-(** Extracting the semantical components out of the raw syntax of mutual
- inductive declarations *)
-
-type structured_one_inductive_expr = {
- ind_name : Id.t;
- ind_univs : Vernacexpr.universe_decl_expr option;
- ind_arity : constr_expr;
- ind_lc : (Id.t * constr_expr) list
-}
-
-type structured_inductive_expr =
- local_binder_expr list * structured_one_inductive_expr list
-
-val extract_mutual_inductive_declaration_components :
- (one_inductive_expr * decl_notation list) list ->
- structured_inductive_expr * (*coercions:*) qualid list * decl_notation list
-
-(** Typing mutual inductive definitions *)
-
-type one_inductive_impls =
- Impargs.manual_implicits (** for inds *)*
- Impargs.manual_implicits list (** for constrs *)
-
-val interp_mutual_inductive :
- structured_inductive_expr -> decl_notation list -> cumulative_inductive_flag ->
- polymorphic -> private_flag -> Decl_kinds.recursivity_kind ->
- mutual_inductive_entry * Universes.universe_binders * one_inductive_impls list
-
-(** Registering a mutual inductive definition together with its
- associated schemes *)
-
-val declare_mutual_inductive_with_eliminations :
- mutual_inductive_entry -> Universes.universe_binders -> one_inductive_impls list ->
- MutInd.t
-
-(** Entry points for the vernacular commands Inductive and CoInductive *)
-
-val do_mutual_inductive :
- (one_inductive_expr * decl_notation list) list -> cumulative_inductive_flag ->
- polymorphic -> private_flag -> Decl_kinds.recursivity_kind -> unit
-
-(** {6 Fixpoints and cofixpoints} *)
-
-type structured_fixpoint_expr = {
- fix_name : Id.t;
- fix_univs : Vernacexpr.universe_decl_expr option;
- fix_annot : Id.t Loc.located option;
- fix_binders : local_binder_expr list;
- fix_body : constr_expr option;
- fix_type : constr_expr
-}
-
-(** Extracting the semantical components out of the raw syntax of
- (co)fixpoints declarations *)
-
-val extract_fixpoint_components : bool ->
- (fixpoint_expr * decl_notation list) list ->
- structured_fixpoint_expr list * decl_notation list
-
-val extract_cofixpoint_components :
- (cofixpoint_expr * decl_notation list) list ->
- structured_fixpoint_expr list * decl_notation list
-
-(** Typing global fixpoints and cofixpoint_expr *)
-
-type recursive_preentry =
- Id.t list * constr option list * types list
-
-val interp_fixpoint :
- structured_fixpoint_expr list -> decl_notation list ->
- recursive_preentry * Univdecls.universe_decl * UState.t *
- (EConstr.rel_context * Impargs.manual_implicits * int option) list
-
-val interp_cofixpoint :
- structured_fixpoint_expr list -> decl_notation list ->
- recursive_preentry * Univdecls.universe_decl * UState.t *
- (EConstr.rel_context * Impargs.manual_implicits * int option) list
-
-(** Registering fixpoints and cofixpoints in the environment *)
-
-val declare_fixpoint :
- locality -> polymorphic ->
- recursive_preentry * Univdecls.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 *
- (Context.Rel.t * Impargs.manual_implicits * int option) list ->
- decl_notation list -> unit
-
-(** Entry points for the vernacular commands Fixpoint and CoFixpoint *)
-
-val do_fixpoint :
- (* When [false], assume guarded. *)
- locality -> polymorphic -> (fixpoint_expr * decl_notation list) list -> unit
-
-val do_cofixpoint :
- (* When [false], assume guarded. *)
- locality -> polymorphic -> (cofixpoint_expr * decl_notation list) list -> unit
-
-(** Utils *)
-
-val check_mutuality : Environ.env -> Evd.evar_map -> bool -> (Id.t * types) list -> unit
diff --git a/vernac/declareDef.ml b/vernac/declareDef.ml
index 980db4109..dfac78c04 100644
--- a/vernac/declareDef.ml
+++ b/vernac/declareDef.ml
@@ -36,7 +36,7 @@ let declare_global_definition ident ce local k pl imps =
let kn = declare_constant ident ~local (DefinitionEntry ce, IsDefinition k) in
let gr = ConstRef kn in
let () = maybe_declare_manual_implicits false gr imps in
- let () = Universes.register_universe_binders gr pl in
+ let () = Declare.declare_univ_binders gr pl in
let () = definition_message ident in
gr
@@ -49,6 +49,7 @@ let declare_definition ident (local, p, k) ce pl imps hook =
let () = definition_message ident in
let gr = VarRef ident in
let () = maybe_declare_manual_implicits false gr imps in
+ let () = Declare.declare_univ_binders gr pl in
let () = if Proof_global.there_are_pending_proofs () then
warn_definition_not_visible ident
in
diff --git a/vernac/explainErr.ml b/vernac/explainErr.ml
index 3a8e8fb43..fc3495796 100644
--- a/vernac/explainErr.ml
+++ b/vernac/explainErr.ml
@@ -32,6 +32,7 @@ let explain_exn_default = function
| Sys_error msg -> hov 0 (str "System error: " ++ guill msg)
| Out_of_memory -> hov 0 (str "Out of memory.")
| Stack_overflow -> hov 0 (str "Stack overflow.")
+ | Dynlink.Error e -> hov 0 (str "Dynlink error: " ++ str Dynlink.(error_message e))
| Timeout -> hov 0 (str "Timeout!")
| Sys.Break -> hov 0 (fnl () ++ str "User interrupt.")
(* Exceptions with pre-evaluated error messages *)
@@ -75,8 +76,7 @@ let process_vernac_interp_error exn = match fst exn with
wrap_vernac_error exn (Himsg.explain_pattern_matching_error env sigma e)
| Tacred.ReductionTacticError e ->
wrap_vernac_error exn (Himsg.explain_reduction_tactic_error e)
- | Logic.RefinerError e ->
- let sigma, env = Pfedit.get_current_context () in
+ | Logic.RefinerError (env, sigma, e) ->
wrap_vernac_error exn (Himsg.explain_refiner_error env sigma e)
| Nametab.GlobalizationError q ->
wrap_vernac_error exn
diff --git a/vernac/himsg.ml b/vernac/himsg.ml
index e8c5aeedd..f00c1e604 100644
--- a/vernac/himsg.ml
+++ b/vernac/himsg.ml
@@ -83,12 +83,12 @@ let rec contract3' env sigma a b c = function
(** Ad-hoc reductions *)
-let j_nf_betaiotaevar sigma j =
+let j_nf_betaiotaevar env sigma j =
{ uj_val = j.uj_val;
- uj_type = Reductionops.nf_betaiota sigma j.uj_type }
+ uj_type = Reductionops.nf_betaiota env sigma j.uj_type }
-let jv_nf_betaiotaevar sigma jl =
- Array.map (fun j -> j_nf_betaiotaevar sigma j) jl
+let jv_nf_betaiotaevar env sigma jl =
+ Array.map (fun j -> j_nf_betaiotaevar env sigma j) jl
(** Printers *)
@@ -258,7 +258,7 @@ let explain_number_branches env sigma cj expn =
str "expects " ++ int expn ++ str " branches."
let explain_ill_formed_branch env sigma c ci actty expty =
- let simp t = Reductionops.nf_betaiota sigma t in
+ let simp t = Reductionops.nf_betaiota env sigma t in
let env = make_all_name_different env sigma in
let pc = pr_leconstr_env env sigma c in
let pa, pe = pr_explicit env sigma (simp actty) (simp expty) in
@@ -295,8 +295,8 @@ let explain_unification_error env sigma p1 p2 = function
| NotSameArgSize | NotSameHead | NoCanonicalStructure ->
(* Error speaks from itself *) []
| ConversionFailed (env,t1,t2) ->
- let t1 = Reductionops.nf_betaiota sigma t1 in
- let t2 = Reductionops.nf_betaiota sigma t2 in
+ let t1 = Reductionops.nf_betaiota env sigma t1 in
+ let t2 = Reductionops.nf_betaiota env sigma t2 in
if EConstr.eq_constr sigma t1 p1 && EConstr.eq_constr sigma t2 p2 then [] else
let env = make_all_name_different env sigma in
if not (EConstr.eq_constr sigma t1 p1) || not (EConstr.eq_constr sigma t2 p2) then
@@ -336,8 +336,8 @@ let explain_unification_error env sigma p1 p2 = function
let explain_actual_type env sigma j t reason =
let env = make_all_name_different env sigma in
- let j = j_nf_betaiotaevar sigma j in
- let t = Reductionops.nf_betaiota sigma t in
+ let j = j_nf_betaiotaevar env sigma j in
+ let t = Reductionops.nf_betaiota env sigma t in
(** Actually print *)
let pe = pr_ne_context_of (str "In environment") env sigma in
let pc = pr_leconstr_env env sigma (Environ.j_val j) in
@@ -351,8 +351,8 @@ let explain_actual_type env sigma j t reason =
ppreason ++ str ".")
let explain_cant_apply_bad_type env sigma (n,exptyp,actualtyp) rator randl =
- let randl = jv_nf_betaiotaevar sigma randl in
- let actualtyp = Reductionops.nf_betaiota sigma actualtyp in
+ let randl = jv_nf_betaiotaevar env sigma randl in
+ let actualtyp = Reductionops.nf_betaiota env sigma actualtyp in
let env = make_all_name_different env sigma in
let actualtyp, exptyp = pr_explicit env sigma actualtyp exptyp in
let nargs = Array.length randl in
diff --git a/vernac/indschemes.ml b/vernac/indschemes.ml
index e4ca98749..2f4c95aff 100644
--- a/vernac/indschemes.ml
+++ b/vernac/indschemes.ml
@@ -258,7 +258,7 @@ let declare_one_induction_scheme ind =
let declare_induction_schemes kn =
let mib = Global.lookup_mind kn in
- if mib.mind_finite <> Decl_kinds.CoFinite then begin
+ if mib.mind_finite <> Declarations.CoFinite then begin
for i = 0 to Array.length mib.mind_packets - 1 do
declare_one_induction_scheme (kn,i);
done;
@@ -268,7 +268,7 @@ let declare_induction_schemes kn =
let declare_eq_decidability_gen internal names kn =
let mib = Global.lookup_mind kn in
- if mib.mind_finite <> Decl_kinds.CoFinite then
+ if mib.mind_finite <> Declarations.CoFinite then
ignore (define_mutual_scheme eq_dec_scheme_kind internal names kn)
let eq_dec_scheme_msg ind = (* TODO: mutual inductive case *)
@@ -512,7 +512,7 @@ let map_inductive_block f kn n = for i=0 to n-1 do f (kn,i) done
let declare_default_schemes kn =
let mib = Global.lookup_mind kn in
let n = Array.length mib.mind_packets in
- if !elim_flag && (mib.mind_finite <> BiFinite || !bifinite_elim_flag)
+ if !elim_flag && (mib.mind_finite <> Declarations.BiFinite || !bifinite_elim_flag)
&& mib.mind_typing_flags.check_guarded then
declare_induction_schemes kn;
if !case_flag then map_inductive_block declare_one_case_analysis_scheme kn n;
diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml
index 42631a15b..57436784b 100644
--- a/vernac/lemmas.ml
+++ b/vernac/lemmas.ml
@@ -13,12 +13,10 @@ open CErrors
open Util
open Pp
open Names
-open Term
open Constr
open Declarations
open Declareops
open Entries
-open Environ
open Nameops
open Globnames
open Decls
@@ -88,31 +86,31 @@ let adjust_guardness_conditions const = function
(mkFix ((indexes,0),fixdecls), ctx), eff
| _ -> (body, ctx), eff) }
-let find_mutually_recursive_statements thms =
+let find_mutually_recursive_statements sigma thms =
let n = List.length thms in
let inds = List.map (fun (id,(t,impls)) ->
- let (hyps,ccl) = decompose_prod_assum t in
+ let (hyps,ccl) = EConstr.decompose_prod_assum sigma t in
let x = (id,(t,impls)) in
- let whnf_hyp_hds = map_rel_context_in_env
- (fun env c -> EConstr.Unsafe.to_constr (fst (whd_all_stack env Evd.empty (EConstr.of_constr c))))
+ let whnf_hyp_hds = EConstr.map_rel_context_in_env
+ (fun env c -> fst (Reductionops.whd_all_stack env sigma c))
(Global.env()) hyps in
let ind_hyps =
List.flatten (List.map_i (fun i decl ->
let t = RelDecl.get_type decl in
- match Constr.kind t with
+ match EConstr.kind sigma t with
| Ind ((kn,_ as ind),u) when
let mind = Global.lookup_mind kn in
- mind.mind_finite <> Decl_kinds.CoFinite ->
+ mind.mind_finite <> Declarations.CoFinite ->
[ind,x,i]
| _ ->
- []) 0 (List.rev (List.filter RelDecl.is_local_assum whnf_hyp_hds))) in
+ []) 0 (List.rev (List.filter Context.Rel.Declaration.is_local_assum whnf_hyp_hds))) in
let ind_ccl =
- let cclenv = push_rel_context hyps (Global.env()) in
- let whnf_ccl,_ = whd_all_stack cclenv Evd.empty (EConstr.of_constr ccl) in
- match Constr.kind (EConstr.Unsafe.to_constr whnf_ccl) with
+ let cclenv = EConstr.push_rel_context hyps (Global.env()) in
+ let whnf_ccl,_ = whd_all_stack cclenv Evd.empty ccl in
+ match EConstr.kind sigma whnf_ccl with
| Ind ((kn,_ as ind),u) when
let mind = Global.lookup_mind kn in
- Int.equal mind.mind_ntypes n && mind.mind_finite == Decl_kinds.CoFinite ->
+ Int.equal mind.mind_ntypes n && mind.mind_finite == Declarations.CoFinite ->
[ind,x,0]
| _ ->
[] in
@@ -163,21 +161,21 @@ let find_mutually_recursive_statements thms =
in
(finite,guard,None), ordered_inds
-let look_for_possibly_mutual_statements = function
+let look_for_possibly_mutual_statements sigma = function
| [id,(t,impls)] ->
(* One non recursively proved theorem *)
None,[id,(t,impls)],None
| _::_ as thms ->
(* More than one statement and/or an explicit decreasing mark: *)
(* we look for a common inductive hyp or a common coinductive conclusion *)
- let recguard,ordered_inds = find_mutually_recursive_statements thms in
+ let recguard,ordered_inds = find_mutually_recursive_statements sigma thms in
let thms = List.map pi2 ordered_inds in
Some recguard,thms, Some (List.map (fun (_,_,i) -> succ i) ordered_inds)
| [] -> anomaly (Pp.str "Empty list of theorems.")
(* Saving a goal *)
-let save ?export_seff id const cstrs pl do_guard (locality,poly,kind) hook =
+let save ?export_seff id const uctx do_guard (locality,poly,kind) hook =
let fix_exn = Future.fix_exn_of const.Entries.const_entry_body in
try
let const = adjust_guardness_conditions const do_guard in
@@ -204,7 +202,7 @@ let save ?export_seff id const cstrs pl do_guard (locality,poly,kind) hook =
(locality, ConstRef kn)
in
definition_message id;
- Universes.register_universe_binders r (Option.default Universes.empty_binders pl);
+ Declare.declare_univ_binders r (UState.universe_binders uctx);
call_hook (fun exn -> exn) hook l r
with e when CErrors.noncritical e ->
let e = CErrors.push e in
@@ -212,17 +210,16 @@ let save ?export_seff id const cstrs pl do_guard (locality,poly,kind) hook =
let default_thm_id = Id.of_string "Unnamed_thm"
-let compute_proof_name locality = function
- | Some ((loc,id),pl) ->
- (* We check existence here: it's a bit late at Qed time *)
- if Nametab.exists_cci (Lib.make_path id) || is_section_variable id ||
- locality == Global && Nametab.exists_cci (Lib.make_path_except_section id)
- then
- user_err ?loc (Id.print id ++ str " already exists.");
- id
- | None ->
- let avoid = Id.Set.of_list (Proof_global.get_all_proof_names ()) in
- next_global_ident_away default_thm_id avoid
+let fresh_name_for_anonymous_theorem () =
+ let avoid = Id.Set.of_list (Proof_global.get_all_proof_names ()) in
+ next_global_ident_away default_thm_id avoid
+
+let check_name_freshness locality (loc,id) : unit =
+ (* We check existence here: it's a bit late at Qed time *)
+ if Nametab.exists_cci (Lib.make_path id) || is_section_variable id ||
+ locality == Global && Nametab.exists_cci (Lib.make_path_except_section id)
+ then
+ user_err ?loc (Id.print id ++ str " already exists.")
let save_remaining_recthms (locality,p,kind) norm univs body opaq i (id,(t_i,(_,imps))) =
let t_i = norm t_i in
@@ -286,17 +283,17 @@ let save_hook = ref ignore
let set_save_hook f = save_hook := f
let save_named ?export_seff proof =
- let id,const,(cstrs,pl),do_guard,persistence,hook = proof in
- save ?export_seff id const cstrs pl do_guard persistence hook
+ let id,const,uctx,do_guard,persistence,hook = proof in
+ save ?export_seff id const uctx do_guard persistence hook
let check_anonymity id save_ident =
if not (String.equal (atompart_of_id id) (Id.to_string (default_thm_id))) then
user_err Pp.(str "This command can only be used for unnamed theorem.")
let save_anonymous ?export_seff proof save_ident =
- let id,const,(cstrs,pl),do_guard,persistence,hook = proof in
+ let id,const,uctx,do_guard,persistence,hook = proof in
check_anonymity id save_ident;
- save ?export_seff save_ident const cstrs pl do_guard persistence hook
+ save ?export_seff save_ident const uctx do_guard persistence hook
(* Admitted *)
@@ -312,7 +309,7 @@ let admit (id,k,e) pl hook () =
| Local, _, _ | Discharge, _, _ -> warn_let_as_axiom id
in
let () = assumption_message id in
- Universes.register_universe_binders (ConstRef kn) (Option.default Universes.empty_binders pl);
+ Declare.declare_univ_binders (ConstRef kn) pl;
call_hook (fun exn -> exn) hook Global (ConstRef kn)
(* Starting a goal *)
@@ -330,8 +327,8 @@ let get_proof proof do_guard hook opacity =
let universe_proof_terminator compute_guard hook =
let open Proof_global in
make_terminator begin function
- | Admitted (id,k,pe,(ctx,pl)) ->
- admit (id,k,pe) pl (hook (Some ctx)) ();
+ | Admitted (id,k,pe,ctx) ->
+ admit (id,k,pe) (UState.universe_binders ctx) (hook (Some ctx)) ();
Feedback.feedback Feedback.AddedAxiom
| Proved (opaque,idopt,proof) ->
let is_opaque, export_seff = match opaque with
@@ -339,7 +336,7 @@ let universe_proof_terminator compute_guard hook =
| Vernacexpr.Opaque -> true, false
in
let proof = get_proof proof compute_guard
- (hook (Some (fst proof.Proof_global.universes))) is_opaque in
+ (hook (Some (proof.Proof_global.universes))) is_opaque in
begin match idopt with
| None -> save_named ~export_seff proof
| Some (_,id) -> save_anonymous ~export_seff proof id
@@ -377,7 +374,7 @@ let start_proof_univs id ?pl kind sigma ?terminator ?sign c ?init_tac ?(compute_
let rec_tac_initializer finite guard thms snl =
if finite then
- match List.map (fun (id,(t,_)) -> (id,EConstr.of_constr t)) thms with
+ match List.map (fun (id,(t,_)) -> (id,t)) thms with
| (id,_)::l -> Tactics.mutual_cofix id l 0
| _ -> assert false
else
@@ -385,11 +382,11 @@ let rec_tac_initializer finite guard thms snl =
let nl = match snl with
| None -> List.map succ (List.map List.last guard)
| Some nl -> nl
- in match List.map2 (fun (id,(t,_)) n -> (id,n, EConstr.of_constr t)) thms nl with
+ in match List.map2 (fun (id,(t,_)) n -> (id,n, t)) thms nl with
| (id,n,_)::l -> Tactics.mutual_fix id n l 0
| _ -> assert false
-let start_proof_with_initialization kind ctx decl recguard thms snl hook =
+let start_proof_with_initialization kind sigma decl recguard thms snl hook =
let intro_tac (_, (_, (ids, _))) =
Tacticals.New.tclMAP (function
| Name id -> Tactics.intro_mustbe_force id
@@ -417,47 +414,45 @@ let start_proof_with_initialization kind ctx decl recguard thms snl hook =
| (id,(t,(_,imps)))::other_thms ->
let hook ctx strength ref =
let ctx = match ctx with
- | None -> Evd.empty_evar_universe_context
+ | None -> UState.empty
| Some ctx -> ctx
in
let other_thms_data =
if List.is_empty other_thms then [] else
(* there are several theorems defined mutually *)
let body,opaq = retrieve_first_recthm ctx ref in
- let subst = Evd.evar_universe_context_subst ctx in
- let norm c = Universes.subst_opt_univs_constr subst c in
- let ctx = UState.check_univ_decl ~poly:(pi2 kind) ctx decl in
- let body = Option.map norm body in
- List.map_i (save_remaining_recthms kind norm ctx body opaq) 1 other_thms in
+ let norm c = EConstr.to_constr (Evd.from_ctx ctx) c in
+ let body = Option.map EConstr.of_constr body in
+ let uctx = UState.check_univ_decl ~poly:(pi2 kind) ctx decl in
+ List.map_i (save_remaining_recthms kind norm uctx body opaq) 1 other_thms in
let thms_data = (strength,ref,imps)::other_thms_data in
List.iter (fun (strength,ref,imps) ->
maybe_declare_manual_implicits false ref imps;
call_hook (fun exn -> exn) hook strength ref) thms_data in
- start_proof_univs id ~pl:decl kind ctx (EConstr.of_constr t) ?init_tac (fun ctx -> mk_hook (hook ctx)) ~compute_guard:guard
+ start_proof_univs id ~pl:decl kind sigma t ?init_tac (fun ctx -> mk_hook (hook ctx)) ~compute_guard:guard
let start_proof_com ?inference_hook kind thms hook =
let env0 = Global.env () in
let decl = fst (List.hd thms) in
- let evd, decl =
- match decl with
- | None -> Evd.from_env env0, Univdecls.default_univ_decl
- | Some decl ->
- Univdecls.interp_univ_decl_opt env0 (snd decl) in
- let evdref = ref evd in
- let thms = List.map (fun (sopt,(bl,t)) ->
- let impls, ((env, ctx), imps) = interp_context_evars env0 evdref bl in
- let t', imps' = interp_type_evars_impls ~impls env evdref t in
+ let evd, decl = Univdecls.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
let flags = all_and_fail_flags in
let flags = { flags with use_hook = inference_hook } in
- evdref := solve_remaining_evars flags env !evdref Evd.empty;
+ let evd = solve_remaining_evars flags env evd Evd.empty in
let ids = List.map RelDecl.get_name ctx in
- (compute_proof_name (pi1 kind) sopt,
- (EConstr.Unsafe.to_constr (nf_evar !evdref (EConstr.it_mkProd_or_LetIn t' ctx)),
- (ids, imps @ lift_implicits (Context.Rel.nhyps ctx) imps'))))
- thms in
- let recguard,thms,snl = look_for_possibly_mutual_statements thms in
- let evd, nf = Evarutil.nf_evars_and_universes !evdref in
- let thms = List.map (fun (n, (t, info)) -> (n, (nf t, info))) thms in
+ check_name_freshness (pi1 kind) id;
+ (* XXX: The nf_evar is critical !! *)
+ evd, (snd id,
+ (Evarutil.nf_evar evd (EConstr.it_mkProd_or_LetIn t' ctx),
+ (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
+ (* 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
if not (decl.univdecl_extensible_instance && decl.univdecl_extensible_constraints) then
@@ -470,7 +465,6 @@ let start_proof_com ?inference_hook kind thms hook =
in
start_proof_with_initialization kind evd decl recguard thms snl hook
-
(* Saving a proof *)
let keep_admitted_vars = ref true
@@ -496,7 +490,7 @@ let save_proof ?proof = function
if const_entry_type = None then
user_err Pp.(str "Admitted requires an explicit statement");
let typ = Option.get const_entry_type in
- let ctx = UState.const_univ_entry ~poly:(pi2 k) (fst universes) in
+ let ctx = UState.const_univ_entry ~poly:(pi2 k) universes in
let sec_vars = if !keep_admitted_vars then const_entry_secctx else None in
Admitted(id, k, (sec_vars, (typ, ctx), None), universes)
| None ->
@@ -518,12 +512,9 @@ let save_proof ?proof = function
Some (Environ.keep_hyps env (Id.Set.union ids_typ ids_def))
| _ -> None in
let decl = Proof_global.get_universe_decl () in
- let evd = Evd.from_ctx universes in
let poly = pi2 k in
- let ctx = Evd.check_univ_decl ~poly evd decl in
- let binders = if poly then Some (UState.universe_binders universes) else None in
- Admitted(id,k,(sec_vars, (typ, ctx), None),
- (universes, binders))
+ let ctx = UState.check_univ_decl ~poly universes decl in
+ Admitted(id,k,(sec_vars, (typ, ctx), None), universes)
in
Proof_global.apply_terminator (Proof_global.get_terminator ()) pe
| Vernacexpr.Proved (is_opaque,idopt) ->
diff --git a/vernac/lemmas.mli b/vernac/lemmas.mli
index 1b1304db5..126dcd8b0 100644
--- a/vernac/lemmas.mli
+++ b/vernac/lemmas.mli
@@ -7,7 +7,6 @@
(************************************************************************)
open Names
-open Constr
open Decl_kinds
type 'a declaration_hook
@@ -37,11 +36,11 @@ val start_proof_com :
goal_kind -> Vernacexpr.proof_expr list ->
unit declaration_hook -> unit
-val start_proof_with_initialization :
+val start_proof_with_initialization :
goal_kind -> Evd.evar_map -> Univdecls.universe_decl ->
(bool * Proof_global.lemma_possible_guards * unit Proofview.tactic list option) option ->
(Id.t (* name of thm *) *
- (types (* type of thm *) * (Name.t list (* names to pre-introduce *) * Impargs.manual_explicitation list))) list
+ (EConstr.types (* type of thm *) * (Name.t list (* names to pre-introduce *) * Impargs.manual_explicitation list))) list
-> int list option -> unit declaration_hook -> unit
val universe_proof_terminator :
@@ -53,10 +52,12 @@ val standard_proof_terminator :
Proof_global.lemma_possible_guards -> unit declaration_hook ->
Proof_global.proof_terminator
+val fresh_name_for_anonymous_theorem : unit -> Id.t
+
(** {6 ... } *)
(** A hook the next three functions pass to cook_proof *)
-val set_save_hook : (Proof.proof -> unit) -> unit
+val set_save_hook : (Proof.t -> unit) -> unit
val save_proof : ?proof:Proof_global.closed_proof -> Vernacexpr.proof_end -> unit
diff --git a/vernac/locality.ml b/vernac/locality.ml
index 681b1ab20..87b411625 100644
--- a/vernac/locality.ml
+++ b/vernac/locality.ml
@@ -6,36 +6,14 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Decl_kinds
+
(** * Managing locality *)
let local_of_bool = function
- | true -> Decl_kinds.Local
- | false -> Decl_kinds.Global
-
-(** Extracting the locality flag *)
-
-(* Commands which supported an inlined Local flag *)
-
-let warn_deprecated_local_syntax =
- CWarnings.create ~name:"deprecated-local-syntax" ~category:"deprecated"
- (fun () ->
- Pp.strbrk "Deprecated syntax: use \"Local\" as a prefix.")
-
-let enforce_locality_full locality_flag local =
- let local =
- match locality_flag with
- | Some false when local ->
- CErrors.user_err Pp.(str "Cannot be simultaneously Local and Global.")
- | Some true when local ->
- CErrors.user_err Pp.(str "Use only prefix \"Local\".")
- | None ->
- if local then begin
- warn_deprecated_local_syntax ();
- Some true
- end else
- None
- | Some b -> Some b in
- local
+ | true -> Local
+ | false -> Global
+
(** Positioning locality for commands supporting discharging and export
outside of modules *)
@@ -48,15 +26,16 @@ let make_non_locality = function Some false -> false | _ -> true
let make_locality = function Some true -> true | _ -> false
-let enforce_locality locality_flag local =
- make_locality (enforce_locality_full locality_flag local)
+let enforce_locality_exp locality_flag discharge =
+ match locality_flag, discharge with
+ | Some b, NoDischarge -> local_of_bool b
+ | None, NoDischarge -> Global
+ | None, DoDischarge -> Discharge
+ | Some true, DoDischarge -> CErrors.user_err Pp.(str "Local not allowed in this case")
+ | Some false, DoDischarge -> CErrors.user_err Pp.(str "Global not allowed in this case")
-let enforce_locality_exp locality_flag local =
- match locality_flag, local with
- | None, Some local -> local
- | Some b, None -> local_of_bool b
- | None, None -> Decl_kinds.Global
- | Some _, Some _ -> CErrors.user_err Pp.(str "Local non allowed in this case")
+let enforce_locality locality_flag =
+ make_locality locality_flag
(* For commands whose default is to not discharge but to export:
Global in sections forces discharge, Global not in section is the default;
@@ -65,8 +44,8 @@ let enforce_locality_exp locality_flag local =
let make_section_locality =
function Some b -> b | None -> Lib.sections_are_opened ()
-let enforce_section_locality locality_flag local =
- make_section_locality (enforce_locality_full locality_flag local)
+let enforce_section_locality locality_flag =
+ make_section_locality locality_flag
(** Positioning locality for commands supporting export but not discharge *)
@@ -83,5 +62,5 @@ let make_module_locality = function
| Some true -> true
| None -> false
-let enforce_module_locality locality_flag local =
- make_module_locality (enforce_locality_full locality_flag local)
+let enforce_module_locality locality_flag =
+ make_module_locality locality_flag
diff --git a/vernac/locality.mli b/vernac/locality.mli
index bef66d8bc..922538b23 100644
--- a/vernac/locality.mli
+++ b/vernac/locality.mli
@@ -8,10 +8,6 @@
(** * Managing locality *)
-(** Commands which supported an inlined Local flag *)
-
-val enforce_locality_full : bool option -> bool -> bool option
-
(** * Positioning locality for commands supporting discharging and export
outside of modules *)
@@ -22,16 +18,15 @@ val enforce_locality_full : bool option -> bool -> bool option
val make_locality : bool option -> bool
val make_non_locality : bool option -> bool
-val enforce_locality : bool option -> bool -> bool
-val enforce_locality_exp :
- bool option -> Decl_kinds.locality option -> Decl_kinds.locality
+val enforce_locality_exp : bool option -> Decl_kinds.discharge -> Decl_kinds.locality
+val enforce_locality : bool option -> bool
(** For commands whose default is to not discharge but to export:
Global in sections forces discharge, Global not in section is the default;
Local in sections is the default, Local not in section forces non-export *)
val make_section_locality : bool option -> bool
-val enforce_section_locality : bool option -> bool -> bool
+val enforce_section_locality : bool option -> bool
(** * Positioning locality for commands supporting export but not discharge *)
@@ -40,4 +35,4 @@ val enforce_section_locality : bool option -> bool -> bool
Local in sections is the default, Local not in section forces non-export *)
val make_module_locality : bool option -> bool
-val enforce_module_locality : bool option -> bool -> bool
+val enforce_module_locality : bool option -> bool
diff --git a/vernac/metasyntax.ml b/vernac/metasyntax.ml
index 6c3dfec7d..bcffe640c 100644
--- a/vernac/metasyntax.ml
+++ b/vernac/metasyntax.ml
@@ -43,13 +43,6 @@ let add_token_obj s = Lib.add_anonymous_leaf (inToken s)
let entry_buf = Buffer.create 64
-type any_entry = AnyEntry : 'a Pcoq.Gram.entry -> any_entry
-
-let grammars : any_entry list String.Map.t ref = ref String.Map.empty
-
-let register_grammar name grams =
- grammars := String.Map.add name grams !grammars
-
let pr_entry e =
let () = Buffer.clear entry_buf in
let ft = Format.formatter_of_buffer entry_buf in
@@ -57,11 +50,11 @@ let pr_entry e =
str (Buffer.contents entry_buf)
let pr_registered_grammar name =
- let gram = try Some (String.Map.find name !grammars) with Not_found -> None in
+ let gram = try Some (Pcoq.find_grammars_by_name name) with Not_found -> None in
match gram with
| None -> user_err Pp.(str "Unknown or unprintable grammar entry.")
| Some entries ->
- let pr_one (AnyEntry e) =
+ let pr_one (Pcoq.AnyEntry e) =
str "Entry " ++ str (Pcoq.Gram.Entry.name e) ++ str " is" ++ fnl () ++
pr_entry e
in
@@ -80,8 +73,8 @@ let pr_grammar = function
| "pattern" ->
pr_entry Pcoq.Constr.pattern
| "vernac" ->
- str "Entry vernac is" ++ fnl () ++
- pr_entry Pcoq.Vernac_.vernac ++
+ str "Entry vernac_control is" ++ fnl () ++
+ pr_entry Pcoq.Vernac_.vernac_control ++
str "Entry command is" ++ fnl () ++
pr_entry Pcoq.Vernac_.command ++
str "Entry syntax is" ++ fnl () ++
@@ -199,36 +192,6 @@ let parse_format ((loc, str) : lstring) =
(***********************)
(* Analyzing notations *)
-type symbol_token = WhiteSpace of int | String of string
-
-let split_notation_string str =
- let push_token beg i l =
- if Int.equal beg i then l else
- let s = String.sub str beg (i - beg) in
- String s :: l
- in
- let push_whitespace beg i l =
- if Int.equal beg i then l else WhiteSpace (i-beg) :: l
- in
- let rec loop beg i =
- if i < String.length str then
- if str.[i] == ' ' then
- push_token beg i (loop_on_whitespace (i+1) (i+1))
- else
- loop beg (i+1)
- else
- push_token beg i []
- and loop_on_whitespace beg i =
- if i < String.length str then
- if str.[i] != ' ' then
- push_whitespace beg i (loop i (i+1))
- else
- loop_on_whitespace beg (i+1)
- else
- push_whitespace beg i []
- in
- loop 0 0
-
(* Interpret notations with a recursive component *)
let out_nt = function NonTerminal x -> x | _ -> assert false
@@ -284,17 +247,6 @@ let quote_notation_token x =
if (n > 0 && norm) || (n > 2 && x.[0] == '\'') then "'"^x^"'"
else x
-let rec raw_analyze_notation_tokens = function
- | [] -> []
- | String ".." :: sl -> NonTerminal ldots_var :: raw_analyze_notation_tokens sl
- | String "_" :: _ -> user_err Pp.(str "_ must be quoted.")
- | String x :: sl when CLexer.is_ident x ->
- NonTerminal (Names.Id.of_string x) :: raw_analyze_notation_tokens sl
- | String s :: sl ->
- Terminal (String.drop_simple_quotes s) :: raw_analyze_notation_tokens sl
- | WhiteSpace n :: sl ->
- Break n :: raw_analyze_notation_tokens sl
-
let is_numeral symbs =
match List.filter (function Break _ -> false | _ -> true) symbs with
| ([Terminal "-"; Terminal x] | [Terminal x]) ->
@@ -315,8 +267,8 @@ let rec get_notation_vars onlyprint = function
| (Terminal _ | Break _) :: sl -> get_notation_vars onlyprint sl
| SProdList _ :: _ -> assert false
-let analyze_notation_tokens ~onlyprint l =
- let l = raw_analyze_notation_tokens l in
+let analyze_notation_tokens ~onlyprint ntn =
+ let l = decompose_raw_notation ntn in
let vars = get_notation_vars onlyprint l in
let recvars,l = interp_list_parser [] l in
recvars, List.subtract Id.equal vars (List.map snd recvars), l
@@ -333,13 +285,17 @@ let prec_assoc = function
| LeftA -> (E,L)
| NonA -> (L,L)
-let precedence_of_entry_type from = function
- | ETConstr (NumLevel n,BorderProd (_,None)) -> n, Prec n
- | ETConstr (NumLevel n,BorderProd (b,Some a)) ->
+let precedence_of_position_and_level from = function
+ | NumLevel n, BorderProd (_,None) -> n, Prec n
+ | NumLevel n, BorderProd (b,Some a) ->
n, let (lp,rp) = prec_assoc a in if b == Left then lp else rp
- | ETConstr (NumLevel n,InternalProd) -> n, Prec n
- | ETConstr (NextLevel,_) -> from, L
- | _ -> 0, E (* ?? *)
+ | NumLevel n, InternalProd -> n, Prec n
+ | NextLevel, _ -> from, L
+
+let precedence_of_entry_type from = function
+ | ETConstr x | ETConstrAsBinder (_,x) -> precedence_of_position_and_level from x
+ | ETPattern (_,n) -> let n = match n with None -> 0 | Some n -> n in n, Prec n
+ | _ -> 0, E (* should not matter *)
(* Some breaking examples *)
(* "x = y" : "x /1 = y" (breaks before any symbol) *)
@@ -377,8 +333,8 @@ let is_non_terminal = function
| NonTerminal _ | SProdList _ -> true
| _ -> false
-let is_next_non_terminal = function
-| [] -> false
+let is_next_non_terminal b = function
+| [] -> b
| pr :: _ -> is_non_terminal pr
let is_next_terminal = function Terminal _ :: _ -> true | _ -> false
@@ -387,8 +343,9 @@ let is_next_break = function Break _ :: _ -> true | _ -> false
let add_break n l = (None,UnpCut (PpBrk(n,0))) :: l
-let add_break_if_none n = function
- | (((_,UnpCut (PpBrk _)) :: _) | []) as l -> l
+let add_break_if_none n b = function
+ | (_,UnpCut (PpBrk _)) :: _ as l -> l
+ | [] when not b -> []
| l -> (None,UnpCut (PpBrk(n,0))) :: l
let check_open_binder isopen sl m =
@@ -403,45 +360,59 @@ let check_open_binder isopen sl m =
prlist_with_sep spc pr_token sl
++ strbrk "\" is allowed to occur.")
+let unparsing_metavar i from typs =
+ let x = List.nth typs (i-1) in
+ let prec = snd (precedence_of_entry_type from x) in
+ match x with
+ | ETConstr _ | ETConstrAsBinder _ | ETReference | ETBigint ->
+ UnpMetaVar (i,prec)
+ | ETPattern _ ->
+ UnpBinderMetaVar (i,prec)
+ | ETName ->
+ UnpBinderMetaVar (i,Prec 0)
+ | ETBinder isopen ->
+ assert false
+ | ETOther _ -> failwith "TODO"
+
(* Heuristics for building default printing rules *)
let index_id id l = List.index Id.equal id l
let make_hunks etyps symbols from =
let vars,typs = List.split etyps in
- let rec make = function
+ let rec make b = function
| NonTerminal m :: prods ->
let i = index_id m vars in
- let _,prec = precedence_of_entry_type from (List.nth typs (i-1)) in
- let u = UnpMetaVar (i,prec) in
- if is_next_non_terminal prods then
- (None,u) :: add_break_if_none 1 (make prods)
+ let u = unparsing_metavar i from typs in
+ if is_next_non_terminal b prods then
+ (None, u) :: add_break_if_none 1 b (make b prods)
else
- (None,u) :: make_with_space prods
- | Terminal s :: prods when List.exists is_non_terminal prods ->
+ (None, u) :: make_with_space b prods
+ | Terminal s :: prods
+ when (* true to simulate presence of non-terminal *) b || List.exists is_non_terminal prods ->
if (is_comma s || is_operator s) then
(* Always a breakable space after comma or separator *)
- (None,UnpTerminal s) :: add_break_if_none 1 (make prods)
+ (None, UnpTerminal s) :: add_break_if_none 1 b (make b prods)
else if is_right_bracket s && is_next_terminal prods then
(* Always no space after right bracked, but possibly a break *)
- (None,UnpTerminal s) :: add_break_if_none 0 (make prods)
- else if is_left_bracket s && is_next_non_terminal prods then
- (None,UnpTerminal s) :: make prods
+ (None, UnpTerminal s) :: add_break_if_none 0 b (make b prods)
+ else if is_left_bracket s && is_next_non_terminal b prods then
+ (None, UnpTerminal s) :: make b prods
else if not (is_next_break prods) then
(* Add rigid space, no break, unless user asked for something *)
- (None,UnpTerminal (s^" ")) :: make prods
+ (None, UnpTerminal (s^" ")) :: make b prods
else
(* Rely on user spaces *)
- (None,UnpTerminal s) :: make prods
+ (None, UnpTerminal s) :: make b prods
| Terminal s :: prods ->
(* Separate but do not cut a trailing sequence of terminal *)
(match prods with
- | Terminal _ :: _ -> (None,UnpTerminal (s^" ")) :: make prods
- | _ -> (None,UnpTerminal s) :: make prods)
+ | Terminal _ :: _ -> (None,UnpTerminal (s^" ")) :: make b prods
+ | _ -> (None,UnpTerminal s) :: make b prods)
| Break n :: prods ->
- add_break n (make prods)
+ add_break n (make b prods)
| SProdList (m,sl) :: prods ->
let i = index_id m vars in
@@ -451,47 +422,52 @@ let make_hunks etyps symbols from =
(* If no separator: add a break *)
if List.is_empty sl then add_break 1 []
(* We add NonTerminal for simulation but remove it afterwards *)
- else snd (List.sep_last (make (sl@[NonTerminal m]))) in
+ else make true sl in
let hunk = match typ with
| ETConstr _ -> UnpListMetaVar (i,prec,List.map snd sl')
| ETBinder isopen ->
check_open_binder isopen sl m;
UnpBinderListMetaVar (i,isopen,List.map snd sl')
| _ -> assert false in
- (None,hunk) :: make_with_space prods
+ (None, hunk) :: make_with_space b prods
| [] -> []
- and make_with_space prods =
+ and make_with_space b prods =
match prods with
| Terminal s' :: prods'->
if is_operator s' then
(* A rigid space before operator and a breakable after *)
- (None,UnpTerminal (" "^s')) :: add_break_if_none 1 (make prods')
+ (None,UnpTerminal (" "^s')) :: add_break_if_none 1 b (make b prods')
else if is_comma s' then
(* No space whatsoever before comma *)
- make prods
+ make b prods
else if is_right_bracket s' then
- make prods
+ make b prods
else
(* A breakable space between any other two terminals *)
- add_break_if_none 1 (make prods)
+ add_break_if_none 1 b (make b prods)
| (NonTerminal _ | SProdList _) :: _ ->
(* A breakable space before a non-terminal *)
- add_break_if_none 1 (make prods)
+ add_break_if_none 1 b (make b prods)
| Break _ :: _ ->
(* Rely on user wish *)
- make prods
+ make b prods
| [] -> []
- in make symbols
+ in make false symbols
(* Build default printing rules from explicit format *)
let error_format ?loc () = user_err ?loc Pp.(str "The format does not match the notation.")
+let warn_format_break =
+ CWarnings.create ~name:"notation-both-format-and-spaces" ~category:"parsing"
+ (fun () ->
+ strbrk "Discarding format implicitly indicated by multiple spaces in notation because an explicit format modifier is given.")
+
let rec split_format_at_ldots hd = function
- | (loc,UnpTerminal s) :: fmt when String.equal s (Id.to_string ldots_var) -> loc, List.rev hd, fmt
+ | (loc,UnpTerminal s) :: fmt when String.equal s (Id.to_string Notation_ops.ldots_var) -> loc, List.rev hd, fmt
| u :: fmt ->
check_no_ldots_in_box u;
split_format_at_ldots (u::hd) fmt
@@ -538,8 +514,7 @@ let hunks_of_format (from,(vars,typs)) symfmt =
let symbs, l = aux (symbs,fmt) in symbs, UnpTerminal s :: l
| NonTerminal s :: symbs, (_,UnpTerminal s') :: fmt when Id.equal s (Id.of_string s') ->
let i = index_id s vars in
- let _,prec = precedence_of_entry_type from (List.nth typs (i-1)) in
- let symbs, l = aux (symbs,fmt) in symbs, UnpMetaVar (i,prec) :: l
+ let symbs, l = aux (symbs,fmt) in symbs, unparsing_metavar i from typs :: l
| symbs, (_,UnpBox (a,b)) :: fmt ->
let symbs', b' = aux (symbs,b) in
let symbs', l = aux (symbs',fmt) in
@@ -562,6 +537,7 @@ let hunks_of_format (from,(vars,typs)) symfmt =
| _ -> assert false in
symbs, hunk :: l
| symbs, [] -> symbs, []
+ | Break _ :: symbs, fmt -> warn_format_break (); aux (symbs,fmt)
| _, fmt -> error_format ?loc:(fst (List.hd fmt)) ()
in
match aux symfmt with
@@ -574,8 +550,8 @@ let hunks_of_format (from,(vars,typs)) symfmt =
let assoc_of_type n (_,typ) = precedence_of_entry_type n typ
let is_not_small_constr = function
- ETConstr _ -> true
- | ETOther("constr","binder_constr") -> true
+ ETProdConstr _ -> true
+ | ETProdOther("constr","binder_constr") -> true
| _ -> false
let rec define_keywords_aux = function
@@ -606,15 +582,15 @@ let distribute a ll = List.map (fun l -> a @ l) ll
t;sep;t;...;t;sep;t;...;t;sep;t;LIST1(t,sep) *)
let expand_list_rule typ tkl x n p ll =
- let camlp4_message_name = Some (add_suffix x ("_"^string_of_int n)) in
- let main = GramConstrNonTerminal (ETConstr typ, camlp4_message_name) in
+ let camlp5_message_name = Some (add_suffix x ("_"^string_of_int n)) in
+ let main = GramConstrNonTerminal (ETProdConstr typ, camlp5_message_name) in
let tks = List.map (fun x -> GramConstrTerminal x) tkl in
let rec aux i hds ll =
if i < p then aux (i+1) (main :: tks @ hds) ll
else if Int.equal i (p+n) then
let hds =
GramConstrListMark (p+n,true,p) :: hds
- @ [GramConstrNonTerminal (ETConstrList (typ,tkl), Some x)] in
+ @ [GramConstrNonTerminal (ETProdConstrList (typ,tkl), Some x)] in
distribute hds ll
else
distribute (GramConstrListMark (i+1,false,p) :: hds @ [main]) ll @
@@ -623,7 +599,7 @@ let expand_list_rule typ tkl x n p ll =
let is_constr_typ typ x etyps =
match List.assoc x etyps with
- | ETConstr typ' -> typ = typ'
+ | ETConstr typ' | ETConstrAsBinder (_,typ') -> typ = typ'
| _ -> false
let include_possible_similar_trailing_pattern typ etyps sl l =
@@ -636,12 +612,21 @@ let include_possible_similar_trailing_pattern typ etyps sl l =
with Exit -> n,l in
try_aux 0 l
+let prod_entry_type = function
+ | ETName -> ETProdName
+ | ETReference -> ETProdReference
+ | ETBigint -> ETProdBigint
+ | ETBinder _ -> assert false (* See check_binder_type *)
+ | ETConstr p | ETConstrAsBinder (_,p) -> ETProdConstr p
+ | ETPattern (_,n) -> ETProdPattern (match n with None -> 0 | Some n -> n)
+ | ETOther (s,t) -> ETProdOther (s,t)
+
let make_production etyps symbols =
let rec aux = function
| [] -> [[]]
| NonTerminal m :: l ->
let typ = List.assoc m etyps in
- distribute [GramConstrNonTerminal (typ, Some m)] (aux l)
+ distribute [GramConstrNonTerminal (prod_entry_type typ, Some m)] (aux l)
| Terminal s :: l ->
distribute [GramConstrTerminal (CLexer.terminal s)] (aux l)
| Break _ :: l ->
@@ -656,8 +641,10 @@ let make_production etyps symbols =
let p,l' = include_possible_similar_trailing_pattern typ etyps sl l in
expand_list_rule typ tkl x 1 p (aux l')
| ETBinder o ->
- distribute
- [GramConstrNonTerminal (ETBinderList (o,tkl), Some x)] (aux l)
+ check_open_binder o sl x;
+ let typ = if o then (assert (tkl = []); ETBinderOpen) else ETBinderClosed tkl in
+ distribute
+ [GramConstrNonTerminal (ETProdBinderList typ, Some x)] (aux l)
| _ ->
user_err Pp.(str "Components of recursive patterns in notation must be terms or binders.") in
let prods = aux symbols in
@@ -675,6 +662,7 @@ let rec find_symbols c_current c_next c_last = function
let border = function
| (_,ETConstr(_,BorderProd (_,a))) :: _ -> a
+ | (_,(ETConstrAsBinder(_,(_,BorderProd (_,a))))) :: _ -> a
| _ -> None
let recompute_assoc typs =
@@ -688,17 +676,16 @@ let recompute_assoc typs =
(* Registration of syntax extensions (parsing/printing, no interpretation)*)
let pr_arg_level from (lev,typ) =
- let pplev = match lev with
+ let pplev = function
| (n,L) when Int.equal n from -> str "at next level"
| (n,E) -> str "at level " ++ int n
| (n,L) -> str "at level below " ++ int n
| (n,Prec m) when Int.equal m n -> str "at level " ++ int n
| (n,_) -> str "Unknown level" in
- let pptyp = match typ with
- | NtnInternTypeConstr -> mt ()
- | NtnInternTypeBinder -> str " " ++ surround (str "binder")
- | NtnInternTypeIdent -> str " " ++ surround (str "ident") in
- pplev ++ pptyp
+ Ppvernac.pr_set_entry_type (fun _ -> (*TO CHECK*) mt()) typ ++
+ (match typ with
+ | ETConstr _ | ETConstrAsBinder _ | ETPattern _ -> spc () ++ pplev lev
+ | _ -> mt ())
let pr_level ntn (from,args,typs) =
str "at level " ++ int from ++ spc () ++ str "with arguments" ++ spc() ++
@@ -830,15 +817,23 @@ let interp_modifiers modl = let open NotationMods in
interp { acc with etyps = (id,typ) :: acc.etyps; } l
| SetItemLevel ([],n) :: l ->
interp acc l
+ | SetItemLevelAsBinder ([],_,_) :: l ->
+ interp acc l
| SetItemLevel (s::idl,n) :: l ->
let id = Id.of_string s in
if Id.List.mem_assoc id acc.etyps then
user_err ~hdr:"Metasyntax.interp_modifiers"
(str s ++ str " is already assigned to an entry or constr level.");
- let typ = ETConstr (n,()) in
+ let typ = ETConstr (Some n) in
interp { acc with etyps = (id,typ)::acc.etyps; } (SetItemLevel (idl,n)::l)
+ | SetItemLevelAsBinder (s::idl,bk,n) :: l ->
+ let id = Id.of_string s in
+ if Id.List.mem_assoc id acc.etyps then
+ user_err ~hdr:"Metasyntax.interp_modifiers"
+ (str s ++ str " is already assigned to an entry or constr level.");
+ let typ = ETConstrAsBinder (bk,n) in
+ interp { acc with etyps = (id,typ)::acc.etyps; } (SetItemLevelAsBinder (idl,bk,n)::l)
| SetLevel n :: l ->
-
interp { acc with level = Some n; } l
| SetAssoc a :: l ->
if not (Option.is_empty acc.assoc) then user_err Pp.(str "An associativity is given more than once.");
@@ -902,12 +897,17 @@ let get_compat_version mods =
let set_entry_type etyps (x,typ) =
let typ = try
match List.assoc x etyps, typ with
- | ETConstr (n,()), (_,BorderProd (left,_)) ->
+ | ETConstr (Some n), (_,BorderProd (left,_)) ->
ETConstr (n,BorderProd (left,None))
- | ETConstr (n,()), (_,InternalProd) -> ETConstr (n,InternalProd)
- | (ETPattern | ETName | ETBigint | ETOther _ |
- ETReference | ETBinder _ as t), _ -> t
- | (ETBinderList _ |ETConstrList _), _ -> assert false
+ | ETConstr (Some n), (_,InternalProd) -> ETConstr (n,InternalProd)
+ | ETConstrAsBinder (bk, Some n), (_,BorderProd (left,_)) ->
+ ETConstrAsBinder (bk, (n,BorderProd (left,None)))
+ | ETConstrAsBinder (bk, Some n), (_,InternalProd) ->
+ ETConstrAsBinder (bk, (n,InternalProd))
+ | ETPattern (b,n), _ -> ETPattern (b,n)
+ | (ETName | ETBigint | ETReference | ETBinder _ | ETOther _ as x), _ -> x
+ | ETConstr None, _ -> ETConstr typ
+ | ETConstrAsBinder (bk,None), _ -> ETConstrAsBinder (bk,typ)
with Not_found -> ETConstr typ
in (x,typ)
@@ -927,12 +927,9 @@ let join_auxiliary_recursive_types recvars etyps =
recvars etyps
let internalization_type_of_entry_type = function
- | ETConstr _ -> NtnInternTypeConstr
- | ETBigint | ETReference -> NtnInternTypeConstr
- | ETBinder _ -> NtnInternTypeBinder
- | ETName -> NtnInternTypeIdent
- | ETPattern | ETOther _ -> user_err Pp.(str "Not supported.")
- | ETBinderList _ | ETConstrList _ -> assert false
+ | ETBinder _ -> NtnInternTypeOnlyBinder
+ | ETConstr _ | ETConstrAsBinder _ | ETBigint | ETReference
+ | ETName | ETPattern _ | ETOther _ -> NtnInternTypeAny
let set_internalization_type typs =
List.map (fun (_, e) -> internalization_type_of_entry_type e) typs
@@ -943,28 +940,36 @@ let make_internalization_vars recvars mainvars typs =
maintyps @ extratyps
let make_interpretation_type isrec isonlybinding = function
- | NtnInternTypeConstr when isrec -> NtnTypeConstrList
- | NtnInternTypeConstr | NtnInternTypeIdent ->
- if isonlybinding then NtnTypeOnlyBinder else NtnTypeConstr
- | NtnInternTypeBinder when isrec -> NtnTypeBinderList
- | NtnInternTypeBinder -> user_err Pp.(str "Type binder is only for use in recursive notations for binders.")
-
-let make_interpretation_vars recvars allvars =
+ | ETConstr _ ->
+ if isrec then NtnTypeConstrList else
+ if isonlybinding then
+ (* Parsed as constr, but interpreted as a binder: default is to parse it as an ident only *)
+ NtnTypeBinder (NtnBinderParsedAsConstr AsIdent)
+ else NtnTypeConstr
+ | ETConstrAsBinder (bk,_) -> NtnTypeBinder (NtnBinderParsedAsConstr bk)
+ | ETName -> NtnTypeBinder NtnParsedAsIdent
+ | ETPattern (ppstrict,_) -> NtnTypeBinder (NtnParsedAsPattern ppstrict) (* Parsed as ident/pattern, primarily interpreted as binder; maybe strict at printing *)
+ | ETBigint | ETReference | ETOther _ -> NtnTypeConstr
+ | ETBinder _ ->
+ if isrec then NtnTypeBinderList
+ else anomaly Pp.(str "Type binder is only for use in recursive notations for binders.")
+
+let make_interpretation_vars recvars allvars typs =
let eq_subscope (sc1, l1) (sc2, l2) =
Option.equal String.equal sc1 sc2 &&
List.equal String.equal l1 l2
in
let check (x, y) =
- let (_,scope1, _) = Id.Map.find x allvars in
- let (_,scope2, _) = Id.Map.find y allvars in
+ let (_,scope1) = Id.Map.find x allvars in
+ let (_,scope2) = Id.Map.find y allvars in
if not (eq_subscope scope1 scope2) then error_not_same_scope x y
in
let () = List.iter check recvars in
let useless_recvars = List.map snd recvars in
let mainvars =
Id.Map.filter (fun x _ -> not (Id.List.mem x useless_recvars)) allvars in
- Id.Map.mapi (fun x (isonlybinding, sc, typ) ->
- (sc, make_interpretation_type (Id.List.mem_assoc x recvars) isonlybinding typ)) mainvars
+ Id.Map.mapi (fun x (isonlybinding, sc) ->
+ (sc, make_interpretation_type (Id.List.mem_assoc x recvars) isonlybinding (Id.List.assoc x typs))) mainvars
let check_rule_productivity l =
if List.for_all (function NonTerminal _ | Break _ -> true | _ -> false) l then
@@ -979,18 +984,27 @@ let warn_notation_bound_to_variable =
let warn_non_reversible_notation =
CWarnings.create ~name:"non-reversible-notation" ~category:"parsing"
- (fun () ->
- strbrk "This notation will not be used for printing as it is not reversible.")
-
-let is_not_printable onlyparse nonreversible = function
+ (function
+ | APrioriReversible -> assert false
+ | HasLtac ->
+ strbrk "This notation contains Ltac expressions: it will not be used for printing."
+ | NonInjective ids ->
+ let n = List.length ids in
+ strbrk (String.plural n "Variable") ++ spc () ++ pr_enum Id.print ids ++ spc () ++
+ strbrk (if n > 1 then "do" else "does") ++
+ str " not occur in the right-hand side." ++ spc() ++
+ strbrk "The notation will not be used for printing as it is not reversible.")
+
+let is_not_printable onlyparse reversibility = function
| NVar _ ->
if not onlyparse then warn_notation_bound_to_variable ();
true
| _ ->
- if not onlyparse && nonreversible then
- (warn_non_reversible_notation (); true)
+ if not onlyparse && reversibility <> APrioriReversible then
+ (warn_non_reversible_notation reversibility; true)
else onlyparse
+
let find_precedence lev etyps symbols onlyprint =
let first_symbol =
let rec aux = function
@@ -1008,29 +1022,30 @@ let find_precedence lev etyps symbols onlyprint =
match first_symbol with
| None -> [],0
| Some (NonTerminal x) ->
+ let test () =
+ if onlyprint then
+ if Option.is_empty lev then
+ user_err Pp.(str "Explicit level needed in only-printing mode when the level of the leftmost non-terminal is given.")
+ else [],Option.get lev
+ else
+ user_err Pp.(str "The level of the leftmost non-terminal cannot be changed.") in
(try match List.assoc x etyps with
- | ETConstr _ ->
- if onlyprint then
- if Option.is_empty lev then
- user_err Pp.(str "Explicit level needed in only-printing mode when the level of the leftmost non-terminal is given.")
- else [],Option.get lev
- else
- user_err Pp.(str "The level of the leftmost non-terminal cannot be changed.")
- | ETName | ETBigint | ETReference ->
+ | ETConstr _ -> test ()
+ | ETConstrAsBinder (_,Some _) -> test ()
+ | (ETName | ETBigint | ETReference) ->
begin match lev with
| None ->
([Feedback.msg_info ?loc:None ,strbrk "Setting notation at level 0."],0)
| Some 0 ->
([],0)
| _ ->
- user_err Pp.(str "A notation starting with an atomic expression must be at level 0.")
+ user_err Pp.(str "A notation starting with an atomic expression must be at level 0.")
end
- | ETPattern | ETBinder _ | ETOther _ -> (* Give a default ? *)
- if Option.is_empty lev then
- user_err Pp.(str "Need an explicit level.")
- else [],Option.get lev
- | ETConstrList _ | ETBinderList _ ->
- assert false (* internally used in grammar only *)
+ | (ETPattern _ | ETBinder _ | ETOther _ | ETConstrAsBinder _) ->
+ (* Give a default ? *)
+ if Option.is_empty lev then
+ user_err Pp.(str "Need an explicit level.")
+ else [],Option.get lev
with Not_found ->
if Option.is_empty lev then
user_err Pp.(str "A left-recursive notation must have an explicit level.")
@@ -1074,7 +1089,7 @@ let remove_curly_brackets l =
module SynData = struct
- type subentry_types = (Id.t * (production_level, production_position) constr_entry_key_gen) list
+ type subentry_types = (Id.t * (production_level * production_position) constr_entry_key_gen) list
(* XXX: Document *)
type syn_data = {
@@ -1128,8 +1143,7 @@ let compute_syntax_data df modifiers =
let onlyparse = mods.only_parsing in
if onlyprint && onlyparse then user_err (str "A notation cannot be both 'only printing' and 'only parsing'.");
let assoc = Option.append mods.assoc (Some NonA) in
- let toks = split_notation_string df in
- let (recvars,mainvars,symbols) = analyze_notation_tokens ~onlyprint toks in
+ let (recvars,mainvars,symbols) = analyze_notation_tokens ~onlyprint df in
let _ = check_useless_entry_types recvars mainvars mods.etyps in
let _ = check_binder_type recvars mods.etyps in
@@ -1152,7 +1166,7 @@ let compute_syntax_data df modifiers =
let i_typs = set_internalization_type sy_typs in
let pa_sy_data = (sy_typs_for_grammar,symbols_for_grammar) in
let pp_sy_data = (sy_typs,symbols) in
- let sy_fulldata = (ntn_for_grammar,(n,prec_for_grammar,i_typs),need_squash) in
+ let sy_fulldata = (ntn_for_grammar,(n,prec_for_grammar,List.map snd sy_typs_for_grammar),need_squash) in
let df' = ((Lib.library_dp(),Lib.current_dirpath true),df) in
let i_data = ntn_for_interp, df' in
@@ -1171,7 +1185,7 @@ let compute_syntax_data df modifiers =
mainvars;
intern_typs = i_typs;
- level = (n,prec,i_typs);
+ level = (n,prec,List.map snd sy_typs);
pa_syntax_data = pa_sy_data;
pp_syntax_data = pp_sy_data;
not_data = sy_fulldata;
@@ -1329,10 +1343,10 @@ let add_notation_in_scope local df env c mods scope =
ninterp_var_type = to_map i_vars;
ninterp_rec_vars = to_map sd.recvars;
} in
- let (acvars, ac, reversible) = interp_notation_constr env nenv c in
- let interp = make_interpretation_vars sd.recvars acvars in
+ let (acvars, ac, reversibility) = interp_notation_constr env nenv c in
+ let interp = make_interpretation_vars sd.recvars acvars (fst sd.pa_syntax_data) in
let map (x, _) = try Some (x, Id.Map.find x interp) with Not_found -> None in
- let onlyparse = is_not_printable sd.only_parsing (not reversible) ac in
+ let onlyparse = is_not_printable sd.only_parsing reversibility ac in
let notation = {
notobj_local = local;
notobj_scope = scope;
@@ -1350,8 +1364,7 @@ let add_notation_in_scope local df env c mods scope =
sd.info
let add_notation_interpretation_core local df env ?(impls=empty_internalization_env) c scope onlyparse onlyprint compat =
- let dfs = split_notation_string df in
- let (recvars,mainvars,symbs) = analyze_notation_tokens ~onlyprint dfs in
+ let (recvars,mainvars,symbs) = analyze_notation_tokens ~onlyprint df in
(* Recover types of variables and pa/pp rules; redeclare them if needed *)
let i_typs, onlyprint = if not (is_numeral symbs) then begin
let sy = recover_notation_syntax (make_notation_key symbs) in
@@ -1363,15 +1376,15 @@ let add_notation_interpretation_core local df env ?(impls=empty_internalization_
(* Declare interpretation *)
let path = (Lib.library_dp(), Lib.current_dirpath true) in
let df' = (make_notation_key symbs, (path,df)) in
- let i_vars = make_internalization_vars recvars mainvars i_typs in
+ let i_vars = make_internalization_vars recvars mainvars (List.map internalization_type_of_entry_type i_typs) in
let nenv = {
ninterp_var_type = to_map i_vars;
ninterp_rec_vars = to_map recvars;
} in
- let (acvars, ac, reversible) = interp_notation_constr env ~impls nenv c in
- let interp = make_interpretation_vars recvars acvars in
+ let (acvars, ac, reversibility) = interp_notation_constr env ~impls nenv c in
+ let interp = make_interpretation_vars recvars acvars (List.combine mainvars i_typs) in
let map (x, _) = try Some (x, Id.Map.find x interp) with Not_found -> None in
- let onlyparse = is_not_printable onlyparse (not reversible) ac in
+ let onlyparse = is_not_printable onlyparse reversibility ac in
let notation = {
notobj_local = local;
notobj_scope = scope;
@@ -1427,8 +1440,7 @@ let add_notation local env c ((loc,df),modifiers) sc =
let add_notation_extra_printing_rule df k v =
let notk =
- let dfs = split_notation_string df in
- let _,_, symbs = analyze_notation_tokens ~onlyprint:true dfs in
+ let _,_, symbs = analyze_notation_tokens ~onlyprint:true df in
make_notation_key symbs in
Notation.add_notation_extra_printing_rule notk k v
@@ -1499,23 +1511,21 @@ let try_interp_name_alias = function
| _ -> raise Not_found
let add_syntactic_definition env ident (vars,c) local onlyparse =
- let nonprintable = ref false in
- let vars,pat =
- try [], NRef (try_interp_name_alias (vars,c))
+ let vars,reversibility,pat =
+ try [], APrioriReversible, NRef (try_interp_name_alias (vars,c))
with Not_found ->
- let fold accu id = Id.Map.add id NtnInternTypeConstr accu in
+ let fold accu id = Id.Map.add id NtnInternTypeAny accu in
let i_vars = List.fold_left fold Id.Map.empty vars in
let nenv = {
ninterp_var_type = i_vars;
ninterp_rec_vars = Id.Map.empty;
} in
- let nvars, pat, reversible = interp_notation_constr env nenv c in
- let () = nonprintable := not reversible in
- let map id = let (_,sc,_) = Id.Map.find id nvars in (id, sc) in
- List.map map vars, pat
+ let nvars, pat, reversibility = interp_notation_constr env nenv c in
+ let map id = let (_,sc) = Id.Map.find id nvars in (id, sc) in
+ List.map map vars, reversibility, pat
in
let onlyparse = match onlyparse with
- | None when (is_not_printable false !nonprintable pat) -> Some Flags.Current
+ | None when (is_not_printable false reversibility pat) -> Some Flags.Current
| p -> p
in
Syntax_def.declare_syntactic_definition local ident onlyparse (vars,pat)
diff --git a/vernac/metasyntax.mli b/vernac/metasyntax.mli
index b3049f1b7..9137f7a7e 100644
--- a/vernac/metasyntax.mli
+++ b/vernac/metasyntax.mli
@@ -51,14 +51,10 @@ val add_syntax_extension :
val add_syntactic_definition : env -> Id.t -> Id.t list * constr_expr ->
bool -> Flags.compat_version option -> unit
-(** Print the Camlp4 state of a grammar *)
+(** Print the Camlp5 state of a grammar *)
val pr_grammar : string -> Pp.t
-type any_entry = AnyEntry : 'a Pcoq.Gram.entry -> any_entry
-
-val register_grammar : string -> any_entry list -> unit
-
val check_infix_modifiers : syntax_modifier list -> unit
val with_syntax_protection : ('a -> 'b) -> 'a -> 'b
diff --git a/vernac/mltop.ml b/vernac/mltop.ml
index d3de10235..053b9d070 100644
--- a/vernac/mltop.ml
+++ b/vernac/mltop.ml
@@ -184,10 +184,28 @@ let warn_cannot_open_path =
type add_ml = AddNoML | AddTopML | AddRecML
-let add_rec_path add_ml ~unix_path ~coq_root ~implicit =
+type vo_path_spec = {
+ unix_path : string;
+ coq_path : Names.DirPath.t;
+ implicit : bool;
+ has_ml : add_ml;
+}
+
+type coq_path_spec =
+ | VoPath of vo_path_spec
+ | MlPath of string
+
+type coq_path = {
+ path_spec: coq_path_spec;
+ recursive: bool;
+}
+
+let add_vo_path ~recursive lp =
+ let unix_path = lp.unix_path in
+ let implicit = lp.implicit in
if exists_dir unix_path then
- let dirs = all_subdirs ~unix_path in
- let prefix = Names.DirPath.repr coq_root in
+ let dirs = if recursive then all_subdirs ~unix_path else [] in
+ let prefix = Names.DirPath.repr lp.coq_path in
let convert_dirs (lp, cp) =
try
let path = List.rev_map convert_string cp @ prefix in
@@ -195,17 +213,23 @@ let add_rec_path add_ml ~unix_path ~coq_root ~implicit =
with Exit -> None
in
let dirs = List.map_filter convert_dirs dirs in
- let () = match add_ml with
+ let () = match lp.has_ml with
| AddNoML -> ()
| AddTopML -> add_ml_dir unix_path
| AddRecML -> List.iter (fun (lp,_) -> add_ml_dir lp) dirs in
let add (path, dir) =
Loadpath.add_load_path path ~implicit dir in
let () = List.iter add dirs in
- Loadpath.add_load_path unix_path ~implicit coq_root
+ Loadpath.add_load_path unix_path ~implicit lp.coq_path
else
warn_cannot_open_path unix_path
+let add_coq_path { recursive; path_spec } = match path_spec with
+ | VoPath lp ->
+ add_vo_path ~recursive lp
+ | MlPath dir ->
+ if recursive then add_rec_ml_dir dir else add_ml_dir dir
+
(* convertit un nom quelconque en nom de fichier ou de module *)
let mod_of_name name =
if Filename.check_suffix name ".cmo" then
@@ -378,7 +402,7 @@ let unfreeze_ml_modules x =
(fun (name,path) -> trigger_ml_object false false false ?path name) x
let _ =
- Summary.declare_summary Summary.ml_modules
+ Summary.declare_ml_modules_summary
{ Summary.freeze_function = (fun _ -> get_loaded_modules ());
Summary.unfreeze_function = unfreeze_ml_modules;
Summary.init_function = reset_loaded_modules }
diff --git a/vernac/mltop.mli b/vernac/mltop.mli
index 324a66d38..e44a7c243 100644
--- a/vernac/mltop.mli
+++ b/vernac/mltop.mli
@@ -42,14 +42,26 @@ val dir_ml_load : string -> unit
(** Dynamic interpretation of .ml *)
val dir_ml_use : string -> unit
-(** Adds a path to the ML paths *)
-val add_ml_dir : string -> unit
-val add_rec_ml_dir : string -> unit
-
+(** Adds a path to the Coq and ML paths *)
type add_ml = AddNoML | AddTopML | AddRecML
-(** Adds a path to the Coq and ML paths *)
-val add_rec_path : add_ml -> unix_path:string -> coq_root:Names.DirPath.t -> implicit:bool -> unit
+type vo_path_spec = {
+ unix_path : string; (* Filesystem path contaning vo/ml files *)
+ coq_path : Names.DirPath.t; (* Coq prefix for the path *)
+ implicit : bool; (* [implicit = true] avoids having to qualify with [coq_path] *)
+ has_ml : add_ml; (* If [has_ml] is true, the directory will also be search for plugins *)
+}
+
+type coq_path_spec =
+ | VoPath of vo_path_spec
+ | MlPath of string
+
+type coq_path = {
+ path_spec: coq_path_spec;
+ recursive: bool;
+}
+
+val add_coq_path : coq_path -> unit
(** List of modules linked to the toplevel *)
val add_known_module : string -> unit
diff --git a/vernac/obligations.ml b/vernac/obligations.ml
index 1046d68f8..58e4b00fc 100644
--- a/vernac/obligations.ml
+++ b/vernac/obligations.ml
@@ -429,8 +429,8 @@ let map_replace k v m = ProgMap.add k (CEphemeron.create v) (ProgMap.remove k m)
let map_keys m = ProgMap.fold (fun k _ l -> k :: l) m []
-let from_prg : program_info ProgMap.t ref =
- Summary.ref ProgMap.empty ~name:"program-tcc-table"
+let from_prg, program_tcc_summary_tag =
+ Summary.ref_tag ProgMap.empty ~name:"program-tcc-table"
let close sec =
if not (ProgMap.is_empty !from_prg) then
@@ -477,7 +477,10 @@ let declare_definition prg =
let fix_exn = Hook.get get_fix_exn () in
let typ = nf typ in
let body = nf body in
- let uvars = Univ.LSet.union (Univops.universes_of_constr typ) (Univops.universes_of_constr body) in
+ let env = Global.env () in
+ let uvars = Univ.LSet.union
+ (Univops.universes_of_constr env typ)
+ (Univops.universes_of_constr env body) in
let uctx = UState.restrict prg.prg_ctx uvars in
let univs = UState.check_univ_decl ~poly:(pi2 prg.prg_kind) uctx prg.prg_univdecl in
let ce = definition_entry ~fix_exn ~opaque ~types:typ ~univs body in
@@ -833,7 +836,7 @@ let obligation_terminator name num guard hook auto pf =
let entry = Safe_typing.inline_private_constants_in_definition_entry env entry in
let ty = entry.Entries.const_entry_type in
let (body, cstr), () = Future.force entry.Entries.const_entry_body in
- let sigma = Evd.from_ctx (fst uctx) in
+ let sigma = Evd.from_ctx uctx in
let sigma = Evd.merge_context_set ~sideff:true Evd.univ_rigid sigma cstr in
Inductiveops.control_only_guard (Global.env ()) body;
(** Declare the obligation ourselves and drop the hook *)
@@ -1182,7 +1185,6 @@ let init_program () =
Coqlib.check_required_library ["Coq";"Init";"Specif"];
Coqlib.check_required_library ["Coq";"Program";"Tactics"]
-
let set_program_mode c =
if c then
if !Flags.program_mode then ()
diff --git a/vernac/obligations.mli b/vernac/obligations.mli
index 0602e52e9..bdc97d48c 100644
--- a/vernac/obligations.mli
+++ b/vernac/obligations.mli
@@ -104,3 +104,6 @@ exception NoObligations of Names.Id.t option
val explain_no_obligations : Names.Id.t option -> Pp.t
val set_program_mode : bool -> unit
+
+type program_info
+val program_tcc_summary_tag : program_info Id.Map.t Summary.Dyn.tag
diff --git a/vernac/record.ml b/vernac/record.ml
index 1d255b08e..16b95a5ef 100644
--- a/vernac/record.ml
+++ b/vernac/record.ml
@@ -60,23 +60,25 @@ let _ =
optread = (fun () -> !typeclasses_unique);
optwrite = (fun b -> typeclasses_unique := b); }
-let interp_fields_evars env evars impls_env nots l =
+let interp_fields_evars env sigma impls_env nots l =
List.fold_left2
- (fun (env, uimpls, params, impls) no ((loc, i), b, t) ->
- let t', impl = interp_type_evars_impls env evars ~impls t in
- let b' = Option.map (fun x -> fst (interp_casted_constr_evars_impls env evars ~impls x t')) b in
+ (fun (env, sigma, uimpls, params, impls) no ((loc, i), b, t) ->
+ let sigma, (t', impl) = interp_type_evars_impls env sigma ~impls t in
+ let sigma, b' =
+ Option.cata (fun x -> on_snd (fun x -> Some (fst x)) @@
+ interp_casted_constr_evars_impls env sigma ~impls x t') (sigma,None) b in
let impls =
match i with
| Anonymous -> impls
- | Name id -> Id.Map.add id (compute_internalization_data env Constrintern.Method (EConstr.to_constr !evars t') impl) impls
+ | Name id -> Id.Map.add id (compute_internalization_data env Constrintern.Method (EConstr.to_constr sigma t') impl) impls
in
let d = match b' with
| None -> LocalAssum (i,t')
| Some b' -> LocalDef (i,b',t')
in
List.iter (Metasyntax.set_notation_for_interpretation env impls) no;
- (EConstr.push_rel d env, impl :: uimpls, d::params, impls))
- (env, [], [], impls_env) nots l
+ (EConstr.push_rel d env, sigma, impl :: uimpls, d::params, impls))
+ (env, sigma, [], [], impls_env) nots l
let compute_constructor_level evars env l =
List.fold_right (fun d (env, univ) ->
@@ -97,10 +99,9 @@ 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 evd, decl = Univdecls.interp_univ_decl_opt env0 pl in
- let evars = ref evd in
- let _ =
- let error bk (loc, name) =
+ let sigma, decl = Univdecls.interp_univ_decl_opt env0 pl in
+ let _ =
+ let error bk (loc, name) =
match bk, name with
| Default _, Anonymous ->
user_err ?loc ~hdr:"record" (str "Record parameters must be named")
@@ -112,63 +113,65 @@ let typecheck_params_and_fields finite def id poly pl t ps nots fs =
| CLocalPattern (loc,(_,_)) ->
Loc.raise ?loc (Stream.Error "pattern with quote not allowed in record parameters.")) ps
in
- let impls_env, ((env1,newps), imps) = interp_context_evars env0 evars ps in
- let typ, sort, template = match t with
+ let sigma, (impls_env, ((env1,newps), imps)) = interp_context_evars env0 sigma ps in
+ let sigma, typ, sort, template = match t with
| Some t ->
let env = EConstr.push_rel_context newps env0 in
let poly =
match t with
| { CAst.v = CSort (Misctypes.GType []) } -> true | _ -> false in
- let s = interp_type_evars env evars ~impls:empty_internalization_env t in
- let sred = Reductionops.whd_all env !evars s in
- (match EConstr.kind !evars sred with
+ let sigma, s = interp_type_evars env sigma ~impls:empty_internalization_env t in
+ let sred = Reductionops.whd_all env sigma s in
+ (match EConstr.kind sigma sred with
| Sort s' ->
- let s' = EConstr.ESorts.kind !evars s' in
+ let s' = EConstr.ESorts.kind sigma s' in
(if poly then
- match Evd.is_sort_variable !evars s' with
- | Some l -> evars := Evd.make_flexible_variable !evars ~algebraic:true l;
- s, s', true
- | None -> s, s', false
- else s, s', false)
+ match Evd.is_sort_variable sigma s' with
+ | Some l ->
+ let sigma = Evd.make_flexible_variable sigma ~algebraic:true l in
+ sigma, s, s', true
+ | None ->
+ sigma, s, s', false
+ else sigma, s, s', false)
| _ -> user_err ?loc:(constr_loc t) (str"Sort expected."))
| None ->
let uvarkind = Evd.univ_flexible_alg in
- let s = Evarutil.evd_comb0 (Evd.new_sort_variable uvarkind) evars in
- EConstr.mkSort s, s, true
+ let sigma, s = Evd.new_sort_variable uvarkind sigma in
+ sigma, EConstr.mkSort s, s, true
in
let arity = EConstr.it_mkProd_or_LetIn typ newps in
let env_ar = EConstr.push_rel_context newps (EConstr.push_rel (LocalAssum (Name id,arity)) env0) in
let assums = List.filter is_local_assum newps in
let params = List.map (RelDecl.get_name %> Name.get_id) assums in
- let ty = Inductive (params,(finite != BiFinite)) in
- let impls_env = compute_internalization_env env0 ~impls:impls_env ty [id] [EConstr.to_constr !evars arity] [imps] in
- let env2,impls,newfs,data =
- interp_fields_evars env_ar evars impls_env nots (binders_of_decls fs)
+ let ty = Inductive (params,(finite != Declarations.BiFinite)) in
+ let impls_env = compute_internalization_env env0 ~impls:impls_env ty [id] [EConstr.to_constr sigma arity] [imps] in
+ let env2,sigma,impls,newfs,data =
+ interp_fields_evars env_ar sigma impls_env nots (binders_of_decls fs)
in
- let evars =
- Pretyping.solve_remaining_evars Pretyping.all_and_fail_flags env_ar !evars Evd.empty in
- let typ, evars =
- let _, univ = compute_constructor_level evars env_ar newfs in
+ let sigma =
+ Pretyping.solve_remaining_evars Pretyping.all_and_fail_flags env_ar sigma Evd.empty in
+ let sigma, typ =
+ let _, univ = compute_constructor_level sigma env_ar newfs in
if not def && (Sorts.is_prop sort ||
(Sorts.is_set sort && is_impredicative_set env0)) then
- typ, evars
+ sigma, typ
else
- let evars = Evd.set_leq_sort env_ar evars (Type univ) sort in
+ let sigma = Evd.set_leq_sort env_ar sigma (Type univ) sort in
if Univ.is_small_univ univ &&
- Option.cata (Evd.is_flexible_level evars) false (Evd.is_sort_variable evars sort) then
+ Option.cata (Evd.is_flexible_level sigma) false (Evd.is_sort_variable sigma sort) then
(* We can assume that the level in aritysort is not constrained
and clear it, if it is flexible *)
- EConstr.mkSort (Sorts.sort_of_univ univ),
- Evd.set_eq_sort env_ar evars (Prop Pos) sort
- else typ, evars
+ Evd.set_eq_sort env_ar sigma (Prop Pos) sort,
+ EConstr.mkSort (Sorts.sort_of_univ univ)
+ else sigma, typ
in
- let evars, nf = Evarutil.nf_evars_and_universes evars in
- let newfs = List.map (EConstr.to_rel_decl evars) newfs in
- let newps = List.map (EConstr.to_rel_decl evars) newps in
- let typ = EConstr.to_constr evars typ in
- let ce t = Pretyping.check_evars env0 Evd.empty evars (EConstr.of_constr t) in
- let univs = Evd.check_univ_decl ~poly evars decl in
- let ubinders = Evd.universe_binders evars in
+ let sigma, _ = Evarutil.nf_evars_and_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 univs = Evd.check_univ_decl ~poly sigma decl in
+ let ubinders = Evd.universe_binders sigma in
List.iter (iter_constr ce) (List.rev newps);
List.iter (iter_constr ce) (List.rev newfs);
ubinders, univs, typ, template, imps, newps, impls, newfs
@@ -366,26 +369,10 @@ let declare_projections indsp ctx ?(kind=StructureComponent) binder_name coers u
(List.length fields,0,[],[],[]) coers (List.rev fields) (List.rev fieldimpls)
in (kinds,sp_projs)
-let structure_signature ctx =
- let rec deps_to_evar evm l =
- match l with [] -> Evd.empty
- | [decl] ->
- let env = Environ.empty_named_context_val in
- let (evm, _) = Evarutil.new_pure_evar env evm (EConstr.of_constr (RelDecl.get_type decl)) in
- evm
- | decl::tl ->
- let env = Environ.empty_named_context_val in
- let (evm, ev) = Evarutil.new_pure_evar env evm (EConstr.of_constr (RelDecl.get_type decl)) in
- let new_tl = Util.List.map_i
- (fun pos decl ->
- RelDecl.map_type (fun t -> EConstr.Unsafe.to_constr (Termops.replace_term evm (EConstr.mkRel pos) (EConstr.mkEvar(ev,[||])) (EConstr.of_constr t))) decl) 1 tl in
- deps_to_evar evm new_tl in
- deps_to_evar Evd.empty (List.rev ctx)
-
open Typeclasses
let declare_structure finite ubinders univs id idbuild paramimpls params arity template
- fieldimpls fields ?(kind=StructureComponent) ?name is_coe coers sign =
+ fieldimpls fields ?(kind=StructureComponent) ?name is_coe coers =
let nparams = List.length params and nfields = List.length fields in
let args = Context.Rel.to_extended_list mkRel nfields params in
let ind = applist (mkRel (1+nparams+nfields), args) in
@@ -420,17 +407,8 @@ let declare_structure finite ubinders univs id idbuild paramimpls params arity t
mind_entry_universes = univs;
}
in
- let mie =
- match ctx with
- | Polymorphic_const_entry ctx ->
- let env = Global.env () in
- let env' = Environ.push_context ctx env in
- let evd = Evd.from_env env' in
- Inductiveops.infer_inductive_subtyping env' evd mie
- | Monomorphic_const_entry _ ->
- mie
- in
- let kn = Command.declare_mutual_inductive_with_eliminations mie ubinders [(paramimpls,[])] in
+ let mie = InferCumulativity.infer_inductive (Global.env ()) mie in
+ let kn = ComInductive.declare_mutual_inductive_with_eliminations mie ubinders [(paramimpls,[])] in
let rsp = (kn,0) in (* This is ind path of idstruc *)
let cstr = (rsp,1) in
let kinds,sp_projs = declare_projections rsp ctx ~kind binder_name coers ubinders fieldimpls fields in
@@ -450,7 +428,7 @@ let implicits_of_context ctx =
1 (List.rev (Anonymous :: (List.map RelDecl.get_name ctx)))
let declare_class finite def cum ubinders univs id idbuild paramimpls params arity
- template fieldimpls fields ?(kind=StructureComponent) is_coe coers priorities sign =
+ template fieldimpls fields ?(kind=StructureComponent) is_coe coers priorities =
let fieldimpls =
(* Make the class implicit in the projections, and the params if applicable. *)
let len = List.length params in
@@ -498,15 +476,15 @@ let declare_class finite def cum ubinders univs id idbuild paramimpls params ari
match univs with
| Polymorphic_const_entry univs ->
if cum then
- Cumulative_ind_entry (Universes.univ_inf_ind_from_universe_context univs)
+ Cumulative_ind_entry (Univ.CumulativityInfo.from_universe_context univs)
else
Polymorphic_ind_entry univs
| Monomorphic_const_entry univs ->
Monomorphic_ind_entry univs
in
- let ind = declare_structure BiFinite ubinders univs (snd id) idbuild paramimpls
+ let ind = declare_structure Declarations.BiFinite ubinders univs (snd id) idbuild paramimpls
params arity template fieldimpls fields
- ~kind:Method ~name:binder_name false (List.map (fun _ -> false) fields) sign
+ ~kind:Method ~name:binder_name false (List.map (fun _ -> false) fields)
in
let coers = List.map2 (fun coe pri ->
Option.map (fun b ->
@@ -520,7 +498,7 @@ let declare_class finite def cum ubinders univs id idbuild paramimpls params ari
let ctx_context =
List.map (fun decl ->
match Typeclasses.class_of_constr Evd.empty (EConstr.of_constr (RelDecl.get_type decl)) with
- | Some (_, ((cl,_), _)) -> Some (cl.cl_impl, true)
+ | Some (_, ((cl,_), _)) -> Some cl.cl_impl
| None -> None)
params, params
in
@@ -528,6 +506,7 @@ let declare_class finite def cum ubinders univs id idbuild paramimpls params ari
match univs with
| Polymorphic_const_entry univs ->
let usubst, auctx = Univ.abstract_universes univs in
+ let usubst = Univ.make_instance_subst usubst in
let map c = Vars.subst_univs_level_constr usubst c in
let fields = Context.Rel.map map fields in
let ctx_context = on_snd (fun d -> Context.Rel.map map d) ctx_context in
@@ -612,12 +591,11 @@ let definition_structure (kind,cum,poly,finite,(is_coe,((loc,idstruc),pl)),ps,cf
let pl, univs, arity, template, implpars, params, implfs, fields =
States.with_state_protection (fun () ->
typecheck_params_and_fields finite (kind = Class true) idstruc poly pl s ps notations fs) () in
- let sign = structure_signature (fields@params) in
- match kind with
+ let gr = match kind with
| Class def ->
let priorities = List.map (fun id -> {hint_priority = id; hint_pattern = None}) priorities in
let gr = declare_class finite def cum pl univs (loc,idstruc) idbuild
- implpars params arity template implfs fields is_coe coers priorities sign in
+ implpars params arity template implfs fields is_coe coers priorities in
gr
| _ ->
let implfs = List.map
@@ -628,7 +606,7 @@ let definition_structure (kind,cum,poly,finite,(is_coe,((loc,idstruc),pl)),ps,cf
match univs with
| Polymorphic_const_entry univs ->
if cum then
- Cumulative_ind_entry (Universes.univ_inf_ind_from_universe_context univs)
+ Cumulative_ind_entry (Univ.CumulativityInfo.from_universe_context univs)
else
Polymorphic_ind_entry univs
| Monomorphic_const_entry univs ->
@@ -636,5 +614,8 @@ let definition_structure (kind,cum,poly,finite,(is_coe,((loc,idstruc),pl)),ps,cf
in
let ind = declare_structure finite pl univs idstruc
idbuild implpars params arity template implfs
- fields is_coe (List.map (fun coe -> not (Option.is_empty coe)) coers) sign in
+ fields is_coe (List.map (fun coe -> not (Option.is_empty coe)) coers) in
IndRef ind
+ in
+ Declare.declare_univ_binders gr pl;
+ gr
diff --git a/vernac/record.mli b/vernac/record.mli
index 9fdd5e1c4..e0a4b8fdd 100644
--- a/vernac/record.mli
+++ b/vernac/record.mli
@@ -13,9 +13,20 @@ open Globnames
val primitive_flag : bool ref
+val declare_projections :
+ inductive ->
+ Entries.constant_universes_entry ->
+ ?kind:Decl_kinds.definition_object_kind ->
+ Id.t ->
+ bool list ->
+ Universes.universe_binders ->
+ Impargs.manual_implicits list ->
+ Context.Rel.t ->
+ (Name.t * bool) list * Constant.t option list
+
val definition_structure :
inductive_kind * Decl_kinds.cumulative_inductive_flag * Decl_kinds.polymorphic *
- Decl_kinds.recursivity_kind * ident_decl with_coercion * local_binder_expr list *
+ 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
diff --git a/vernac/topfmt.ml b/vernac/topfmt.ml
index 6a10eb43a..1ad7ead72 100644
--- a/vernac/topfmt.ml
+++ b/vernac/topfmt.ml
@@ -6,7 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Feedback
open Pp
(** Pp control also belongs here as the terminal is private to the toplevel *)
@@ -138,7 +137,7 @@ let make_body quoter info ?pre_hdr s =
(* The empty quoter *)
let noq x = x
(* Generic logger *)
-let gen_logger dbg warn ?pre_hdr level msg = match level with
+let gen_logger dbg warn ?pre_hdr level msg = let open Feedback in match level with
| Debug -> msgnl_with !std_ft (make_body dbg dbg_hdr ?pre_hdr msg)
| Info -> msgnl_with !std_ft (make_body dbg info_hdr ?pre_hdr msg)
| Notice -> msgnl_with !std_ft (make_body noq info_hdr ?pre_hdr msg)
@@ -288,7 +287,6 @@ let init_terminal_output ~color =
*)
let emacs_logger = gen_logger Emacs.quote_info Emacs.quote_warning
-
(* This is specific to the toplevel *)
let pr_loc loc =
let fname = loc.Loc.fname in
@@ -311,17 +309,23 @@ let print_err_exn ?extra any =
std_logger ~pre_hdr Feedback.Error msg
let with_output_to_file fname func input =
- (* XXX FIXME: redirect std_ft *)
- (* let old_logger = !logger in *)
let channel = open_out (String.concat "." [fname; "out"]) in
- (* logger := ft_logger old_logger (Format.formatter_of_out_channel channel); *)
+ let old_fmt = !std_ft, !err_ft, !deep_ft in
+ let new_ft = Format.formatter_of_out_channel channel in
+ std_ft := new_ft;
+ err_ft := new_ft;
+ deep_ft := new_ft;
try
let output = func input in
- (* logger := old_logger; *)
+ std_ft := Util.pi1 old_fmt;
+ err_ft := Util.pi2 old_fmt;
+ deep_ft := Util.pi3 old_fmt;
close_out channel;
output
with reraise ->
let reraise = Backtrace.add_backtrace reraise in
- (* logger := old_logger; *)
+ std_ft := Util.pi1 old_fmt;
+ err_ft := Util.pi2 old_fmt;
+ deep_ft := Util.pi3 old_fmt;
close_out channel;
Exninfo.iraise reraise
diff --git a/vernac/vernac.mllib b/vernac/vernac.mllib
index 8673155e2..f001b572a 100644
--- a/vernac/vernac.mllib
+++ b/vernac/vernac.mllib
@@ -11,7 +11,11 @@ Search
Indschemes
DeclareDef
Obligations
-Command
+ComDefinition
+ComAssumption
+ComInductive
+ComFixpoint
+ComProgramFixpoint
Classes
Record
Assumptions
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
index 7f3410bc1..fa457c895 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -18,7 +18,6 @@ open Tacmach
open Constrintern
open Prettyp
open Printer
-open Command
open Goptions
open Libnames
open Globnames
@@ -59,20 +58,20 @@ let show_proof () =
let p = Proof_global.give_me_the_proof () in
let sigma, env = Pfedit.get_current_context () in
let pprf = Proof.partial_proof p in
- Feedback.msg_notice (Pp.prlist_with_sep Pp.fnl (Printer.pr_econstr_env env sigma) pprf)
+ Pp.prlist_with_sep Pp.fnl (Printer.pr_econstr_env env sigma) pprf
let show_top_evars () =
(* spiwack: new as of Feb. 2010: shows goal evars in addition to non-goal evars. *)
let pfts = Proof_global.give_me_the_proof () in
let gls,_,_,_,sigma = Proof.proof pfts in
- Feedback.msg_notice (pr_evars_int sigma 1 (Evd.undefined_map sigma))
+ pr_evars_int sigma 1 (Evd.undefined_map sigma)
let show_universes () =
let pfts = Proof_global.give_me_the_proof () in
let gls,_,_,_,sigma = Proof.proof pfts in
let ctx = Evd.universe_context_set (Evd.nf_constraints sigma) in
- Feedback.msg_notice (Termops.pr_evar_universe_context (Evd.evar_universe_context sigma));
- Feedback.msg_notice (str"Normalized constraints: " ++ Univ.pr_universe_context_set (Termops.pr_evd_level sigma) ctx)
+ Termops.pr_evar_universe_context (Evd.evar_universe_context sigma) ++ fnl () ++
+ str "Normalized constraints: " ++ Univ.pr_universe_context_set (Termops.pr_evd_level sigma) ctx
(* Simulate the Intro(s) tactic *)
let show_intro all =
@@ -84,11 +83,12 @@ let show_intro all =
let l,_= decompose_prod_assum sigma (Termops.strip_outer_cast sigma (pf_concl gl)) in
if all then
let lid = Tactics.find_intro_names l gl in
- Feedback.msg_notice (hov 0 (prlist_with_sep spc Id.print lid))
+ hov 0 (prlist_with_sep spc Id.print lid)
else if not (List.is_empty l) then
let n = List.last l in
- Feedback.msg_notice (Id.print (List.hd (Tactics.find_intro_names [n] gl)))
- end
+ Id.print (List.hd (Tactics.find_intro_names [n] gl))
+ else mt ()
+ end else mt ()
(** Prepare a "match" template for a given inductive type.
For each branch of the match, we list the constructor name
@@ -147,8 +147,8 @@ let show_match id =
let pr_branch l =
str "| " ++ hov 1 (prlist_with_sep spc str l) ++ str " =>"
in
- Feedback.msg_notice (v 1 (str "match # with" ++ fnl () ++
- prlist_with_sep fnl pr_branch patterns ++ fnl () ++ str "end" ++ fnl ()))
+ v 1 (str "match # with" ++ fnl () ++
+ prlist_with_sep fnl pr_branch patterns ++ fnl () ++ str "end" ++ fnl ())
(* "Print" commands *)
@@ -186,24 +186,24 @@ let print_module r =
try
let globdir = Nametab.locate_dir qid in
match globdir with
- DirModule (dirpath,(mp,_)) ->
- Feedback.msg_notice (Printmod.print_module (Printmod.printable_body dirpath) mp)
+ DirModule { obj_dir; obj_mp; _ } ->
+ Printmod.print_module (Printmod.printable_body obj_dir) obj_mp
| _ -> raise Not_found
with
- Not_found -> Feedback.msg_error (str"Unknown Module " ++ pr_qualid qid)
+ Not_found -> user_err (str"Unknown Module " ++ pr_qualid qid)
let print_modtype r =
let (loc,qid) = qualid_of_reference r in
try
let kn = Nametab.locate_modtype qid in
- Feedback.msg_notice (Printmod.print_modtype kn)
+ Printmod.print_modtype kn
with Not_found ->
(* Is there a module of this name ? If yes we display its type *)
try
let mp = Nametab.locate_module qid in
- Feedback.msg_notice (Printmod.print_module false mp)
+ Printmod.print_module false mp
with Not_found ->
- Feedback.msg_error (str"Unknown Module Type or Module " ++ pr_qualid qid)
+ user_err (str"Unknown Module Type or Module " ++ pr_qualid qid)
let print_namespace ns =
let ns = List.rev (Names.DirPath.repr ns) in
@@ -273,7 +273,7 @@ let print_namespace ns =
acc
) constants (str"")
in
- Feedback.msg_notice ((print_list Id.print ns)++str":"++fnl()++constants_in_namespace)
+ (print_list Id.print ns)++str":"++fnl()++constants_in_namespace
let print_strategy r =
let open Conv_oracle in
@@ -303,7 +303,7 @@ let print_strategy r =
else str "Constant strategies" ++ fnl () ++
hov 0 (prlist_with_sep fnl pr_strategy cst_lvl)
in
- Feedback.msg_notice (var_msg ++ cst_msg)
+ var_msg ++ cst_msg
| Some r ->
let r = Smartlocate.smart_global r in
let key = match r with
@@ -312,7 +312,7 @@ let print_strategy r =
| IndRef _ | ConstructRef _ -> user_err Pp.(str "The reference is not unfoldable")
in
let lvl = get_strategy oracle key in
- Feedback.msg_notice (pr_strategy (r, lvl))
+ pr_strategy (r, lvl)
let dump_universes_gen g s =
let output = open_out s in
@@ -346,7 +346,7 @@ let dump_universes_gen g s =
try
UGraph.dump_universes output_constraint g;
close ();
- Feedback.msg_info (str "Universes written to file \"" ++ str s ++ str "\".")
+ str "Universes written to file \"" ++ str s ++ str "\"."
with reraise ->
let reraise = CErrors.push reraise in
close ();
@@ -361,12 +361,9 @@ let locate_file f =
let msg_found_library = function
| Library.LibLoaded, fulldir, file ->
- Feedback.msg_info (hov 0
- (DirPath.print fulldir ++ strbrk " has been loaded from file " ++
- str file))
+ hov 0 (DirPath.print fulldir ++ strbrk " has been loaded from file " ++ str file)
| Library.LibInPath, fulldir, file ->
- Feedback.msg_info (hov 0
- (DirPath.print fulldir ++ strbrk " is bound to file " ++ str file))
+ hov 0 (DirPath.print fulldir ++ strbrk " is bound to file " ++ str file)
let err_unmapped_library ?loc ?from qid =
let dir = fst (repr_qualid qid) in
@@ -409,8 +406,8 @@ let dump_global r =
(**********)
(* Syntax *)
-let vernac_syntax_extension atts local infix l =
- let local = enforce_module_locality atts.locality local in
+let vernac_syntax_extension atts infix l =
+ let local = enforce_module_locality atts.locality in
if infix then Metasyntax.check_infix_modifiers (snd l);
Metasyntax.add_syntax_extension local l
@@ -421,20 +418,20 @@ let vernac_delimiters sc = function
let vernac_bind_scope sc cll =
Metasyntax.add_class_scope sc (List.map scope_class_of_qualid cll)
-let vernac_open_close_scope ~atts local (b,s) =
- let local = enforce_section_locality atts.locality local in
+let vernac_open_close_scope ~atts (b,s) =
+ let local = enforce_section_locality atts.locality in
Notation.open_close_scope (local,b,s)
let vernac_arguments_scope ~atts r scl =
let local = make_section_locality atts.locality in
Notation.declare_arguments_scope local (smart_global r) scl
-let vernac_infix ~atts local =
- let local = enforce_module_locality atts.locality local in
+let vernac_infix ~atts =
+ let local = enforce_module_locality atts.locality in
Metasyntax.add_infix local (Global.env())
-let vernac_notation ~atts local =
- let local = enforce_module_locality atts.locality local in
+let vernac_notation ~atts =
+ let local = enforce_module_locality atts.locality in
Metasyntax.add_notation local (Global.env())
(***********)
@@ -446,11 +443,13 @@ let start_proof_and_print k l hook =
let hook env sigma ev =
let tac = !Obligations.default_tactic in
let evi = Evd.find sigma ev in
+ let evi = Evarutil.nf_evar_info sigma evi in
let env = Evd.evar_filtered_env evi in
try
- let concl = Evarutil.nf_evars_universes sigma evi.Evd.evar_concl in
- let concl = EConstr.of_constr concl in
- if Evarutil.has_undefined_evars sigma concl then raise Exit;
+ let concl = EConstr.of_constr evi.Evd.evar_concl in
+ if not (Evarutil.is_ground_env sigma env &&
+ Evarutil.is_ground_term sigma concl)
+ then raise Exit;
let c, _, ctx =
Pfedit.build_by_tactic env (Evd.evar_universe_context sigma)
concl (Tacticals.New.tclCOMPLETE tac)
@@ -472,32 +471,40 @@ let vernac_definition_hook p = function
| SubClass -> Class.add_subclass_hook p
| _ -> no_hook
-let vernac_definition ~atts (local,k) ((loc,id as lid),pl) def =
- let local = enforce_locality_exp atts.locality local in
- let hook = vernac_definition_hook atts.polymorphic k in
- let () = match local with
- | Discharge -> Dumpglob.dump_definition lid true "var"
- | Local | Global -> Dumpglob.dump_definition lid false "def"
+let vernac_definition ~atts discharge kind ((loc, id), pl) def =
+ let local = enforce_locality_exp atts.locality discharge in
+ let hook = vernac_definition_hook atts.polymorphic kind in
+ let () =
+ match id with
+ | Anonymous -> ()
+ | Name n -> let lid = (loc, n) in
+ match local with
+ | Discharge -> Dumpglob.dump_definition lid true "var"
+ | Local | Global -> Dumpglob.dump_definition lid false "def"
+ in
+ let program_mode = Flags.is_program_mode () in
+ let name =
+ match id with
+ | Anonymous -> fresh_name_for_anonymous_theorem ()
+ | Name n -> n
in
(match def with
| ProveBody (bl,t) -> (* local binders, typ *)
- start_proof_and_print (local, atts.polymorphic, DefinitionBody k)
- [Some (lid,pl), (bl,t)] hook
+ start_proof_and_print (local, atts.polymorphic, DefinitionBody kind)
+ [((loc, name), pl), (bl, t)] hook
| DefineBody (bl,red_option,c,typ_opt) ->
- let red_option = match red_option with
+ let red_option = match red_option with
| None -> None
| Some r ->
let sigma, env = Pfedit.get_current_context () in
Some (snd (Hook.get f_interp_redexp env sigma r)) in
- do_definition id (local, atts.polymorphic,k) pl bl red_option c typ_opt hook)
+ ComDefinition.do_definition ~program_mode name
+ (local, atts.polymorphic, kind) pl bl red_option c typ_opt hook)
let vernac_start_proof ~atts kind l =
- let local = enforce_locality_exp atts.locality None in
+ let local = enforce_locality_exp atts.locality NoDischarge in
if Dumpglob.dump () then
- List.iter (fun (id, _) ->
- match id with
- | Some (lid,_) -> Dumpglob.dump_definition lid false "prf"
- | None -> ()) l;
+ List.iter (fun ((id, _), _) -> Dumpglob.dump_definition id false "prf") l;
start_proof_and_print (local, atts.polymorphic, Proof kind) l no_hook
let vernac_end_proof ?proof = function
@@ -511,8 +518,8 @@ let vernac_exact_proof c =
save_proof (Vernacexpr.(Proved(Opaque,None)));
if not status then Feedback.feedback Feedback.AddedAxiom
-let vernac_assumption ~atts (local, kind) l nl =
- let local = enforce_locality_exp atts.locality local in
+let vernac_assumption ~atts discharge kind l nl =
+ let local = enforce_locality_exp atts.locality discharge in
let global = local == Global in
let kind = local, atts.polymorphic, kind in
List.iter (fun (is_coe,(idl,c)) ->
@@ -520,7 +527,7 @@ let vernac_assumption ~atts (local, kind) l nl =
List.iter (fun (lid, _) ->
if global then Dumpglob.dump_definition lid false "ax"
else Dumpglob.dump_definition lid true "var") idl) l;
- let status = do_assumptions kind nl l in
+ let status = ComAssumption.do_assumptions kind nl l in
if not status then Feedback.feedback Feedback.AddedAxiom
let should_treat_as_cumulative cum poly =
@@ -592,18 +599,29 @@ let vernac_inductive ~atts cum lo finite indl =
| _ -> user_err Pp.(str "Cannot handle mutually (co)inductive records.")
in
let indl = List.map unpack indl in
- do_mutual_inductive indl is_cumulative atts.polymorphic lo finite
+ ComInductive.do_mutual_inductive indl is_cumulative atts.polymorphic lo finite
-let vernac_fixpoint ~atts local l =
- let local = enforce_locality_exp atts.locality local in
+let vernac_fixpoint ~atts discharge l =
+ let local = enforce_locality_exp atts.locality discharge in
if Dumpglob.dump () then
List.iter (fun (((lid,_), _, _, _, _), _) -> Dumpglob.dump_definition lid false "def") l;
+ (* XXX: Switch to the attribute system and match on ~atts *)
+ let do_fixpoint = if Flags.is_program_mode () then
+ ComProgramFixpoint.do_fixpoint
+ else
+ ComFixpoint.do_fixpoint
+ in
do_fixpoint local atts.polymorphic l
-let vernac_cofixpoint ~atts local l =
- let local = enforce_locality_exp atts.locality local in
+let vernac_cofixpoint ~atts discharge l =
+ let local = enforce_locality_exp atts.locality discharge in
if Dumpglob.dump () then
List.iter (fun (((lid,_), _, _, _), _) -> Dumpglob.dump_definition lid false "def") l;
+ let do_cofixpoint = if Flags.is_program_mode () then
+ ComProgramFixpoint.do_cofixpoint
+ else
+ ComFixpoint.do_cofixpoint
+ in
do_cofixpoint local atts.polymorphic l
let vernac_scheme l =
@@ -627,14 +645,14 @@ let vernac_universe ~atts l =
user_err ?loc:atts.loc ~hdr:"vernac_universe"
(str"Polymorphic universes can only be declared inside sections, " ++
str "use Monomorphic Universe instead");
- do_universe atts.polymorphic l
+ Declare.do_universe atts.polymorphic l
let vernac_constraint ~atts l =
if atts.polymorphic && not (Lib.sections_are_opened ()) then
user_err ?loc:atts.loc ~hdr:"vernac_constraint"
(str"Polymorphic universe constraints can only be declared"
++ str " inside sections, use Monomorphic Constraint instead");
- do_constraint atts.polymorphic l
+ Declare.do_constraint atts.polymorphic l
(**********************)
(* Modules *)
@@ -812,16 +830,16 @@ let vernac_require from import qidl =
let vernac_canonical r =
Recordops.declare_canonical_structure (smart_global r)
-let vernac_coercion ~atts local ref qids qidt =
- let local = enforce_locality atts.locality local in
+let vernac_coercion ~atts ref qids qidt =
+ let local = enforce_locality atts.locality in
let target = cl_of_qualid qidt in
let source = cl_of_qualid qids in
let ref' = smart_global ref in
Class.try_add_new_coercion_with_target ref' ~local atts.polymorphic ~source ~target;
Flags.if_verbose Feedback.msg_info (pr_global ref' ++ str " is now a coercion")
-let vernac_identity_coercion ~atts local id qids qidt =
- let local = enforce_locality atts.locality local in
+let vernac_identity_coercion ~atts id qids qidt =
+ let local = enforce_locality atts.locality in
let target = cl_of_qualid qidt in
let source = cl_of_qualid qids in
Class.try_add_new_identity_coercion id ~local atts.polymorphic ~source ~target
@@ -831,7 +849,8 @@ let vernac_identity_coercion ~atts local 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";
- ignore(Classes.new_instance ~abstract:abst ~global atts.polymorphic sup inst props pri)
+ let program_mode = Flags.is_program_mode () in
+ ignore(Classes.new_instance ~program_mode ~abstract:abst ~global atts.polymorphic sup inst props pri)
let vernac_context ~atts l =
if not (Classes.context atts.polymorphic l) then Feedback.feedback Feedback.AddedAxiom
@@ -893,9 +912,11 @@ let expand filename =
Envars.expand_path_macros ~warn:(fun x -> Feedback.msg_warning (str x)) filename
let vernac_add_loadpath implicit pdir ldiropt =
+ let open Mltop in
let pdir = expand pdir in
let alias = Option.default Libnames.default_root_prefix ldiropt in
- Mltop.add_rec_path Mltop.AddTopML ~unix_path:pdir ~coq_root:alias ~implicit
+ add_coq_path { recursive = true;
+ path_spec = VoPath { unix_path = pdir; coq_path = alias; has_ml = AddTopML; implicit } }
let vernac_remove_loadpath path =
Loadpath.remove_load_path (expand path)
@@ -903,7 +924,8 @@ let vernac_remove_loadpath path =
(* Coq syntax for ML or system commands *)
let vernac_add_ml_path isrec path =
- (if isrec then Mltop.add_rec_ml_dir else Mltop.add_ml_dir) (expand path)
+ let open Mltop in
+ add_coq_path { recursive = isrec; path_spec = MlPath (expand path) }
let vernac_declare_ml_module ~atts l =
let local = make_locality atts.locality in
@@ -922,7 +944,6 @@ let vernac_chdir = function
end;
Flags.if_verbose Feedback.msg_info (str (Sys.getcwd()))
-
(********************)
(* State management *)
@@ -947,13 +968,13 @@ let vernac_remove_hints ~atts dbs ids =
let local = make_module_locality atts.locality in
Hints.remove_hints local dbs (List.map Smartlocate.global_with_alias ids)
-let vernac_hints ~atts local lb h =
- let local = enforce_module_locality atts.locality local in
+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)
-let vernac_syntactic_definition ~atts lid x local y =
+let vernac_syntactic_definition ~atts lid x y =
Dumpglob.dump_definition lid false "syndef";
- let local = enforce_module_locality atts.locality local in
+ let local = enforce_module_locality atts.locality in
Metasyntax.add_syntactic_definition (Global.env()) (snd lid) x local y
let vernac_declare_implicits ~atts r l =
@@ -1368,11 +1389,13 @@ let _ =
optread = (fun () -> !Flags.program_mode);
optwrite = (fun b -> Flags.program_mode:=b) }
+let universe_polymorphism_option_name = ["Universe"; "Polymorphism"]
+
let _ =
declare_bool_option
{ optdepr = false;
optname = "universe polymorphism";
- optkey = ["Universe"; "Polymorphism"];
+ optkey = universe_polymorphism_option_name;
optread = Flags.is_universe_polymorphism;
optwrite = Flags.make_universe_polymorphism }
@@ -1570,10 +1593,10 @@ let vernac_check_may_eval ~atts redexp glopt rc =
| None ->
let evars_of_term c = Evarutil.undefined_evars_of_term sigma' c in
let l = Evar.Set.union (evars_of_term j.Environ.uj_val) (evars_of_term j.Environ.uj_type) in
- let j = { j with Environ.uj_type = Reductionops.nf_betaiota sigma' j.Environ.uj_type } in
- Feedback.msg_notice (print_judgment env sigma' j ++
- pr_ne_evar_set (fnl () ++ str "where" ++ fnl ()) (mt ()) sigma' l ++
- Printer.pr_universe_ctx_set sigma uctx)
+ let j = { j with Environ.uj_type = Reductionops.nf_betaiota env sigma' j.Environ.uj_type } in
+ print_judgment env sigma' j ++
+ pr_ne_evar_set (fnl () ++ str "where" ++ fnl ()) (mt ()) sigma' l ++
+ Printer.pr_universe_ctx_set sigma uctx
| Some r ->
let (sigma',r_interp) = Hook.get f_interp_redexp env sigma' r in
let redfun env evm c =
@@ -1581,7 +1604,7 @@ let vernac_check_may_eval ~atts redexp glopt rc =
let (_, c) = redfun env evm c in
c
in
- Feedback.msg_notice (print_eval redfun env sigma' rc j)
+ print_eval redfun env sigma' rc j
let vernac_declare_reduction ~atts s r =
let local = make_locality atts.locality in
@@ -1591,13 +1614,14 @@ let vernac_declare_reduction ~atts s r =
let vernac_global_check c =
let env = Global.env() in
let sigma = Evd.from_env env in
- let c,ctx = interp_constr env sigma c in
+ let c,uctx = interp_constr env sigma c in
let senv = Global.safe_env() in
- let cstrs = snd (UState.context_set ctx) in
- let senv = Safe_typing.add_constraints cstrs senv in
+ let uctx = UState.context_set uctx in
+ let senv = Safe_typing.push_context_set false uctx senv in
let j = Safe_typing.typing senv c in
let env = Safe_typing.env_of_safe_env senv in
- Feedback.msg_notice (print_safe_judgment env sigma j)
+ print_safe_judgment env sigma j ++
+ pr_universe_ctx_set sigma uctx
let get_nth_goal n =
@@ -1605,7 +1629,7 @@ let get_nth_goal n =
let gls,_,_,_,sigma = Proof.proof pf in
let gl = {Evd.it=List.nth gls (n-1) ; sigma = sigma; } in
gl
-
+
exception NoHyp
(* Printing "About" information of a hypothesis of the current goal.
We only print the type and a small statement to this comes from the
@@ -1639,31 +1663,32 @@ let print_about_hyp_globs ?loc ref_or_by_not udecl glopt =
print_about env sigma ref_or_by_not udecl
let vernac_print ~atts env sigma =
- let open Feedback in
let loc = atts.loc in
function
- | PrintTables -> msg_notice (print_tables ())
- | PrintFullContext-> msg_notice (print_full_context_typ env sigma)
- | PrintSectionContext qid -> msg_notice (print_sec_context_typ env sigma qid)
- | PrintInspect n -> msg_notice (inspect env sigma n)
- | PrintGrammar ent -> msg_notice (Metasyntax.pr_grammar ent)
- | PrintLoadPath dir -> (* For compatibility ? *) msg_notice (print_loadpath dir)
- | PrintModules -> msg_notice (print_modules ())
+ | PrintTables -> print_tables ()
+ | PrintFullContext-> print_full_context_typ env sigma
+ | PrintSectionContext qid -> print_sec_context_typ env sigma qid
+ | PrintInspect n -> inspect env sigma n
+ | PrintGrammar ent -> Metasyntax.pr_grammar ent
+ | PrintLoadPath dir -> (* For compatibility ? *) print_loadpath dir
+ | PrintModules -> print_modules ()
| PrintModule qid -> print_module qid
| PrintModuleType qid -> print_modtype qid
| PrintNamespace ns -> print_namespace ns
- | PrintMLLoadPath -> msg_notice (Mltop.print_ml_path ())
- | PrintMLModules -> msg_notice (Mltop.print_ml_modules ())
- | PrintDebugGC -> msg_notice (Mltop.print_gc ())
- | PrintName (qid,udecl) -> dump_global qid; msg_notice (print_name env sigma qid udecl)
- | PrintGraph -> msg_notice (Prettyp.print_graph())
- | PrintClasses -> msg_notice (Prettyp.print_classes())
- | PrintTypeClasses -> msg_notice (Prettyp.print_typeclasses())
- | PrintInstances c -> msg_notice (Prettyp.print_instances (smart_global c))
- | PrintCoercions -> msg_notice (Prettyp.print_coercions env sigma)
+ | PrintMLLoadPath -> Mltop.print_ml_path ()
+ | PrintMLModules -> Mltop.print_ml_modules ()
+ | PrintDebugGC -> Mltop.print_gc ()
+ | PrintName (qid,udecl) ->
+ dump_global qid;
+ print_name env sigma qid udecl
+ | PrintGraph -> Prettyp.print_graph env sigma
+ | PrintClasses -> Prettyp.print_classes()
+ | PrintTypeClasses -> Prettyp.print_typeclasses()
+ | PrintInstances c -> Prettyp.print_instances (smart_global c)
+ | PrintCoercions -> Prettyp.print_coercions env sigma
| PrintCoercionPaths (cls,clt) ->
- msg_notice (Prettyp.print_path_between (cl_of_qualid cls) (cl_of_qualid clt))
- | PrintCanonicalConversions -> msg_notice (Prettyp.print_canonical_projections env sigma)
+ Prettyp.print_path_between env sigma (cl_of_qualid cls) (cl_of_qualid clt)
+ | PrintCanonicalConversions -> Prettyp.print_canonical_projections env sigma
| PrintUniverses (b, dst) ->
let univ = Global.universes () in
let univ = if b then UGraph.sort_universes univ else univ in
@@ -1672,23 +1697,24 @@ let vernac_print ~atts env sigma =
else str"There may remain asynchronous universe constraints"
in
begin match dst with
- | None -> msg_notice (UGraph.pr_universes Universes.pr_with_global_universes univ ++ pr_remaining)
+ | None -> UGraph.pr_universes Universes.pr_with_global_universes univ ++ pr_remaining
| Some s -> dump_universes_gen univ s
end
- | PrintHint r -> msg_notice (Hints.pr_hint_ref env sigma (smart_global r))
- | PrintHintGoal -> msg_notice (Hints.pr_applicable_hint ())
- | PrintHintDbName s -> msg_notice (Hints.pr_hint_db_by_name env sigma s)
- | PrintHintDb -> msg_notice (Hints.pr_searchtable env sigma)
+ | PrintHint r -> Hints.pr_hint_ref env sigma (smart_global r)
+ | PrintHintGoal -> Hints.pr_applicable_hint ()
+ | PrintHintDbName s -> Hints.pr_hint_db_by_name env sigma s
+ | PrintHintDb -> Hints.pr_searchtable env sigma
| PrintScopes ->
- msg_notice (Notation.pr_scopes (Constrextern.without_symbols (pr_lglob_constr_env env)))
+ Notation.pr_scopes (Constrextern.without_symbols (pr_lglob_constr_env env))
| PrintScope s ->
- msg_notice (Notation.pr_scope (Constrextern.without_symbols (pr_lglob_constr_env env)) s)
+ Notation.pr_scope (Constrextern.without_symbols (pr_lglob_constr_env env)) s
| PrintVisibility s ->
- msg_notice (Notation.pr_visibility (Constrextern.without_symbols (pr_lglob_constr_env env)) s)
+ Notation.pr_visibility (Constrextern.without_symbols (pr_lglob_constr_env env)) s
| PrintAbout (ref_or_by_not,udecl,glnumopt) ->
- msg_notice (print_about_hyp_globs ?loc ref_or_by_not udecl glnumopt)
+ print_about_hyp_globs ?loc ref_or_by_not udecl glnumopt
| PrintImplicit qid ->
- dump_global qid; msg_notice (print_impargs qid)
+ dump_global qid;
+ print_impargs qid
| PrintAssumptions (o,t,r) ->
(* Prints all the axioms and section variables used by a term *)
let gr = smart_global r in
@@ -1696,7 +1722,7 @@ let vernac_print ~atts env sigma =
let st = Conv_oracle.get_transp_state (Environ.oracle (Global.env())) in
let nassums =
Assumptions.assumptions st ~add_opaque:o ~add_transparent:t gr cstr in
- msg_notice (Printer.pr_assumptionset (Global.env ()) nassums)
+ Printer.pr_assumptionset env sigma nassums
| PrintStrategy r -> print_strategy r
let global_module r =
@@ -1782,19 +1808,18 @@ let vernac_search ~atts s gopt r =
| SearchAbout sl ->
(Search.search_about gopt (List.map (on_snd (interp_search_about_item env)) sl) r |> Search.prioritize_search) pr_search
-let vernac_locate = let open Feedback in function
- | LocateAny (AN qid) -> msg_notice (print_located_qualid qid)
- | LocateTerm (AN qid) -> msg_notice (print_located_term qid)
+let vernac_locate = function
+ | LocateAny (AN qid) -> print_located_qualid qid
+ | LocateTerm (AN qid) -> print_located_term qid
| LocateAny (ByNotation (_, (ntn, sc))) (** TODO : handle Ltac notations *)
| LocateTerm (ByNotation (_, (ntn, sc))) ->
let _, env = Pfedit.get_current_context () in
- msg_notice
- (Notation.locate_notation
- (Constrextern.without_symbols (pr_lglob_constr_env env)) ntn sc)
+ Notation.locate_notation
+ (Constrextern.without_symbols (pr_lglob_constr_env env)) ntn sc
| LocateLibrary qid -> print_located_library qid
- | LocateModule qid -> msg_notice (print_located_module qid)
- | LocateOther (s, qid) -> msg_notice (print_located_other s qid)
- | LocateFile f -> msg_notice (locate_file f)
+ | LocateModule qid -> print_located_module qid
+ | LocateOther (s, qid) -> print_located_other s qid
+ | LocateFile f -> locate_file f
let vernac_register id r =
if Proof_global.there_are_pending_proofs () then
@@ -1826,16 +1851,13 @@ let vernac_unfocus () =
let vernac_unfocused () =
let p = Proof_global.give_me_the_proof () in
if Proof.unfocused p then
- Feedback.msg_notice (str"The proof is indeed fully unfocused.")
+ str"The proof is indeed fully unfocused."
else
user_err Pp.(str "The proof is not fully unfocused.")
-(* BeginSubproof / EndSubproof.
- BeginSubproof (vernac_subproof) focuses on the first goal, or the goal
- given as argument.
- EndSubproof (vernac_end_subproof) unfocuses from a BeginSubproof, provided
- that the proof of the goal has been completed.
+(* "{" focuses on the first goal, "n: {" focuses on the n-th goal
+ "}" unfocuses, provided that the proof of the goal has been completed.
*)
let subproof_kind = Proof.new_focus_kind ()
let subproof_cond = Proof.done_cond subproof_kind
@@ -1844,7 +1866,9 @@ let vernac_subproof gln =
Proof_global.simple_with_current_proof (fun _ p ->
match gln with
| None -> Proof.focus subproof_cond () 1 p
- | Some n -> Proof.focus subproof_cond () n p)
+ | Some (SelectNth n) -> Proof.focus subproof_cond () n p
+ | _ -> user_err ~hdr:"bracket_selector"
+ (str "Brackets only support the single numbered goal selector."))
let vernac_end_subproof () =
Proof_global.simple_with_current_proof (fun _ p ->
@@ -1854,21 +1878,20 @@ let vernac_bullet (bullet : Proof_bullet.t) =
Proof_global.simple_with_current_proof (fun _ p ->
Proof_bullet.put p bullet)
-let vernac_show = let open Feedback in function
+let vernac_show = function
| ShowScript -> assert false (* Only the stm knows the script *)
| ShowGoal goalref ->
let proof = Proof_global.give_me_the_proof () in
- let info = match goalref with
- | OpenSubgoals -> pr_open_subgoals ~proof
- | NthGoal n -> pr_nth_open_subgoal ~proof n
- | GoalId id -> pr_goal_by_id ~proof id
- in
- msg_notice info
+ begin match goalref with
+ | OpenSubgoals -> pr_open_subgoals ~proof
+ | NthGoal n -> pr_nth_open_subgoal ~proof n
+ | GoalId id -> pr_goal_by_id ~proof id
+ end
| ShowProof -> show_proof ()
| ShowExistentials -> show_top_evars ()
| ShowUniverses -> show_universes ()
| ShowProofNames ->
- msg_notice (pr_sequence Id.print (Proof_global.get_all_proof_names()))
+ pr_sequence Id.print (Proof_global.get_all_proof_names())
| ShowIntros all -> show_intro all
| ShowMatch id -> show_match id
@@ -1883,8 +1906,7 @@ let vernac_check_guard () =
(str "The condition holds up to here")
with UserError(_,s) ->
(str ("Condition violated: ") ++s)
- in
- Feedback.msg_notice message
+ in message
exception End_of_input
@@ -1920,19 +1942,11 @@ let vernac_load interp fname =
* loc is the Loc.t of the vernacular command being interpreted. *)
let interp ?proof ~atts ~st c =
let open Vernacinterp in
- vernac_pperr_endline (fun () -> str "interpreting: " ++ Ppvernac.pr_vernac c);
+ vernac_pperr_endline (fun () -> str "interpreting: " ++ Ppvernac.pr_vernac_expr c);
match c with
- (* The below vernac are candidates for removal from the main type
- and to be put into a new doc_command datatype: *)
| VernacLoad _ -> assert false
- (* Done later in this file *)
- | VernacFail _ -> assert false
- | VernacTime _ -> assert false
- | VernacRedirect _ -> assert false
- | VernacTimeout _ -> assert false
-
(* The STM should handle that, but LOAD bypasses the STM... *)
| VernacAbortAll -> CErrors.user_err (str "AbortAll cannot be used through the Load command")
| VernacRestart -> CErrors.user_err (str "Restart cannot be used through the Load command")
@@ -1952,33 +1966,30 @@ let interp ?proof ~atts ~st c =
(* This one is possible to handle here *)
| VernacAbort id -> CErrors.user_err (str "Abort cannot be used through the Load command")
- (* Handled elsewhere *)
- | VernacProgram _
- | VernacPolymorphic _
- | VernacLocal _ -> assert false
-
(* Syntax *)
- | VernacSyntaxExtension (infix, local,sl) ->
- vernac_syntax_extension atts local infix sl
+ | VernacSyntaxExtension (infix, sl) ->
+ vernac_syntax_extension atts infix sl
| VernacDelimiters (sc,lr) -> vernac_delimiters sc lr
| VernacBindScope (sc,rl) -> vernac_bind_scope sc rl
- | VernacOpenCloseScope (local, s) -> vernac_open_close_scope ~atts local s
+ | VernacOpenCloseScope (b, s) -> vernac_open_close_scope ~atts (b,s)
| VernacArgumentsScope (qid,scl) -> vernac_arguments_scope ~atts qid scl
- | VernacInfix (local,mv,qid,sc) -> vernac_infix ~atts local mv qid sc
- | VernacNotation (local,c,infpl,sc) ->
- vernac_notation ~atts local c infpl sc
+ | VernacInfix (mv,qid,sc) -> vernac_infix ~atts mv qid sc
+ | VernacNotation (c,infpl,sc) ->
+ vernac_notation ~atts c infpl sc
| VernacNotationAddFormat(n,k,v) ->
Metasyntax.add_notation_extra_printing_rule n k v
(* Gallina *)
- | VernacDefinition (k,lid,d) -> vernac_definition ~atts k lid d
+ | VernacDefinition ((discharge,kind),lid,d) ->
+ vernac_definition ~atts discharge kind lid d
| VernacStartTheoremProof (k,l) -> vernac_start_proof ~atts k l
| VernacEndProof e -> vernac_end_proof ?proof e
| VernacExactProof c -> vernac_exact_proof c
- | VernacAssumption (stre,nl,l) -> vernac_assumption ~atts stre l nl
+ | VernacAssumption ((discharge,kind),nl,l) ->
+ vernac_assumption ~atts discharge kind l nl
| VernacInductive (cum, priv,finite,l) -> vernac_inductive ~atts cum priv finite l
- | VernacFixpoint (local, l) -> vernac_fixpoint ~atts local l
- | VernacCoFixpoint (local, l) -> vernac_cofixpoint ~atts local l
+ | VernacFixpoint (discharge, l) -> vernac_fixpoint ~atts discharge l
+ | VernacCoFixpoint (discharge, l) -> vernac_cofixpoint ~atts discharge l
| VernacScheme l -> vernac_scheme l
| VernacCombinedScheme (id, l) -> vernac_combined_scheme id l
| VernacUniverse l -> vernac_universe ~atts l
@@ -2003,9 +2014,9 @@ let interp ?proof ~atts ~st c =
| VernacRequire (from, export, qidl) -> vernac_require from export qidl
| VernacImport (export,qidl) -> vernac_import export qidl
| VernacCanonical qid -> vernac_canonical qid
- | VernacCoercion (local,r,s,t) -> vernac_coercion ~atts local r s t
- | VernacIdentityCoercion (local,(_,id),s,t) ->
- vernac_identity_coercion ~atts local id s t
+ | VernacCoercion (r,s,t) -> vernac_coercion ~atts r s t
+ | VernacIdentityCoercion ((_,id),s,t) ->
+ vernac_identity_coercion ~atts id s t
(* Type classes *)
| VernacInstance (abst, sup, inst, props, info) ->
@@ -2031,10 +2042,10 @@ let interp ?proof ~atts ~st c =
(* Commands *)
| VernacCreateHintDb (dbname,b) -> vernac_create_hintdb ~atts dbname b
| VernacRemoveHints (dbnames,ids) -> vernac_remove_hints ~atts dbnames ids
- | VernacHints (local,dbnames,hints) ->
- vernac_hints ~atts local dbnames hints
- | VernacSyntacticDefinition (id,c,local,b) ->
- vernac_syntactic_definition ~atts id c local b
+ | VernacHints (dbnames,hints) ->
+ vernac_hints ~atts dbnames hints
+ | VernacSyntacticDefinition (id,c,b) ->
+ vernac_syntactic_definition ~atts id c b
| VernacDeclareImplicits (qid,l) ->
vernac_declare_implicits ~atts qid l
| VernacArguments (qid, args, more_implicits, nargs, flags) ->
@@ -2050,27 +2061,32 @@ let interp ?proof ~atts ~st c =
| VernacAddOption (key,v) -> vernac_add_option key v
| VernacMemOption (key,v) -> vernac_mem_option key v
| VernacPrintOption key -> vernac_print_option key
- | VernacCheckMayEval (r,g,c) -> vernac_check_may_eval ~atts r g c
+ | VernacCheckMayEval (r,g,c) ->
+ Feedback.msg_notice @@ vernac_check_may_eval ~atts r g c
| VernacDeclareReduction (s,r) -> vernac_declare_reduction ~atts s r
- | VernacGlobalCheck c -> vernac_global_check c
+ | VernacGlobalCheck c ->
+ Feedback.msg_notice @@ vernac_global_check c
| VernacPrint p ->
let sigma, env = Pfedit.get_current_context () in
- vernac_print ~atts env sigma p
+ Feedback.msg_notice @@ vernac_print ~atts env sigma p
| VernacSearch (s,g,r) -> vernac_search ~atts s g r
- | VernacLocate l -> vernac_locate l
+ | VernacLocate l ->
+ Feedback.msg_notice @@ vernac_locate l
| VernacRegister (id, r) -> vernac_register id r
| VernacComments l -> Flags.if_verbose Feedback.msg_info (str "Comments ok\n")
(* Proof management *)
- | VernacGoal t -> vernac_start_proof ~atts Theorem [None,([],t)]
| VernacFocus n -> vernac_focus n
| VernacUnfocus -> vernac_unfocus ()
- | VernacUnfocused -> vernac_unfocused ()
+ | VernacUnfocused ->
+ Feedback.msg_notice @@ vernac_unfocused ()
| VernacBullet b -> vernac_bullet b
| VernacSubproof n -> vernac_subproof n
| VernacEndSubproof -> vernac_end_subproof ()
- | VernacShow s -> vernac_show s
- | VernacCheckGuard -> vernac_check_guard ()
+ | VernacShow s ->
+ Feedback.msg_notice @@ vernac_show s
+ | VernacCheckGuard ->
+ Feedback.msg_notice @@ vernac_check_guard ()
| VernacProof (tac, using) ->
let using = Option.append using (Proof_using.get_default_proof_using ()) in
let tacs = if Option.is_empty tac then "tac:no" else "tac:yes" in
@@ -2123,10 +2139,6 @@ let check_vernac_supports_polymorphism c p =
| VernacExtend _ | VernacUniverse _ | VernacConstraint _) -> ()
| Some _, _ -> user_err Pp.(str "This command does not support Polymorphism")
-let enforce_polymorphism = function
- | None -> Flags.is_universe_polymorphism ()
- | Some b -> Flags.make_polymorphic_flag b; b
-
(** A global default timeout, controlled by option "Set Default Timeout n".
Use "Unset Default Timeout" to deactivate it (or set it to 0). *)
@@ -2193,56 +2205,64 @@ let with_fail st b f =
end
let interp ?(verbosely=true) ?proof ~st (loc,c) =
+ let orig_univ_poly = Flags.is_universe_polymorphism () in
let orig_program_mode = Flags.is_program_mode () in
- let rec aux ?polymorphism ~atts isprogcmd = function
-
- | VernacProgram c when not isprogcmd ->
- aux ?polymorphism ~atts true c
-
- | VernacProgram _ ->
- user_err Pp.(str "Program mode specified twice")
-
- | VernacPolymorphic (b, c) when polymorphism = None ->
- aux ~polymorphism:b ~atts:atts isprogcmd c
-
- | VernacPolymorphic (b, c) ->
- user_err Pp.(str "Polymorphism specified twice")
-
- | VernacLocal (b, c) when Option.is_empty atts.locality ->
- aux ?polymorphism ~atts:{atts with locality = Some b} isprogcmd c
-
- | VernacLocal _ ->
- user_err Pp.(str "Locality specified twice")
-
- | VernacFail v ->
- with_fail st true (fun () -> aux ?polymorphism ~atts isprogcmd v)
-
- | VernacTimeout (n,v) ->
+ let flags f atts =
+ List.fold_left
+ (fun (polymorphism, atts) f ->
+ match f with
+ | VernacProgram when not atts.program ->
+ (polymorphism, { atts with program = true })
+ | VernacProgram ->
+ user_err Pp.(str "Program mode specified twice")
+ | VernacPolymorphic b when polymorphism = None ->
+ (Some b, atts)
+ | VernacPolymorphic _ ->
+ user_err Pp.(str "Polymorphism specified twice")
+ | VernacLocal b when Option.is_empty atts.locality ->
+ (polymorphism, { atts with locality = Some b })
+ | VernacLocal _ ->
+ user_err Pp.(str "Locality specified twice")
+ )
+ (None, atts)
+ f
+ in
+ let rec control = function
+ | VernacExpr (f, v) ->
+ let (polymorphism, atts) = flags f { loc; locality = None; polymorphic = false; program = orig_program_mode; } in
+ aux ~polymorphism ~atts v
+ | VernacFail v -> with_fail st true (fun () -> control v)
+ | VernacTimeout (n,v) ->
current_timeout := Some n;
- aux ?polymorphism ~atts isprogcmd v
-
- | VernacRedirect (s, (_,v)) ->
- Topfmt.with_output_to_file s (aux ?polymorphism ~atts isprogcmd) v
+ control v
+ | VernacRedirect (s, (_,v)) ->
+ Topfmt.with_output_to_file s control v
+ | VernacTime (batch, (_loc,v)) ->
+ System.with_time ~batch control v;
- | VernacTime (_,v) ->
- System.with_time !Flags.time (aux ?polymorphism ~atts isprogcmd) v;
+ and aux ~polymorphism ~atts : _ -> unit =
+ function
- | VernacLoad (_,fname) -> vernac_load (aux ?polymorphism ~atts false) fname
+ | VernacLoad (_,fname) -> vernac_load control fname
| c ->
check_vernac_supports_locality c atts.locality;
check_vernac_supports_polymorphism c polymorphism;
- let polymorphic = enforce_polymorphism polymorphism in
- Obligations.set_program_mode isprogcmd;
+ let polymorphic = Option.default (Flags.is_universe_polymorphism ()) polymorphism in
+ Flags.make_universe_polymorphism polymorphic;
+ Obligations.set_program_mode atts.program;
try
vernac_timeout begin fun () ->
let atts = { atts with polymorphic } in
if verbosely
then Flags.verbosely (interp ?proof ~atts ~st) c
else Flags.silently (interp ?proof ~atts ~st) c;
- if orig_program_mode || not !Flags.program_mode || isprogcmd then
+ (* If the command is `(Un)Set Program Mode` or `(Un)Set Universe Polymorphism`,
+ we should not restore the previous state of the flag... *)
+ if orig_program_mode || not !Flags.program_mode || atts.program then
Flags.program_mode := orig_program_mode;
- ignore (Flags.use_polymorphic_flag ())
+ if (Flags.is_universe_polymorphism() = polymorphic) then
+ Flags.make_universe_polymorphism orig_univ_poly;
end
with
| reraise when
@@ -2253,18 +2273,21 @@ let interp ?(verbosely=true) ?proof ~st (loc,c) =
let e = CErrors.push reraise in
let e = locate_if_not_already ?loc e in
let () = restore_timeout () in
+ Flags.make_universe_polymorphism orig_univ_poly;
Flags.program_mode := orig_program_mode;
- ignore (Flags.use_polymorphic_flag ());
iraise e
in
- let atts = { loc; locality = None; polymorphic = false; } in
if verbosely
- then Flags.verbosely (aux ~atts orig_program_mode) c
- else aux ~atts orig_program_mode c
+ then Flags.verbosely control c
+ else control c
-(* XXX: There is a bug here in case of an exception, see @gares
- comments on the PR *)
+(* Be careful with the cache here in case of an exception. *)
let interp ?verbosely ?proof ~st cmd =
Vernacstate.unfreeze_interp_state st;
- interp ?verbosely ?proof ~st cmd;
- Vernacstate.freeze_interp_state `No
+ try
+ interp ?verbosely ?proof ~st cmd;
+ Vernacstate.freeze_interp_state `No
+ with exn ->
+ let exn = CErrors.push exn in
+ Vernacstate.invalidate_cache ();
+ iraise exn
diff --git a/vernac/vernacentries.mli b/vernac/vernacentries.mli
index a559912a0..e99a62fe6 100644
--- a/vernac/vernacentries.mli
+++ b/vernac/vernacentries.mli
@@ -18,7 +18,7 @@ val vernac_require :
val interp :
?verbosely:bool ->
?proof:Proof_global.closed_proof ->
- st:Vernacstate.t -> Vernacexpr.vernac_expr Loc.located -> Vernacstate.t
+ st:Vernacstate.t -> Vernacexpr.vernac_control Loc.located -> Vernacstate.t
(** Prepare a "match" template for a given inductive type.
For each branch of the match, we list the constructor name
@@ -36,3 +36,5 @@ val command_focus : unit Proof.focus_kind
val interp_redexp_hook : (Environ.env -> Evd.evar_map -> Genredexpr.raw_red_expr ->
Evd.evar_map * Redexpr.red_expr) Hook.t
+
+val universe_polymorphism_option_name : string list
diff --git a/vernac/vernacinterp.ml b/vernac/vernacinterp.ml
index c0b93c163..c40ca27db 100644
--- a/vernac/vernacinterp.ml
+++ b/vernac/vernacinterp.ml
@@ -16,6 +16,7 @@ type atts = {
loc : Loc.t option;
locality : bool option;
polymorphic : bool;
+ program : bool;
}
type 'a vernac_command = 'a -> atts:atts -> st:Vernacstate.t -> Vernacstate.t
diff --git a/vernac/vernacinterp.mli b/vernac/vernacinterp.mli
index ab3d4bfc5..c5e610f89 100644
--- a/vernac/vernacinterp.mli
+++ b/vernac/vernacinterp.mli
@@ -14,6 +14,7 @@ type atts = {
loc : Loc.t option;
locality : bool option;
polymorphic : bool;
+ program : bool;
}
type 'a vernac_command = 'a -> atts:atts -> st:Vernacstate.t -> Vernacstate.t
diff --git a/vernac/vernacprop.ml b/vernac/vernacprop.ml
index 3cff1f14c..3932d1c7b 100644
--- a/vernac/vernacprop.ml
+++ b/vernac/vernacprop.ml
@@ -11,42 +11,48 @@
open Vernacexpr
+let rec under_control = function
+ | VernacExpr (_, c) -> c
+ | VernacRedirect (_,(_,c))
+ | VernacTime (_,(_,c))
+ | VernacFail c
+ | VernacTimeout (_,c) -> under_control c
+
+let rec has_Fail = function
+ | VernacExpr _ -> false
+ | VernacRedirect (_,(_,c))
+ | VernacTime (_,(_,c))
+ | VernacTimeout (_,c) -> has_Fail c
+ | VernacFail _ -> true
+
(* Navigation commands are allowed in a coqtop session but not in a .v file *)
-let rec is_navigation_vernac = function
+let is_navigation_vernac_expr = function
| VernacResetInitial
| VernacResetName _
| VernacBacktrack _
| VernacBackTo _
| VernacBack _ -> true
- | VernacRedirect (_, (_,c))
- | VernacTime (_,c) ->
- is_navigation_vernac c (* Time Back* is harmless *)
- | c -> is_deep_navigation_vernac c
+ | _ -> false
+
+let is_navigation_vernac c =
+ is_navigation_vernac_expr (under_control c)
-and is_deep_navigation_vernac = function
+let rec is_deep_navigation_vernac = function
+ | VernacTime (_,(_,c)) -> is_deep_navigation_vernac c
+ | VernacRedirect (_, (_,c))
| VernacTimeout (_,c) | VernacFail c -> is_navigation_vernac c
- | _ -> false
+ | VernacExpr _ -> false
(* NB: Reset is now allowed again as asked by A. Chlipala *)
let is_reset = function
- | VernacResetInitial | VernacResetName _ -> true
+ | VernacExpr ( _, VernacResetInitial)
+ | VernacExpr (_, VernacResetName _) -> true
| _ -> false
-let is_debug cmd = match cmd with
+let is_debug cmd = match under_control cmd with
| VernacSetOption (["Ltac";"Debug"], _) -> true
| _ -> false
-let is_query cmd = match cmd with
- | VernacChdir None
- | VernacMemOption _
- | VernacPrintOption _
- | VernacCheckMayEval _
- | VernacGlobalCheck _
- | VernacPrint _
- | VernacSearch _
- | VernacLocate _ -> true
- | _ -> false
-
-let is_undo cmd = match cmd with
+let is_undo cmd = match under_control cmd with
| VernacUndo _ | VernacUndoTo _ -> true
| _ -> false
diff --git a/vernac/vernacprop.mli b/vernac/vernacprop.mli
index fbdba6bac..df739f96a 100644
--- a/vernac/vernacprop.mli
+++ b/vernac/vernacprop.mli
@@ -11,9 +11,16 @@
open Vernacexpr
-val is_navigation_vernac : vernac_expr -> bool
-val is_deep_navigation_vernac : vernac_expr -> bool
-val is_reset : vernac_expr -> bool
-val is_query : vernac_expr -> bool
-val is_debug : vernac_expr -> bool
-val is_undo : vernac_expr -> bool
+(* Return the vernacular command below control (Time, Timeout, Redirect, Fail).
+ Beware that Fail can change many properties of the underlying command, since
+ a success of Fail means the command was backtracked over. *)
+val under_control : vernac_control -> vernac_expr
+
+val has_Fail : vernac_control -> bool
+
+val is_navigation_vernac : vernac_control -> bool
+val is_deep_navigation_vernac : vernac_control -> bool
+val is_reset : vernac_control -> bool
+val is_debug : vernac_control -> bool
+val is_undo : vernac_control -> bool
+
diff --git a/vernac/vernacstate.ml b/vernac/vernacstate.ml
index eb1359d52..4980333b5 100644
--- a/vernac/vernacstate.ml
+++ b/vernac/vernacstate.ml
@@ -8,22 +8,34 @@
type t = {
system : States.state; (* summary + libstack *)
- proof : Proof_global.state; (* proof state *)
+ proof : Proof_global.t; (* proof state *)
shallow : bool (* is the state trimmed down (libstack) *)
}
-let s_cache = ref (States.freeze ~marshallable:`No)
-let s_proof = ref (Proof_global.freeze ~marshallable:`No)
+let s_cache = ref None
+let s_proof = ref None
let invalidate_cache () =
- s_cache := Obj.magic 0;
- s_proof := Obj.magic 0
+ s_cache := None;
+ s_proof := None
+
+let update_cache rf v =
+ rf := Some v; v
+
+let do_if_not_cached rf f v =
+ match !rf with
+ | None ->
+ rf := Some v; f v
+ | Some vc when vc != v ->
+ rf := Some v; f v
+ | Some _ ->
+ ()
let freeze_interp_state marshallable =
- { system = (s_cache := States.freeze ~marshallable; !s_cache);
- proof = (s_proof := Proof_global.freeze ~marshallable; !s_proof);
+ { system = update_cache s_cache (States.freeze ~marshallable);
+ proof = update_cache s_proof (Proof_global.freeze ~marshallable);
shallow = marshallable = `Shallow }
let unfreeze_interp_state { system; proof } =
- if (!s_cache != system) then (s_cache := system; States.unfreeze system);
- if (!s_proof != proof) then (s_proof := proof; Proof_global.unfreeze proof)
+ do_if_not_cached s_cache States.unfreeze system;
+ do_if_not_cached s_proof Proof_global.unfreeze proof
diff --git a/vernac/vernacstate.mli b/vernac/vernacstate.mli
index bcfa49aa3..3ed27ddb7 100644
--- a/vernac/vernacstate.mli
+++ b/vernac/vernacstate.mli
@@ -8,7 +8,7 @@
type t = {
system : States.state; (* summary + libstack *)
- proof : Proof_global.state; (* proof state *)
+ proof : Proof_global.t; (* proof state *)
shallow : bool (* is the state trimmed down (libstack) *)
}