summaryrefslogtreecommitdiff
path: root/toplevel
diff options
context:
space:
mode:
Diffstat (limited to 'toplevel')
-rw-r--r--toplevel/assumptions.ml126
-rw-r--r--toplevel/assumptions.mli3
-rw-r--r--toplevel/auto_ind_decl.ml125
-rw-r--r--toplevel/auto_ind_decl.mli2
-rw-r--r--toplevel/class.ml20
-rw-r--r--toplevel/classes.ml129
-rw-r--r--toplevel/classes.mli17
-rw-r--r--toplevel/command.ml281
-rw-r--r--toplevel/command.mli4
-rw-r--r--toplevel/coqinit.ml38
-rw-r--r--toplevel/coqloop.ml116
-rw-r--r--toplevel/coqloop.mli2
-rw-r--r--toplevel/coqtop.ml183
-rw-r--r--toplevel/discharge.ml25
-rw-r--r--toplevel/discharge.mli3
-rw-r--r--toplevel/explainErr.ml (renamed from toplevel/cerrors.ml)49
-rw-r--r--toplevel/explainErr.mli (renamed from toplevel/cerrors.mli)6
-rw-r--r--toplevel/g_obligations.ml4135
-rw-r--r--toplevel/himsg.ml173
-rw-r--r--toplevel/himsg.mli3
-rw-r--r--toplevel/ind_tables.ml2
-rw-r--r--toplevel/indschemes.ml57
-rw-r--r--toplevel/locality.ml19
-rw-r--r--toplevel/metasyntax.ml416
-rw-r--r--toplevel/metasyntax.mli11
-rw-r--r--toplevel/mltop.ml50
-rw-r--r--toplevel/mltop.mli4
-rw-r--r--toplevel/obligations.ml440
-rw-r--r--toplevel/obligations.mli14
-rw-r--r--toplevel/record.ml175
-rw-r--r--toplevel/record.mli13
-rw-r--r--toplevel/search.ml148
-rw-r--r--toplevel/search.mli21
-rw-r--r--toplevel/toplevel.mllib2
-rw-r--r--toplevel/usage.ml19
-rw-r--r--toplevel/usage.mli1
-rw-r--r--toplevel/vernac.ml287
-rw-r--r--toplevel/vernac.mli15
-rw-r--r--toplevel/vernacentries.ml861
-rw-r--r--toplevel/vernacentries.mli5
-rw-r--r--toplevel/vernacinterp.ml13
41 files changed, 2061 insertions, 1952 deletions
diff --git a/toplevel/assumptions.ml b/toplevel/assumptions.ml
index 2a3e6536..45c539e2 100644
--- a/toplevel/assumptions.ml
+++ b/toplevel/assumptions.ml
@@ -15,7 +15,7 @@
Module-traversing code: Pierre Letouzey *)
open Pp
-open Errors
+open CErrors
open Util
open Names
open Term
@@ -23,6 +23,7 @@ open Declarations
open Mod_subst
open Globnames
open Printer
+open Context.Named.Declaration
(** For a constant c in a module sealed by an interface (M:T and
not M<:T), [Global.lookup_constant] may return a [constant_body]
@@ -141,22 +142,20 @@ let label_of = function
| ConstructRef ((kn,_),_) -> pi3 (repr_mind kn)
| VarRef id -> Label.of_id id
-let push (r : Context.rel_declaration) (ctx : Context.rel_context) = r :: ctx
-
let rec traverse current ctx accu t = match kind_of_term t with
| Var id ->
- let body () = match Global.lookup_named id with (_, body, _) -> body in
+ let body () = Global.lookup_named id |> get_value in
traverse_object accu body (VarRef id)
| Const (kn, _) ->
let body () = Global.body_of_constant_body (lookup_constant kn) in
traverse_object accu body (ConstRef kn)
-| Ind (ind, _) ->
- traverse_object accu (fun () -> None) (IndRef ind)
-| Construct (cst, _) ->
- traverse_object accu (fun () -> None) (ConstructRef cst)
+| Ind ((mind, _) as ind, _) ->
+ traverse_inductive accu mind (IndRef ind)
+| Construct (((mind, _), _) as cst, _) ->
+ traverse_inductive accu mind (ConstructRef cst)
| Meta _ | Evar _ -> assert false
| Case (_,oty,c,[||]) ->
- (* non dependent match on an inductive with no constructors *)
+ (* non dependent match on an inductive with no constructors *)
begin match Constr.(kind oty, kind c) with
| Lambda(_,_,oty), Const (kn, _)
when Vars.noccurn 1 oty &&
@@ -165,9 +164,11 @@ let rec traverse current ctx accu t = match kind_of_term t with
traverse_object
~inhabits:(current,ctx,Vars.subst1 mkProp oty) accu body (ConstRef kn)
| _ ->
- Termops.fold_constr_with_full_binders push (traverse current) ctx accu t
+ Termops.fold_constr_with_full_binders
+ Context.Rel.add (traverse current) ctx accu t
end
-| _ -> Termops.fold_constr_with_full_binders push (traverse current) ctx accu t
+| _ -> Termops.fold_constr_with_full_binders
+ Context.Rel.add (traverse current) ctx accu t
and traverse_object ?inhabits (curr, data, ax2ty) body obj =
let data, ax2ty =
@@ -185,14 +186,87 @@ and traverse_object ?inhabits (curr, data, ax2ty) body obj =
| Some body ->
if already_in then data, ax2ty else
let contents,data,ax2ty =
- traverse (label_of obj) [] (Refset_env.empty,data,ax2ty) body in
+ traverse (label_of obj) Context.Rel.empty
+ (Refset_env.empty,data,ax2ty) body in
Refmap_env.add obj contents data, ax2ty
in
(Refset_env.add obj curr, data, ax2ty)
+(** Collects the references occurring in the declaration of mutual inductive
+ definitions. All the constructors and names of a mutual inductive
+ definition share exactly the same dependencies. Also, there is no explicit
+ dependency between mutually defined inductives and constructors. *)
+and traverse_inductive (curr, data, ax2ty) mind obj =
+ let firstind_ref = (IndRef (mind, 0)) in
+ let label = label_of obj in
+ let data, ax2ty =
+ (* Invariant : I_0 \in data iff I_i \in data iff c_ij \in data
+ where I_0, I_1, ... are in the same mutual definition and c_ij
+ are all their constructors. *)
+ if Refmap_env.mem firstind_ref data then data, ax2ty else
+ let mib = Global.lookup_mind mind in
+ (* Collects references of parameters *)
+ let param_ctx = mib.mind_params_ctxt in
+ let nparam = List.length param_ctx in
+ let accu =
+ traverse_context label Context.Rel.empty
+ (Refset_env.empty, data, ax2ty) param_ctx
+ in
+ (* Build the context of all arities *)
+ let arities_ctx =
+ let global_env = Global.env () in
+ Array.fold_left (fun accu oib ->
+ let pspecif = Univ.in_punivs (mib, oib) in
+ let ind_type = Inductive.type_of_inductive global_env pspecif in
+ let ind_name = Name oib.mind_typename in
+ Context.Rel.add (Context.Rel.Declaration.LocalAssum (ind_name, ind_type)) accu)
+ Context.Rel.empty mib.mind_packets
+ in
+ (* For each inductive, collects references in their arity and in the type
+ of constructors*)
+ let (contents, data, ax2ty) = Array.fold_left (fun accu oib ->
+ let arity_wo_param =
+ List.rev (List.skipn nparam (List.rev oib.mind_arity_ctxt))
+ in
+ let accu =
+ traverse_context
+ label param_ctx accu arity_wo_param
+ in
+ Array.fold_left (fun accu cst_typ ->
+ let param_ctx, cst_typ_wo_param = Term.decompose_prod_n_assum nparam cst_typ in
+ let ctx = Context.(Rel.fold_outside Context.Rel.add ~init:arities_ctx param_ctx) in
+ traverse label ctx accu cst_typ_wo_param)
+ accu oib.mind_user_lc)
+ accu mib.mind_packets
+ in
+ (* Maps all these dependencies to inductives and constructors*)
+ let data = Array.fold_left_i (fun n data oib ->
+ let ind = (mind, n) in
+ let data = Refmap_env.add (IndRef ind) contents data in
+ Array.fold_left_i (fun k data _ ->
+ Refmap_env.add (ConstructRef (ind, k+1)) contents data
+ ) data oib.mind_consnames) data mib.mind_packets
+ in
+ data, ax2ty
+ in
+ (Refset_env.add obj curr, data, ax2ty)
+
+(** Collects references in a rel_context. *)
+and traverse_context current ctx accu ctxt =
+ snd (Context.Rel.fold_outside (fun decl (ctx, accu) ->
+ match decl with
+ | Context.Rel.Declaration.LocalDef (_,c,t) ->
+ let accu = traverse current ctx (traverse current ctx accu t) c in
+ let ctx = Context.Rel.add decl ctx in
+ ctx, accu
+ | Context.Rel.Declaration.LocalAssum (_,t) ->
+ let accu = traverse current ctx accu t in
+ let ctx = Context.Rel.add decl ctx in
+ ctx, accu) ctxt ~init:(ctx, accu))
+
let traverse current t =
let () = modcache := MPmap.empty in
- traverse current [] (Refset_env.empty, Refmap_env.empty, Refmap_env.empty) t
+ traverse current Context.Rel.empty (Refset_env.empty, Refmap_env.empty, Refmap_env.empty) t
(** Hopefully bullet-proof function to recover the type of a constant. It just
ignores all the universe stuff. There are many issues that can arise when
@@ -208,15 +282,23 @@ let assumptions ?(add_opaque=false) ?(add_transparent=false) st gr t =
let (_, graph, ax2ty) = traverse (label_of gr) t in
let fold obj _ accu = match obj with
| VarRef id ->
- let (_, body, t) = Global.lookup_named id in
- if Option.is_empty body then ContextObjectMap.add (Variable id) t accu
+ let decl = Global.lookup_named id in
+ if is_local_assum decl then
+ let t = Context.Named.Declaration.get_type decl in
+ ContextObjectMap.add (Variable id) t accu
else accu
| ConstRef kn ->
- let cb = lookup_constant kn in
- if not (Declareops.constant_has_body cb) then
+ let cb = lookup_constant kn in
+ let accu =
+ if cb.const_typing_flags.check_guarded then accu
+ else
+ let l = try Refmap_env.find obj ax2ty with Not_found -> [] in
+ ContextObjectMap.add (Axiom (Guarded kn, l)) Constr.mkProp accu
+ in
+ if not (Declareops.constant_has_body cb) || not cb.const_typing_flags.check_universes then
let t = type_of_constant cb in
let l = try Refmap_env.find obj ax2ty with Not_found -> [] in
- ContextObjectMap.add (Axiom (kn,l)) t accu
+ ContextObjectMap.add (Axiom (Constant kn,l)) t accu
else if add_opaque && (Declareops.is_opaque cb || not (Cpred.mem kn knst)) then
let t = type_of_constant cb in
ContextObjectMap.add (Opaque kn) t accu
@@ -225,6 +307,12 @@ let assumptions ?(add_opaque=false) ?(add_transparent=false) st gr t =
ContextObjectMap.add (Transparent kn) t accu
else
accu
- | IndRef _ | ConstructRef _ -> accu
+ | IndRef (m,_) | ConstructRef ((m,_),_) ->
+ let mind = Global.lookup_mind m in
+ if mind.mind_typing_flags.check_guarded then
+ accu
+ else
+ let l = try Refmap_env.find obj ax2ty with Not_found -> [] in
+ ContextObjectMap.add (Axiom (Positive m, l)) Constr.mkProp accu
in
Refmap_env.fold fold graph ContextObjectMap.empty
diff --git a/toplevel/assumptions.mli b/toplevel/assumptions.mli
index 666218fe..07267578 100644
--- a/toplevel/assumptions.mli
+++ b/toplevel/assumptions.mli
@@ -6,7 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Util
open Names
open Term
open Globnames
@@ -22,7 +21,7 @@ open Printer
val traverse :
Label.t -> constr ->
(Refset_env.t * Refset_env.t Refmap_env.t *
- (label * Context.rel_context * types) list Refmap_env.t)
+ (label * Context.Rel.t * types) list Refmap_env.t)
(** Collects all the assumptions (optionally including opaque definitions)
on which a term relies (together with their type). The above warning of
diff --git a/toplevel/auto_ind_decl.ml b/toplevel/auto_ind_decl.ml
index b3144fa9..0561fc4b 100644
--- a/toplevel/auto_ind_decl.ml
+++ b/toplevel/auto_ind_decl.ml
@@ -10,7 +10,7 @@
decidable equality, created by Vincent Siles, Oct 2007 *)
open Tacmach
-open Errors
+open CErrors
open Util
open Pp
open Term
@@ -25,6 +25,7 @@ open Tactics
open Ind_tables
open Misctypes
open Proofview.Notations
+open Context.Rel.Declaration
let out_punivs = Univ.out_punivs
@@ -53,7 +54,7 @@ exception EqUnknown of string
exception UndefinedCst of string
exception InductiveWithProduct
exception InductiveWithSort
-exception ParameterWithoutEquality of constant
+exception ParameterWithoutEquality of global_reference
exception NonSingletonProp of inductive
exception DecidabilityMutualNotSupported
@@ -85,7 +86,7 @@ let destruct_on c = destruct false None c None None
let destruct_on_using c id =
destruct false None c
- (Some (dl,[[dl,IntroNaming IntroAnonymous];
+ (Some (dl,IntroOrPattern [[dl,IntroNaming IntroAnonymous];
[dl,IntroNaming (IntroIdentifier id)]]))
None
@@ -102,12 +103,12 @@ let mkFullInd (ind,u) n =
context_chop (nparams-nparrec) mib.mind_params_ctxt in
if nparrec > 0
then mkApp (mkIndU (ind,u),
- Array.of_list(extended_rel_list (nparrec+n) lnamesparrec))
+ Array.of_list(Context.Rel.to_extended_list (nparrec+n) lnamesparrec))
else mkIndU (ind,u)
let check_bool_is_defined () =
try let _ = Global.type_of_global_unsafe Coqlib.glob_bool in ()
- with e when Errors.noncritical e -> raise (UndefinedCst "bool")
+ with e when CErrors.noncritical e -> raise (UndefinedCst "bool")
let beq_scheme_kind_aux = ref (fun _ -> failwith "Undefined")
@@ -137,7 +138,7 @@ let build_beq_scheme mode kn =
| Name s -> Id.of_string ("eq_"^(Id.to_string s))
| Anonymous -> Id.of_string "eq_A"
in
- let ext_rel_list = extended_rel_list 0 lnamesparrec in
+ let ext_rel_list = Context.Rel.to_extended_list 0 lnamesparrec in
let lift_cnt = ref 0 in
let eqs_typ = List.map (fun aa ->
let a = lift !lift_cnt aa in
@@ -146,17 +147,17 @@ let build_beq_scheme mode kn =
) ext_rel_list in
let eq_input = List.fold_left2
- ( fun a b (n,_,_) -> (* mkLambda(n,b,a) ) *)
+ ( fun a b decl -> (* mkLambda(n,b,a) ) *)
(* here I leave the Naming thingy so that the type of
the function is more readable for the user *)
- mkNamedLambda (eqName n) b a )
+ mkNamedLambda (eqName (get_name decl)) b a )
c (List.rev eqs_typ) lnamesparrec
in
- List.fold_left (fun a (n,_,t) ->(* mkLambda(n,t,a)) eq_input rel_list *)
+ List.fold_left (fun a decl ->(* mkLambda(n,t,a)) eq_input rel_list *)
(* Same here , hoping the auto renaming will do something good ;) *)
mkNamedLambda
- (match n with Name s -> s | Anonymous -> Id.of_string "A")
- t a) eq_input lnamesparrec
+ (match get_name decl with Name s -> s | Anonymous -> Id.of_string "A")
+ (get_type decl) a) eq_input lnamesparrec
in
let make_one_eq cur =
let u = Univ.Instance.empty in
@@ -180,7 +181,13 @@ let build_beq_scheme mode kn =
let (c,a) = Reductionops.whd_betaiota_stack Evd.empty c in
match kind_of_term c with
| Rel x -> mkRel (x-nlist+ndx), Safe_typing.empty_private_constants
- | Var x -> mkVar (id_of_string ("eq_"^(string_of_id x))), Safe_typing.empty_private_constants
+ | Var x ->
+ let eid = id_of_string ("eq_"^(string_of_id x)) in
+ let () =
+ try ignore (Environ.lookup_named eid env)
+ with Not_found -> raise (ParameterWithoutEquality (VarRef x))
+ in
+ mkVar eid, Safe_typing.empty_private_constants
| Cast (x,_,_) -> aux (applist (x,a))
| App _ -> assert false
| Ind ((kn',i as ind'),u) (*FIXME: universes *) ->
@@ -208,7 +215,7 @@ let build_beq_scheme mode kn =
| LetIn _ -> raise (EqUnknown "LetIn")
| Const kn ->
(match Environ.constant_opt_value_in env kn with
- | None -> raise (ParameterWithoutEquality (fst kn))
+ | None -> raise (ParameterWithoutEquality (ConstRef (fst kn)))
| Some c -> aux (applist (c,a)))
| Proj _ -> raise (EqUnknown "Proj")
| Construct _ -> raise (EqUnknown "Construct")
@@ -233,7 +240,7 @@ let build_beq_scheme mode kn =
Cn => match Y with ... end |] part *)
let ci = make_case_info env (fst ind) MatchStyle in
let constrs n = get_constructors env (make_ind_family (ind,
- extended_rel_list (n+nb_ind-1) mib.mind_params_ctxt)) in
+ Context.Rel.to_extended_list (n+nb_ind-1) mib.mind_params_ctxt)) in
let constrsi = constrs (3+nparrec) in
let n = Array.length constrsi in
let ar = Array.make n (Lazy.force ff) in
@@ -248,7 +255,7 @@ let build_beq_scheme mode kn =
| 0 -> Lazy.force tt
| _ -> let eqs = Array.make nb_cstr_args (Lazy.force tt) in
for ndx = 0 to nb_cstr_args-1 do
- let _,_,cc = List.nth constrsi.(i).cs_args ndx in
+ let cc = get_type (List.nth constrsi.(i).cs_args ndx) in
let eqA, eff' = compute_A_equality rel_list
nparrec
(nparrec+3+2*nb_cstr_args)
@@ -267,14 +274,14 @@ let build_beq_scheme mode kn =
(Array.sub eqs 1 (nb_cstr_args - 1))
)
in
- (List.fold_left (fun a (p,q,r) -> mkLambda (p,r,a)) cc
+ (List.fold_left (fun a decl -> mkLambda (get_name decl, get_type decl, a)) cc
(constrsj.(j).cs_args)
)
- else ar2.(j) <- (List.fold_left (fun a (p,q,r) ->
- mkLambda (p,r,a)) (Lazy.force ff) (constrsj.(j).cs_args) )
+ else ar2.(j) <- (List.fold_left (fun a decl ->
+ mkLambda (get_name decl, get_type decl, a)) (Lazy.force ff) (constrsj.(j).cs_args) )
done;
- ar.(i) <- (List.fold_left (fun a (p,q,r) -> mkLambda (p,r,a))
+ ar.(i) <- (List.fold_left (fun a decl -> mkLambda (get_name decl, get_type decl, a))
(mkCase (ci,do_predicate rel_list nb_cstr_args,
mkVar (Id.of_string "Y") ,ar2))
(constrsi.(i).cs_args))
@@ -343,7 +350,7 @@ let do_replace_lb mode lb_scheme_key aavoid narg p q =
(str "Var " ++ pr_id s ++ str " seems unknown.")
)
in mkVar (find 1)
- with e when Errors.noncritical e ->
+ with e when CErrors.noncritical e ->
(* if this happen then the args have to be already declared as a
Parameter*)
(
@@ -354,7 +361,7 @@ let do_replace_lb mode lb_scheme_key aavoid narg p q =
)))
)
in
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let type_of_pq = Tacmach.New.of_old (fun gl -> pf_unsafe_type_of gl p) gl in
let u,v = destruct_ind type_of_pq
in let lb_type_of_p =
@@ -384,7 +391,7 @@ let do_replace_lb mode lb_scheme_key aavoid narg p q =
Tacticals.New.tclTHENLIST [
Proofview.tclEFFECTS eff;
Equality.replace p q ; apply app ; Auto.default_auto]
- end
+ end }
(* used in the bool -> leib side *)
let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt =
@@ -401,7 +408,7 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt =
(str "Var " ++ pr_id s ++ str " seems unknown.")
)
in mkVar (find 1)
- with e when Errors.noncritical e ->
+ with e when CErrors.noncritical e ->
(* if this happen then the args have to be already declared as a
Parameter*)
(
@@ -416,13 +423,13 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt =
let rec aux l1 l2 =
match (l1,l2) with
| (t1::q1,t2::q2) ->
- Proofview.Goal.enter begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let tt1 = Tacmach.New.pf_unsafe_type_of gl t1 in
if eq_constr t1 t2 then aux q1 q2
else (
let u,v = try destruct_ind tt1
(* trick so that the good sequence is returned*)
- with e when Errors.noncritical e -> indu,[||]
+ with e when CErrors.noncritical e -> indu,[||]
in if eq_ind (fst u) ind
then Tacticals.New.tclTHENLIST [Equality.replace t1 t2; Auto.default_auto ; aux q1 q2 ]
else (
@@ -457,7 +464,7 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt =
aux q1 q2 ]
)
)
- end
+ end }
| ([],[]) -> Proofview.tclUNIT ()
| _ -> Tacticals.New.tclZEROMSG (str "Both side of the equality must have the same arity.")
in
@@ -487,8 +494,8 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt =
create, from a list of ids [i1,i2,...,in] the list
[(in,eq_in,in_bl,in_al),,...,(i1,eq_i1,i1_bl_i1_al )]
*)
-let list_id l = List.fold_left ( fun a (n,_,t) -> let s' =
- match n with
+let list_id l = List.fold_left ( fun a decl -> let s' =
+ match get_name decl with
Name s -> Id.to_string s
| Anonymous -> "A" in
(Id.of_string s',Id.of_string ("eq_"^s'),
@@ -535,9 +542,9 @@ let compute_bl_goal ind lnamesparrec nparrec =
let eq_input = List.fold_left2 ( fun a (s,seq,_,_) b ->
mkNamedProd seq b a
) bl_input (List.rev list_id) (List.rev eqs_typ) in
- List.fold_left (fun a (n,_,t) -> mkNamedProd
- (match n with Name s -> s | Anonymous -> Id.of_string "A")
- t a) eq_input lnamesparrec
+ List.fold_left (fun a decl -> mkNamedProd
+ (match get_name decl with Name s -> s | Anonymous -> Id.of_string "A")
+ (get_type decl) a) eq_input lnamesparrec
in
let n = Id.of_string "x" and
m = Id.of_string "y" in
@@ -564,7 +571,7 @@ let compute_bl_tact mode bl_scheme_key ind lnamesparrec nparrec =
avoid := fresh::(!avoid); fresh
end gl
in
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let fresh_first_intros = List.map (fun id -> fresh_id id gl) first_intros in
let freshn = fresh_id (Id.of_string "x") gl in
let freshm = fresh_id (Id.of_string "y") gl in
@@ -580,25 +587,25 @@ let compute_bl_tact mode bl_scheme_key ind lnamesparrec nparrec =
Tacticals.New.tclTRY (
Tacticals.New.tclORELSE reflexivity (Equality.discr_tac false None)
);
- Proofview.V82.tactic (simpl_in_hyp (freshz,Locus.InHyp));
+ simpl_in_hyp (freshz,Locus.InHyp);
(*
repeat ( apply andb_prop in z;let z1:= fresh "Z" in destruct z as [z1 z]).
*)
Tacticals.New.tclREPEAT (
Tacticals.New.tclTHENLIST [
Simple.apply_in freshz (andb_prop());
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let fresht = fresh_id (Id.of_string "Z") gl in
- (destruct_on_as (mkVar freshz)
- [[dl,IntroNaming (IntroIdentifier fresht);
+ destruct_on_as (mkVar freshz)
+ (IntroOrPattern [[dl,IntroNaming (IntroIdentifier fresht);
dl,IntroNaming (IntroIdentifier freshz)]])
- end
+ end }
]);
(*
Ci a1 ... an = Ci b1 ... bn
replace bi with ai; auto || replace bi with ai by apply typeofbi_prod ; auto
*)
- Proofview.Goal.nf_enter begin fun gls ->
+ Proofview.Goal.nf_enter { enter = begin fun gls ->
let gl = Proofview.Goal.concl gls in
match (kind_of_term gl) with
| App (c,ca) -> (
@@ -617,10 +624,10 @@ repeat ( apply andb_prop in z;let z1:= fresh "Z" in destruct z as [z1 z]).
| _ -> Tacticals.New.tclZEROMSG (str" Failure while solving Boolean->Leibniz.")
)
| _ -> Tacticals.New.tclZEROMSG (str "Failure while solving Boolean->Leibniz.")
- end
+ end }
]
- end
+ end }
let bl_scheme_kind_aux = ref (fun _ -> failwith "Undefined")
@@ -678,9 +685,9 @@ let compute_lb_goal ind lnamesparrec nparrec =
let eq_input = List.fold_left2 ( fun a (s,seq,_,_) b ->
mkNamedProd seq b a
) lb_input (List.rev list_id) (List.rev eqs_typ) in
- List.fold_left (fun a (n,_,t) -> mkNamedProd
- (match n with Name s -> s | Anonymous -> Id.of_string "A")
- t a) eq_input lnamesparrec
+ List.fold_left (fun a decl -> mkNamedProd
+ (match (get_name decl) with Name s -> s | Anonymous -> Id.of_string "A")
+ (get_type decl) a) eq_input lnamesparrec
in
let n = Id.of_string "x" and
m = Id.of_string "y" in
@@ -707,7 +714,7 @@ let compute_lb_tact mode lb_scheme_key ind lnamesparrec nparrec =
avoid := fresh::(!avoid); fresh
end gl
in
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let fresh_first_intros = List.map (fun id -> fresh_id id gl) first_intros in
let freshn = fresh_id (Id.of_string "x") gl in
let freshm = fresh_id (Id.of_string "y") gl in
@@ -724,13 +731,13 @@ let compute_lb_tact mode lb_scheme_key ind lnamesparrec nparrec =
Tacticals.New.tclORELSE reflexivity (Equality.discr_tac false None)
);
Equality.inj None false None (mkVar freshz,NoBindings);
- intros; (Proofview.V82.tactic simpl_in_concl);
+ intros; simpl_in_concl;
Auto.default_auto;
Tacticals.New.tclREPEAT (
Tacticals.New.tclTHENLIST [apply (andb_true_intro());
simplest_split ;Auto.default_auto ]
);
- Proofview.Goal.nf_enter begin fun gls ->
+ Proofview.Goal.nf_enter { enter = begin fun gls ->
let gl = Proofview.Goal.concl gls in
(* assume the goal to be eq (eq_type ...) = true *)
match (kind_of_term gl) with
@@ -746,9 +753,9 @@ let compute_lb_tact mode lb_scheme_key ind lnamesparrec nparrec =
)
| _ ->
Tacticals.New.tclZEROMSG (str "Failure while solving Leibniz->Boolean.")
- end
+ end }
]
- end
+ end }
let lb_scheme_kind_aux = ref (fun () -> failwith "Undefined")
@@ -779,7 +786,7 @@ let _ = lb_scheme_kind_aux := fun () -> lb_scheme_kind
let check_not_is_defined () =
try ignore (Coqlib.build_coq_not ())
- with e when Errors.noncritical e -> raise (UndefinedCst "not")
+ with e when CErrors.noncritical e -> raise (UndefinedCst "not")
(* {n=m}+{n<>m} part *)
let compute_dec_goal ind lnamesparrec nparrec =
@@ -819,9 +826,9 @@ let compute_dec_goal ind lnamesparrec nparrec =
let eq_input = List.fold_left2 ( fun a (s,seq,_,_) b ->
mkNamedProd seq b a
) bl_input (List.rev list_id) (List.rev eqs_typ) in
- List.fold_left (fun a (n,_,t) -> mkNamedProd
- (match n with Name s -> s | Anonymous -> Id.of_string "A")
- t a) eq_input lnamesparrec
+ List.fold_left (fun a decl -> mkNamedProd
+ (match get_name decl with Name s -> s | Anonymous -> Id.of_string "A")
+ (get_type decl) a) eq_input lnamesparrec
in
let n = Id.of_string "x" and
m = Id.of_string "y" in
@@ -854,7 +861,7 @@ let compute_dec_tact ind lnamesparrec nparrec =
avoid := fresh::(!avoid); fresh
end gl
in
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let fresh_first_intros = List.map (fun id -> fresh_id id gl) first_intros in
let freshn = fresh_id (Id.of_string "x") gl in
let freshm = fresh_id (Id.of_string "y") gl in
@@ -885,7 +892,7 @@ let compute_dec_tact ind lnamesparrec nparrec =
)
(Tacticals.New.tclTHEN (destruct_on eqbnm) Auto.default_auto);
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let freshH2 = fresh_id (Id.of_string "H") gl in
Tacticals.New.tclTHENS (destruct_on_using (mkVar freshH) freshH2) [
(* left *)
@@ -897,11 +904,11 @@ let compute_dec_tact ind lnamesparrec nparrec =
;
(*right *)
- Proofview.Goal.nf_enter begin fun gl ->
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
let freshH3 = fresh_id (Id.of_string "H") gl in
Tacticals.New.tclTHENLIST [
simplest_right ;
- Proofview.V82.tactic (unfold_constr (Lazy.force Coqlib.coq_not_ref));
+ unfold_constr (Lazy.force Coqlib.coq_not_ref);
intro;
Equality.subst_all ();
assert_by (Name freshH3)
@@ -919,11 +926,11 @@ let compute_dec_tact ind lnamesparrec nparrec =
true;
Equality.discr_tac false None
]
- end
+ end }
]
- end
+ end }
]
- end
+ end }
let make_eq_decidability mode mind =
let mib = Global.lookup_mind mind in
diff --git a/toplevel/auto_ind_decl.mli b/toplevel/auto_ind_decl.mli
index b6c66a1e..fa5c6148 100644
--- a/toplevel/auto_ind_decl.mli
+++ b/toplevel/auto_ind_decl.mli
@@ -21,7 +21,7 @@ exception EqUnknown of string
exception UndefinedCst of string
exception InductiveWithProduct
exception InductiveWithSort
-exception ParameterWithoutEquality of constant
+exception ParameterWithoutEquality of Globnames.global_reference
exception NonSingletonProp of inductive
exception DecidabilityMutualNotSupported
diff --git a/toplevel/class.ml b/toplevel/class.ml
index 3d6d567c..6d53ec9d 100644
--- a/toplevel/class.ml
+++ b/toplevel/class.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Errors
+open CErrors
open Util
open Pp
open Names
@@ -32,7 +32,6 @@ type coercion_error_kind =
| NotAFunction
| NoSource of cl_typ option
| ForbiddenSourceClass of cl_typ
- | NotUniform
| NoTarget
| WrongTarget of cl_typ * cl_typ
| NotAClass of global_reference
@@ -51,9 +50,6 @@ let explain_coercion_error g = function
(str ": cannot find the source class of " ++ Printer.pr_global g)
| ForbiddenSourceClass cl ->
pr_class cl ++ str " cannot be a source class"
- | NotUniform ->
- (Printer.pr_global g ++
- str" does not respect the uniform inheritance condition");
| NoTarget ->
(str"Cannot find the target class")
| WrongTarget (clt,cl) ->
@@ -197,13 +193,13 @@ let build_id_coercion idf_opt source poly =
let val_f =
it_mkLambda_or_LetIn
(mkLambda (Name Namegen.default_dependent_ident,
- applistc vs (extended_rel_list 0 lams),
+ applistc vs (Context.Rel.to_extended_list 0 lams),
mkRel 1))
lams
in
let typ_f =
it_mkProd_wo_LetIn
- (mkProd (Anonymous, applistc vs (extended_rel_list 0 lams), lift 1 t))
+ (mkProd (Anonymous, applistc vs (Context.Rel.to_extended_list 0 lams), lift 1 t))
lams
in
(* juste pour verification *)
@@ -247,6 +243,12 @@ booleen "coercion identite'?"
lorque source est None alors target est None aussi.
*)
+let warn_uniform_inheritance =
+ CWarnings.create ~name:"uniform-inheritance" ~category:"typechecker"
+ (fun g ->
+ Printer.pr_global g ++
+ strbrk" does not respect the uniform inheritance condition")
+
let add_new_coercion_core coef stre poly source target isid =
check_source source;
let t = Global.type_of_global_unsafe coef in
@@ -262,7 +264,7 @@ let add_new_coercion_core coef stre poly source target isid =
in
check_source (Some cls);
if not (uniform_cond (llp-ind) lvs) then
- msg_warning (explain_coercion_error coef NotUniform);
+ warn_uniform_inheritance coef;
let clt =
try
get_target tg ind
@@ -310,7 +312,7 @@ let add_coercion_hook poly local ref =
in
let () = try_add_new_coercion ref stre poly in
let msg = pr_global_env Id.Set.empty ref ++ str " is now a coercion" in
- Flags.if_verbose msg_info msg
+ Flags.if_verbose Feedback.msg_info msg
let add_coercion_hook poly = Lemmas.mk_hook (add_coercion_hook poly)
diff --git a/toplevel/classes.ml b/toplevel/classes.ml
index 3a0b5f24..1528cbb2 100644
--- a/toplevel/classes.ml
+++ b/toplevel/classes.ml
@@ -12,7 +12,7 @@ open Term
open Vars
open Environ
open Nametab
-open Errors
+open CErrors
open Util
open Typeclasses_errors
open Typeclasses
@@ -20,6 +20,8 @@ open Libnames
open Globnames
open Constrintern
open Constrexpr
+open Sigma.Notations
+open Context.Rel.Declaration
(*i*)
open Decl_kinds
@@ -44,25 +46,34 @@ let set_typeclass_transparency c local b =
let _ =
Hook.set Typeclasses.add_instance_hint_hook
- (fun inst path local pri poly ->
+ (fun inst path local info poly ->
let inst' = match inst with IsConstr c -> Hints.IsConstr (c, Univ.ContextSet.empty)
| IsGlobal gr -> Hints.IsGlobRef gr
in
- Flags.silently (fun () ->
+ let info =
+ let open Vernacexpr in
+ { info with hint_pattern =
+ Option.map
+ (Constrintern.intern_constr_pattern (Global.env()))
+ info.hint_pattern } in
+ Flags.silently (fun () ->
Hints.add_hints local [typeclasses_db]
(Hints.HintsResolveEntry
- [pri, poly, false, Hints.PathHints path, inst'])) ());
+ [info, poly, false, Hints.PathHints path, inst'])) ());
Hook.set Typeclasses.set_typeclass_transparency_hook set_typeclass_transparency;
Hook.set Typeclasses.classes_transparent_state_hook
(fun () -> Hints.Hint_db.transparent_state (Hints.searchtable_map typeclasses_db))
-
+
+open Vernacexpr
+
(** TODO: add subinstances *)
-let existing_instance glob g pri =
+let existing_instance glob g info =
let c = global g in
+ let info = Option.default Hints.empty_hint_info info in
let instance = Global.type_of_global_unsafe c in
let _, r = decompose_prod_assum instance in
match class_of_constr r with
- | Some (_, ((tc,u), _)) -> add_instance (new_instance tc pri glob
+ | Some (_, ((tc,u), _)) -> add_instance (new_instance tc info glob
(*FIXME*) (Flags.use_polymorphic_flag ()) c)
| None -> user_err_loc (loc_of_reference g, "declare_instance",
Pp.str "Constant does not build instances of a declared type class.")
@@ -74,14 +85,14 @@ let mismatched_props env n m = mismatched_ctx_inst env Properties n m
let type_ctx_instance evars env ctx inst subst =
let rec aux (subst, instctx) l = function
- (na, b, t) :: ctx ->
- let t' = substl subst t in
+ decl :: ctx ->
+ let t' = substl subst (get_type decl) in
let c', l =
- match b with
- | None -> interp_casted_constr_evars env evars (List.hd l) t', List.tl l
- | Some b -> substl subst b, l
+ match decl with
+ | LocalAssum _ -> interp_casted_constr_evars env evars (List.hd l) t', List.tl l
+ | LocalDef (_,b,_) -> substl subst b, l
in
- let d = na, Some c', t' in
+ let d = get_name decl, Some c', t' in
aux (c' :: subst, d :: instctx) l ctx
| [] -> subst
in aux (subst, []) inst (List.rev ctx)
@@ -96,37 +107,41 @@ let id_of_class cl =
open Pp
-let instance_hook k pri global imps ?hook cst =
+let instance_hook k info global imps ?hook cst =
Impargs.maybe_declare_manual_implicits false cst ~enriching:false imps;
- Typeclasses.declare_instance pri (not global) cst;
+ Typeclasses.declare_instance (Some info) (not global) cst;
(match hook with Some h -> h cst | None -> ())
-let declare_instance_constant k pri global imps ?hook id poly uctx term termtype =
+let declare_instance_constant k info global imps ?hook id pl poly evm term termtype =
let kind = IsDefinition Instance in
- let uctx =
+ let evm =
let levels = Univ.LSet.union (Universes.universes_of_constr termtype)
(Universes.universes_of_constr term) in
- Universes.restrict_universe_context uctx levels
+ Evd.restrict_universe_context evm levels
in
+ let pl, uctx = Evd.universe_context ?names:pl evm in
let entry =
- Declare.definition_entry ~types:termtype ~poly ~univs:(Univ.ContextSet.to_context uctx) term
+ Declare.definition_entry ~types:termtype ~poly ~univs:uctx term
in
let cdecl = (DefinitionEntry entry, kind) in
let kn = Declare.declare_constant id cdecl in
Declare.definition_message id;
- instance_hook k pri global imps ?hook (ConstRef kn);
+ Universes.register_universe_binders (ConstRef kn) pl;
+ instance_hook k info global imps ?hook (ConstRef kn);
id
-let new_instance ?(abstract=false) ?(global=false) poly ctx (instid, bk, cl) props
+let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) poly ctx (instid, bk, cl) props
?(generalize=true)
?(tac:unit Proofview.tactic option) ?hook pri =
let env = Global.env() in
- let evars = ref (Evd.from_env env) in
+ let ((loc, instid), pl) = instid in
+ let uctx = Evd.make_evar_universe_context env pl in
+ let evars = ref (Evd.from_ctx uctx) in
let tclass, ids =
match bk with
- | Implicit ->
+ | Decl_kinds.Implicit ->
Implicit_quantifiers.implicit_application Id.Set.empty ~allow_partial:false
- (fun avoid (clname, (id, _, t)) ->
+ (fun avoid (clname, _) ->
match clname with
| Some (cl, b) ->
let t = CHole (Loc.ghost, None, Misctypes.IntroAnonymous, None) in
@@ -149,16 +164,17 @@ let new_instance ?(abstract=false) ?(global=false) poly ctx (instid, bk, cl) pro
let k, args = Typeclasses.dest_class_app (push_rel_context ctx'' env) c in
let cl, u = Typeclasses.typeclass_univ_instance k in
let _, args =
- List.fold_right (fun (na, b, t) (args, args') ->
- match b with
- | None -> (List.tl args, List.hd args :: args')
- | Some b -> (args, substl args' b :: args'))
+ List.fold_right (fun decl (args, args') ->
+ let open Context.Rel.Declaration in
+ match decl with
+ | LocalAssum _ -> (List.tl args, List.hd args :: args')
+ | LocalDef (_,b,_) -> (args, substl args' b :: args'))
(snd cl.cl_context) (args, [])
in
cl, u, c', ctx', ctx, len, imps, args
in
let id =
- match snd instid with
+ match instid with
Name id ->
let sp = Lib.make_path id in
if Nametab.exists_cci sp then
@@ -175,7 +191,7 @@ let new_instance ?(abstract=false) ?(global=false) poly ctx (instid, bk, cl) pro
if abstract then
begin
let subst = List.fold_left2
- (fun subst' s (_, b, _) -> if Option.is_empty b then s :: subst' else subst')
+ (fun subst' s decl -> if is_local_assum decl then s :: subst' else subst')
[] subst (snd k.cl_context)
in
let (_, ty_constr) = instance_constructor (k,u) subst in
@@ -184,17 +200,19 @@ let new_instance ?(abstract=false) ?(global=false) poly ctx (instid, bk, cl) pro
let t = it_mkProd_or_LetIn ty_constr (ctx' @ ctx) in
nf t
in
- Evarutil.check_evars env Evd.empty !evars termtype;
- let pl, ctx = Evd.universe_context !evars in
+ Pretyping.check_evars env Evd.empty !evars termtype;
+ let pl, ctx = Evd.universe_context ?names:pl !evars in
let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id
(ParameterEntry
(None,poly,(termtype,ctx),None), Decl_kinds.IsAssumption Decl_kinds.Logical)
- in instance_hook k None global imps ?hook (ConstRef cst); id
+ in
+ Universes.register_universe_binders (ConstRef cst) pl;
+ instance_hook k pri global imps ?hook (ConstRef cst); id
end
else (
let props =
match props with
- | Some (true, CRecord (loc, _, fs)) ->
+ | Some (true, CRecord (loc, fs)) ->
if List.length fs > List.length k.cl_props then
mismatched_props env' (List.map snd fs) k.cl_props;
Some (Inl fs)
@@ -217,10 +235,10 @@ let new_instance ?(abstract=false) ?(global=false) poly ctx (instid, bk, cl) pro
in
let props, rest =
List.fold_left
- (fun (props, rest) (id,b,_) ->
- if Option.is_empty b then
+ (fun (props, rest) decl ->
+ if is_local_assum decl then
try
- let is_id (id', _) = match id, get_id id' with
+ let is_id (id', _) = match get_name decl, get_id id' with
| Name id, (_, id') -> Id.equal id id'
| Anonymous, _ -> false
in
@@ -254,7 +272,7 @@ let new_instance ?(abstract=false) ?(global=false) poly ctx (instid, bk, cl) pro
None, termtype
| Some (Inl subst) ->
let subst = List.fold_left2
- (fun subst' s (_, b, _) -> if Option.is_empty b then s :: subst' else subst')
+ (fun subst' s decl -> if is_local_assum decl then s :: subst' else subst')
[] subst (k.cl_props @ snd k.cl_context)
in
let (app, ty_constr) = instance_constructor (k,u) subst in
@@ -278,20 +296,19 @@ let new_instance ?(abstract=false) ?(global=false) poly ctx (instid, bk, cl) pro
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. *)
- Evarutil.check_evars env Evd.empty evm termtype
+ Pretyping.check_evars env Evd.empty evm termtype
in
let term = Option.map nf term in
if not (Evd.has_undefined evm) && not (Option.is_empty term) then
- let ctx = Evd.universe_context_set evm in
- declare_instance_constant k pri global imps ?hook id
- poly ctx (Option.get term) termtype
- else if !refine_instance || Option.is_empty term then begin
+ declare_instance_constant k pri global imps ?hook id pl
+ poly evm (Option.get term) termtype
+ else if Flags.is_program_mode () || refine || Option.is_empty term then begin
let kind = Decl_kinds.Global, poly, Decl_kinds.DefinitionBody Decl_kinds.Instance in
if Flags.is_program_mode () then
let hook vis gr _ =
let cst = match gr with ConstRef kn -> kn | _ -> assert false in
Impargs.declare_manual_implicits false gr ~enriching:false [imps];
- Typeclasses.declare_instance pri (not global) (ConstRef cst)
+ Typeclasses.declare_instance (Some pri) (not global) (ConstRef cst)
in
let obls, constr, typ =
match term with
@@ -304,7 +321,7 @@ let new_instance ?(abstract=false) ?(global=false) poly ctx (instid, bk, cl) pro
let hook = Lemmas.mk_hook hook in
let ctx = Evd.evar_universe_context evm in
ignore (Obligations.add_definition id ?term:constr
- typ ctx ~kind:(Global,poly,Instance) ~hook obls);
+ ?pl typ ctx ~kind:(Global,poly,Instance) ~hook obls);
id
else
(Flags.silently
@@ -322,7 +339,7 @@ let new_instance ?(abstract=false) ?(global=false) poly ctx (instid, bk, cl) pro
if not (Option.is_empty term) then
let init_refine =
Tacticals.New.tclTHENLIST [
- Proofview.Refine.refine (fun evm -> evm, Option.get term);
+ Refine.refine { run = fun evm -> Sigma (Option.get term, evm, Sigma.refl) };
Proofview.Unsafe.tclNEWGOALS gls;
Tactics.New.reduce_after_refine;
]
@@ -333,14 +350,16 @@ let new_instance ?(abstract=false) ?(global=false) poly ctx (instid, bk, cl) pro
(match tac with Some tac -> ignore (Pfedit.by tac) | None -> ())) ();
id)
end
- else Errors.error "Unsolved obligations remaining.")
+ else CErrors.error "Unsolved obligations remaining.")
let named_of_rel_context l =
let acc, ctx =
List.fold_right
- (fun (na, b, t) (subst, ctx) ->
- let id = match na with Anonymous -> invalid_arg "named_of_rel_context" | Name id -> id in
- let d = (id, Option.map (substl subst) b, substl subst t) in
+ (fun decl (subst, ctx) ->
+ let id = match get_name decl with Anonymous -> invalid_arg "named_of_rel_context" | Name id -> id in
+ let d = match decl with
+ | LocalAssum (_,t) -> id, None, substl subst t
+ | LocalDef (_,b,t) -> id, Some (substl subst b), substl subst t in
(mkVar id :: subst, d :: ctx))
l ([], [])
in ctx
@@ -350,12 +369,12 @@ let context poly l =
let evars = ref (Evd.from_env env) in
let _, ((env', fullctx), impls) = interp_context_evars env evars l in
let subst = Evarutil.evd_comb0 Evarutil.nf_evars_and_universes evars in
- let fullctx = Context.map_rel_context subst fullctx in
- let ce t = Evarutil.check_evars env Evd.empty !evars t in
- let () = List.iter (fun (n, b, t) -> Option.iter ce b; ce t) fullctx in
+ let fullctx = Context.Rel.map subst fullctx in
+ let ce t = Pretyping.check_evars env Evd.empty !evars t in
+ let () = List.iter (fun decl -> Context.Rel.Declaration.iter_constr ce decl) fullctx in
let ctx =
try named_of_rel_context fullctx
- with e when Errors.noncritical e ->
+ with e when CErrors.noncritical e ->
error "Anonymous variables not allowed in contexts."
in
let uctx = ref (Evd.universe_context_set !evars) in
@@ -368,7 +387,7 @@ let context poly l =
let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id decl in
match class_of_constr t with
| Some (rels, ((tc,_), args) as _cl) ->
- add_instance (Typeclasses.new_instance tc None false (*FIXME*)
+ add_instance (Typeclasses.new_instance tc Hints.empty_hint_info false (*FIXME*)
poly (ConstRef cst));
status
(* declare_subclasses (ConstRef cst) cl *)
diff --git a/toplevel/classes.mli b/toplevel/classes.mli
index 24c51b31..d2cb788e 100644
--- a/toplevel/classes.mli
+++ b/toplevel/classes.mli
@@ -7,7 +7,6 @@
(************************************************************************)
open Names
-open Context
open Environ
open Constrexpr
open Typeclasses
@@ -15,24 +14,25 @@ open Libnames
(** Errors *)
-val mismatched_params : env -> constr_expr list -> rel_context -> 'a
+val mismatched_params : env -> constr_expr list -> Context.Rel.t -> 'a
-val mismatched_props : env -> constr_expr list -> rel_context -> 'a
+val mismatched_props : env -> constr_expr list -> Context.Rel.t -> 'a
(** Instance declaration *)
-val existing_instance : bool -> reference -> int option -> unit
-(** globality, reference, priority *)
+val existing_instance : bool -> reference -> Vernacexpr.hint_info_expr option -> unit
+(** globality, reference, optional priority and pattern information *)
val declare_instance_constant :
typeclass ->
- int option -> (** priority *)
+ Vernacexpr.hint_info_expr -> (** priority *)
bool -> (** globality *)
Impargs.manual_explicitation list -> (** implicits *)
?hook:(Globnames.global_reference -> unit) ->
Id.t -> (** name *)
+ Id.t Loc.located list option ->
bool -> (* polymorphic *)
- Univ.universe_context_set -> (* Universes *)
+ Evd.evar_map -> (* Universes *)
Constr.t -> (** body *)
Term.types -> (** type *)
Names.Id.t
@@ -40,6 +40,7 @@ val declare_instance_constant :
val new_instance :
?abstract:bool -> (** Not abstract by default. *)
?global:bool -> (** Not global by default. *)
+ ?refine:bool -> (** Allow refinement *)
Decl_kinds.polymorphic ->
local_binder list ->
typeclass_constraint ->
@@ -47,7 +48,7 @@ val new_instance :
?generalize:bool ->
?tac:unit Proofview.tactic ->
?hook:(Globnames.global_reference -> unit) ->
- int option ->
+ Vernacexpr.hint_info_expr ->
Id.t
(** Setting opacity *)
diff --git a/toplevel/command.ml b/toplevel/command.ml
index 5d2a7638..a9f2598e 100644
--- a/toplevel/command.ml
+++ b/toplevel/command.ml
@@ -7,14 +7,12 @@
(************************************************************************)
open Pp
-open Errors
+open CErrors
open Util
open Flags
open Term
open Vars
-open Context
open Termops
-open Entries
open Environ
open Redexpr
open Declare
@@ -37,17 +35,20 @@ open Evarconv
open Indschemes
open Misctypes
open Vernacexpr
+open Sigma.Notations
+open Context.Rel.Declaration
+open Entries
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 snd (f env sigma c) else
+ if Int.equal n 0 then f env sigma c else
match kind_of_term c with
| Lambda (x,t,c) ->
- mkLambda (x,t,under_binders (push_rel (x,None,t) env) sigma f (n-1) 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 (x,Some b,t) env) sigma f (n-1) c)
+ mkLetIn (x,b,t,under_binders (push_rel (LocalDef (x,b,t)) env) sigma f (n-1) c)
| _ -> assert false
let rec complete_conclusion a cs = function
@@ -72,11 +73,21 @@ let red_constant_entry n ce sigma = function
| Some red ->
let proof_out = ce.const_entry_body in
let env = Global.env () in
+ let (redfun, _) = reduction_of_red_expr env red in
+ let redfun env sigma c =
+ let sigma = Sigma.Unsafe.of_evar_map sigma in
+ let Sigma (c, _, _) = redfun.e_redfun env sigma c in
+ c
+ in
{ ce with const_entry_body = Future.chain ~greedy:true ~pure:true proof_out
- (fun ((body,ctx),eff) ->
- (under_binders env sigma
- (fst (reduction_of_red_expr env red)) n body,ctx),eff) }
-
+ (fun ((body,ctx),eff) -> (under_binders env sigma redfun n body,ctx),eff) }
+
+let warn_implicits_in_term =
+ CWarnings.create ~name:"implicits-in-term" ~category:"implicits"
+ (fun () ->
+ strbrk "Implicit arguments declaration relies on type." ++ spc () ++
+ strbrk "The term declares more implicits than the type here.")
+
let interp_definition pl bl p red_option c ctypopt =
let env = Global.env() in
let ctx = Evd.make_evar_universe_context env pl in
@@ -87,7 +98,7 @@ let interp_definition pl bl p red_option c ctypopt =
match ctypopt with
None ->
let subst = evd_comb0 Evd.nf_univ_variables evdref in
- let ctx = map_rel_context (Vars.subst_univs_constr subst) ctx in
+ let ctx = Context.Rel.map (Vars.subst_univs_constr subst) ctx in
let env_bl = push_rel_context ctx env in
let c, imps2 = interp_constr_evars_impls ~impls env_bl evdref c in
let nf,subst = Evarutil.e_nf_evars_and_universes evdref in
@@ -100,7 +111,7 @@ let interp_definition pl bl p red_option c ctypopt =
| 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 = map_rel_context (Vars.subst_univs_constr subst) ctx in
+ let ctx = Context.Rel.map (Vars.subst_univs_constr subst) ctx in
let env_bl = push_rel_context ctx env in
let c, imps2 = interp_casted_constr_evars_impls ~impls env_bl evdref c ty in
let nf, subst = Evarutil.e_nf_evars_and_universes evdref in
@@ -114,9 +125,7 @@ let interp_definition pl bl p red_option c ctypopt =
impl_eq (List.assoc_f Pervasives.(=) key impsty) va (* FIXME *)
in
if not (try List.for_all chk imps2 with Not_found -> false)
- then msg_warning
- (strbrk "Implicit arguments declaration relies on type." ++ spc () ++
- strbrk "The term declares more implicits than the type here.");
+ then warn_implicits_in_term ();
let vars = Univ.LSet.union (Universes.universes_of_constr body)
(Universes.universes_of_constr typ) in
let ctx = Evd.restrict_universe_context !evdref vars in
@@ -125,23 +134,27 @@ let interp_definition pl bl p red_option c ctypopt =
definition_entry ~types:typ ~poly:p
~univs:uctx body
in
- red_constant_entry (rel_context_length ctx) ce !evdref red_option, !evdref, pl, imps
+ red_constant_entry (Context.Rel.length ctx) ce !evdref red_option, !evdref, pl, imps
let check_definition (ce, evd, _, imps) =
check_evars_are_solved (Global.env ()) evd (Evd.empty,evd);
ce
-let get_locality id = function
+let warn_local_declaration =
+ CWarnings.create ~name:"local-declaration" ~category:"scope"
+ (fun (id,kind) ->
+ pr_id id ++ strbrk " is declared as a local " ++ str kind)
+
+let get_locality id ~kind = function
| Discharge ->
(** If a Let is defined outside a section, then we consider it as a local definition *)
- let msg = pr_id id ++ strbrk " is declared as a local definition" in
- let () = msg_warning msg in
+ warn_local_declaration (id,kind);
true
| Local -> true
| Global -> false
let declare_global_definition ident ce local k pl imps =
- let local = get_locality ident local in
+ let local = get_locality ident ~kind:"definition" local in
let kn = declare_constant ident ~local (DefinitionEntry ce, IsDefinition k) in
let gr = ConstRef kn in
let () = maybe_declare_manual_implicits false gr imps in
@@ -153,6 +166,12 @@ let declare_definition_hook = ref ignore
let set_declare_definition_hook = (:=) declare_definition_hook
let get_declare_definition_hook () = !declare_definition_hook
+let warn_definition_not_visible =
+ CWarnings.create ~name:"definition-not-visible" ~category:"implicits"
+ (fun ident ->
+ strbrk "Section definition " ++
+ pr_id ident ++ strbrk " is not visible from current goals")
+
let declare_definition ident (local, p, k) ce pl imps hook =
let fix_exn = Future.fix_exn_of ce.const_entry_body in
let () = !declare_definition_hook ce in
@@ -164,9 +183,7 @@ let declare_definition ident (local, p, k) ce pl imps hook =
let gr = VarRef ident in
let () = maybe_declare_manual_implicits false gr imps in
let () = if Pfedit.refining () then
- let msg = strbrk "Section definition " ++
- pr_id ident ++ strbrk " is not visible from current goals" in
- msg_warning msg
+ warn_definition_not_visible ident
in
gr
| Discharge | Local | Global ->
@@ -177,7 +194,9 @@ let _ = Obligations.declare_definition_ref :=
(fun i k c imps hook -> declare_definition i k c [] imps hook)
let do_definition ident k pl bl red_option c ctypopt hook =
- let (ce, evd, pl, imps as def) = interp_definition pl bl (pi2 k) red_option c ctypopt in
+ let (ce, evd, pl', imps as def) =
+ interp_definition pl bl (pi2 k) red_option c ctypopt
+ in
if Flags.is_program_mode () then
let env = Global.env () in
let (c,ctx), sideff = Future.force ce.const_entry_body in
@@ -194,9 +213,9 @@ let do_definition ident k pl bl red_option c ctypopt hook =
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 ~implicits:imps ~kind:k ~hook obls)
+ ident ~term:c cty ctx ?pl ~implicits:imps ~kind:k ~hook obls)
else let ce = check_definition def in
- ignore(declare_definition ident k ce pl imps
+ ignore(declare_definition ident k ce pl' imps
(Lemmas.mk_hook
(fun l r -> Lemmas.call_hook (fun exn -> exn) hook l r;r)))
@@ -210,7 +229,7 @@ match local with
let () = assumption_message ident in
let () =
if is_verbose () && Pfedit.refining () then
- msg_warning (str"Variable" ++ spc () ++ pr_id ident ++
+ Feedback.msg_info (str"Variable" ++ spc () ++ pr_id ident ++
strbrk " is not visible from current goals")
in
let r = VarRef ident in
@@ -219,7 +238,7 @@ match local with
(r,Univ.Instance.empty,true)
| Global | Local | Discharge ->
- let local = get_locality ident local in
+ let local = get_locality ident ~kind:"axiom" local in
let inl = match nl with
| NoInline -> None
| DefaultInline -> Some (Flags.get_inline_level())
@@ -242,10 +261,7 @@ match local with
let interp_assumption evdref env impls bl c =
let c = prod_constr_expr c bl in
- let ty, impls = interp_type_evars_impls env evdref ~impls c in
- let evd, nf = nf_evars_and_universes !evdref in
- let ctx = Evd.universe_context_set evd in
- ((nf ty, ctx), impls)
+ interp_type_evars_impls env evdref ~impls c
let declare_assumptions idl is_coe k (c,ctx) pl imps impl_is_on nl =
let refs, status, _ =
@@ -258,6 +274,7 @@ let declare_assumptions idl is_coe k (c,ctx) pl imps impl_is_on nl =
List.rev refs, status
let do_assumptions_unbound_univs (_, poly, _ as kind) nl l =
+ let open Context.Named.Declaration in
let env = Global.env () in
let evdref = ref (Evd.from_env env) in
let l =
@@ -269,26 +286,32 @@ let do_assumptions_unbound_univs (_, poly, _ as kind) nl l =
l []
else l
in
+ (* We intepret all declarations in the same evar_map, i.e. as a telescope. *)
let _,l = List.fold_map (fun (env,ienv) (is_coe,(idl,c)) ->
- let (t,ctx),imps = interp_assumption evdref env ienv [] c in
+ let t,imps = interp_assumption evdref env ienv [] c in
let env =
- push_named_context (List.map (fun (_,id) -> (id,None,t)) idl) env in
+ 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,(ctx,imps))))
+ ((env,ienv),((is_coe,idl),t,imps)))
(env,empty_internalization_env) l
in
let evd = solve_remaining_evars all_and_fail_flags env !evdref (Evd.empty,!evdref) in
+ (* The universe constraints come from the whole telescope. *)
+ let evd = Evd.nf_constraints evd in
+ let ctx = Evd.universe_context_set evd in
let l = List.map (on_pi2 (nf_evar evd)) l in
- snd (List.fold_left (fun (subst,status) ((is_coe,idl),t,(ctx,imps)) ->
+ pi2 (List.fold_left (fun (subst,status,ctx) ((is_coe,idl),t,imps) ->
let t = replace_vars subst t in
let (refs,status') = declare_assumptions idl is_coe kind (t,ctx) [] imps false nl in
let subst' = List.map2
(fun (_,id) (c,u) -> (id,Universes.constr_of_global_univ (c,u)))
idl refs
in
- (subst'@subst, status' && status)) ([],true) l)
+ (subst'@subst, status' && status,
+ (* The universe constraints are declared with the first declaration only. *)
+ Univ.ContextSet.empty)) ([],true,ctx) l)
let do_assumptions_bound_univs coe kind nl id pl c =
let env = Global.env () in
@@ -334,7 +357,7 @@ let do_assumptions kind nl l = match l with
(* 3b| Mutual inductive definitions *)
let push_types env idl tl =
- List.fold_left2 (fun env id t -> Environ.push_rel (Name id,None,t) env)
+ List.fold_left2 (fun env id t -> Environ.push_rel (LocalAssum (Name id,t)) env)
env idl tl
type structured_one_inductive_expr = {
@@ -377,8 +400,8 @@ let mk_mltype_data evdref env assums arity indname =
(is_ml_type,indname,assums)
let prepare_param = function
- | (na,None,t) -> out_name na, LocalAssum t
- | (na,Some b,_) -> out_name na, LocalDef b
+ | LocalAssum (na,t) -> out_name na, LocalAssumEntry t
+ | LocalDef (na,b,_) -> out_name na, LocalDefEntry b
(** Make the arity conclusion flexible to avoid generating an upper bound universe now,
only if the universe does not appear anywhere else.
@@ -432,12 +455,12 @@ let interp_cstrs evdref env impls mldata arity ind =
let sign_level env evd sign =
fst (List.fold_right
- (fun (_,b,t as d) (lev,env) ->
- match b with
- | Some _ -> (lev, push_rel d env)
- | None ->
- let s = destSort (Reduction.whd_betadeltaiota env
- (nf_evar evd (Retyping.get_type_of env evd t)))
+ (fun d (lev,env) ->
+ match d with
+ | LocalDef _ -> lev, push_rel d env
+ | LocalAssum _ ->
+ let s = destSort (Reduction.whd_all env
+ (nf_evar evd (Retyping.get_type_of env evd (get_type d))))
in
let u = univ_of_sort s in
(Univ.sup u lev, push_rel d env))
@@ -448,7 +471,7 @@ 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 ((Anonymous, None, concl) :: ctx)) tys
+ sign_level env evd (LocalAssum (Anonymous, concl) :: ctx)) tys
in sup_list min sorts
let is_flexible_sort evd u =
@@ -517,11 +540,9 @@ let inductive_levels env evdref poly arities inds =
in
let duu = Sorts.univ_of_sort du in
let evd =
- if not (Univ.is_small_univ duu) && Evd.check_eq evd cu duu then
- if is_flexible_sort evd duu then
- if Evd.check_leq evd Univ.type0_univ duu then
- evd
- else Evd.set_eq_sort env evd (Prop Null) du
+ 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
@@ -540,6 +561,7 @@ let check_param = function
| LocalRawDef (na, _) -> check_named na
| LocalRawAssum (nas, Default _, _) -> List.iter check_named nas
| LocalRawAssum (nas, Generalized _, _) -> ()
+| LocalPattern _ -> assert false
let interp_mutual_inductive (paramsl,indl) notations poly prv finite =
check_all_names_different indl;
@@ -554,8 +576,8 @@ let interp_mutual_inductive (paramsl,indl) notations poly prv finite =
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(fun (_,b,_) -> Option.is_empty b) ctx_params in
- let params = List.map (fun (na,_,_) -> out_name na) assums in
+ let assums = List.filter is_local_assum ctx_params in
+ let params = List.map (fun decl -> out_name (get_name decl)) assums in
(* Interpret the arities *)
let arities = List.map (interp_ind_arity env_params evdref) indl in
@@ -566,7 +588,7 @@ let interp_mutual_inductive (paramsl,indl) notations poly prv finite =
(* Compute interpretation metadatas *)
let indimpls = List.map (fun (_, _, impls) -> userimpls @
- lift_implicits (rel_context_nhyps ctx_params) impls) arities in
+ lift_implicits (Context.Rel.nhyps ctx_params) impls) arities in
let arities = List.map pi1 arities and aritypoly = List.map pi2 arities in
let impls = compute_internalization_env env0 (Inductive params) indnames fullarities indimpls in
let mldatas = List.map2 (mk_mltype_data evdref env_params params) arities indnames in
@@ -592,11 +614,11 @@ let interp_mutual_inductive (paramsl,indl) notations poly prv finite =
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 = map_rel_context nf ctx_params in
+ let ctx_params = Context.Rel.map nf ctx_params in
let evd = !evdref in
let pl, uctx = Evd.universe_context ?names:pl evd in
List.iter (check_evars env_params Evd.empty evd) arities;
- iter_rel_context (check_evars env0 Evd.empty evd) ctx_params;
+ Context.Rel.iter (check_evars env0 Evd.empty evd) ctx_params;
List.iter (fun (_,ctyps,_) ->
List.iter (check_evars env_ar_params Evd.empty evd) ctyps)
constructors;
@@ -610,7 +632,7 @@ let interp_mutual_inductive (paramsl,indl) notations poly prv finite =
mind_entry_lc = ctypes
}) indl arities aritypoly constructors in
let impls =
- let len = rel_context_nhyps ctx_params in
+ 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
@@ -622,7 +644,8 @@ let interp_mutual_inductive (paramsl,indl) notations poly prv finite =
mind_entry_inds = entries;
mind_entry_polymorphic = poly;
mind_entry_private = if prv then Some false else None;
- mind_entry_universes = uctx },
+ mind_entry_universes = uctx;
+ },
pl, impls
(* Very syntactical equality *)
@@ -698,7 +721,7 @@ let declare_mutual_inductive_with_eliminations mie pl impls =
constrimpls)
impls;
let warn_prim = match mie.mind_entry_record with Some (Some _) -> not prim | _ -> false in
- if_verbose msg_info (minductive_message warn_prim names);
+ if_verbose Feedback.msg_info (minductive_message warn_prim names);
if mie.mind_entry_private == None
then declare_default_schemes mind;
mind
@@ -716,8 +739,10 @@ let do_mutual_inductive indl poly prv finite =
(* Declare the possible notations of inductive types *)
List.iter Metasyntax.add_notation_interpretation ntns;
(* Declare the coercions *)
- List.iter (fun qid -> Class.try_add_new_coercion (locate qid) false poly) coes
-
+ List.iter (fun qid -> Class.try_add_new_coercion (locate qid) false poly) coes;
+ (* If positivity is assumed declares itself as unsafe. *)
+ if Environ.deactivated_guard (Global.env ()) then Feedback.feedback Feedback.AddedAxiom else ()
+
(* 3c| Fixpoints and co-fixpoints *)
(* An (unoptimized) function that maps preorders to partial orders...
@@ -768,20 +793,25 @@ let rec partial_order cmp = function
let non_full_mutual_message x xge y yge isfix rest =
let reason =
if Id.List.mem x yge then
- pr_id y ++ str " depends on " ++ pr_id x ++ str " but not conversely"
+ pr_id y ++ str " depends on " ++ pr_id x ++ strbrk " but not conversely"
else if Id.List.mem y xge then
- pr_id x ++ str " depends on " ++ pr_id y ++ str " but not conversely"
+ pr_id x ++ str " depends on " ++ pr_id y ++ strbrk " but not conversely"
else
- pr_id y ++ str " and " ++ pr_id x ++ str " are not mutually dependent" in
- let e = if List.is_empty rest then reason else str "e.g., " ++ reason in
+ pr_id y ++ str " and " ++ pr_id x ++ strbrk " are not mutually dependent" in
+ let e = if List.is_empty rest then reason else strbrk "e.g., " ++ reason in
let k = if isfix then "fixpoint" else "cofixpoint" in
let w =
if isfix
- then str "Well-foundedness check may fail unexpectedly." ++ fnl()
+ then strbrk "Well-foundedness check may fail unexpectedly." ++ fnl()
else mt () in
- str "Not a fully mutually defined " ++ str k ++ fnl () ++
+ strbrk "Not a fully mutually defined " ++ str k ++ fnl () ++
str "(" ++ e ++ str ")." ++ fnl () ++ w
+let warn_non_full_mutual =
+ CWarnings.create ~name:"non-full-mutual" ~category:"fixpoints"
+ (fun (x,xge,y,yge,isfix,rest) ->
+ non_full_mutual_message x xge y yge isfix rest)
+
let check_mutuality env isfix fixl =
let names = List.map fst fixl in
let preorder =
@@ -791,7 +821,7 @@ let check_mutuality env isfix fixl =
let po = partial_order Id.equal preorder in
match List.filter (function (_,Inr _) -> true | _ -> false) po with
| (x,Inr xge)::(y,Inr yge)::rest ->
- msg_warning (non_full_mutual_message x xge y yge isfix rest)
+ warn_non_full_mutual (x,xge,y,yge,isfix,rest)
| _ -> ()
type structured_fixpoint_expr = {
@@ -826,7 +856,7 @@ let declare_fix ?(opaque = false) (_,poly,_ as kind) pl ctx f ((def,_),eff) t im
declare_definition f kind ce pl imps (Lemmas.mk_hook (fun _ r -> r))
let _ = Obligations.declare_fix_ref :=
- (fun ?opaque k ctx f d t imps -> declare_fix ?opaque k [] ctx f d t imps)
+ (fun ?opaque k ctx f d t imps -> declare_fix ?opaque k [] ctx f d t imps)
let prepare_recursive_declaration fixnames fixtypes fixdefs =
let defs = List.map (subst_vars (List.rev fixnames)) fixdefs in
@@ -868,19 +898,20 @@ let well_founded = init_constant ["Init"; "Wf"] "well_founded"
let mkSubset name typ prop =
mkApp (Universes.constr_of_global (delayed_force build_sigma).typ,
[| typ; mkLambda (name, typ, prop) |])
-let sigT = Lazy.lazy_from_fun build_sigma_type
+let sigT = Lazy.from_fun build_sigma_type
let make_qref s = Qualid (Loc.ghost, qualid_of_string s)
let lt_ref = make_qref "Init.Peano.lt"
let rec telescope = function
| [] -> assert false
- | [(n, None, t)] -> t, [n, Some (mkRel 1), t], mkRel 1
- | (n, None, t) :: tl ->
+ | [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)) (n, b, t) ->
- let pred = mkLambda (n, t, ty) in
+ (fun (ty, tys, (k, constr)) decl ->
+ let t = get_type decl in
+ let pred = mkLambda (get_name decl, t, ty) in
let ty = Universes.constr_of_global (Lazy.force sigT).typ in
let intro = Universes.constr_of_global (Lazy.force sigT).intro in
let sigty = mkApp (ty, [|t; pred|]) in
@@ -889,26 +920,27 @@ let rec telescope = function
(t, [], (2, mkRel 1)) tl
in
let (last, subst) = List.fold_right2
- (fun pred (n, b, t) (prev, subst) ->
+ (fun pred decl (prev, subst) ->
+ let t = get_type decl in
let p1 = Universes.constr_of_global (Lazy.force sigT).proj1 in
let p2 = Universes.constr_of_global (Lazy.force sigT).proj2 in
let proj1 = applistc p1 [t; pred; prev] in
let proj2 = applistc p2 [t; pred; prev] in
- (lift 1 proj2, (n, Some proj1, t) :: subst))
+ (lift 1 proj2, LocalDef (get_name decl, proj1, t) :: subst))
(List.rev tys) tl (mkRel 1, [])
- in ty, ((n, Some last, t) :: subst), constr
+ in ty, (LocalDef (n, last, t) :: subst), constr
- | (n, Some b, t) :: tl -> let ty, subst, term = telescope tl in
- ty, ((n, Some b, t) :: subst), lift 1 term
+ | LocalDef (n, b, t) :: tl -> let ty, subst, term = telescope tl in
+ ty, (LocalDef (n, b, t) :: subst), lift 1 term
let nf_evar_context sigma ctx =
- List.map (fun (n, b, t) ->
- (n, Option.map (Evarutil.nf_evar sigma) b, Evarutil.nf_evar sigma t)) ctx
+ List.map (map_constr (Evarutil.nf_evar sigma)) ctx
-let build_wellfounded (recname,n,bl,arityc,body) r measure notation =
+let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
Coqlib.check_required_library ["Coq";"Program";"Wf"];
let env = Global.env() in
- let evdref = ref (Evd.from_env env) in
+ let ctx = Evd.make_evar_universe_context env pl in
+ let evdref = ref (Evd.from_ctx ctx) in
let _, ((env', binders_rel), impls) = interp_context_evars env evdref bl in
let len = List.length binders_rel in
let top_env = push_rel_context binders_rel env in
@@ -916,11 +948,10 @@ let build_wellfounded (recname,n,bl,arityc,body) r measure notation =
let full_arity = it_mkProd_or_LetIn top_arity binders_rel in
let argtyp, letbinders, make = telescope binders_rel in
let argname = Id.of_string "recarg" in
- let arg = (Name argname, None, argtyp) 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 () = check_evars_are_solved env !evdref (Evd.empty,!evdref) in
let relty = Typing.unsafe_type_of env !evdref rel in
let relargty =
let error () =
@@ -931,10 +962,10 @@ let build_wellfounded (recname,n,bl,arityc,body) r measure notation =
try
let ctx, ar = Reductionops.splay_prod_n env !evdref 2 relty in
match ctx, kind_of_term ar with
- | [(_, None, t); (_, None, u)], Sort (Prop Null)
+ | [LocalAssum (_,t); LocalAssum (_,u)], Sort (Prop Null)
when Reductionops.is_conv env !evdref t u -> t
| _, _ -> error ()
- with e when Errors.noncritical e -> error ()
+ with e when CErrors.noncritical e -> error ()
in
let measure = interp_casted_constr_evars binders_env evdref measure relargty in
let wf_rel, wf_rel_fun, measure_fn =
@@ -951,9 +982,9 @@ let build_wellfounded (recname,n,bl,arityc,body) r measure notation =
in
let wf_proof = mkApp (delayed_force well_founded, [| argtyp ; wf_rel |]) in
let argid' = Id.of_string (Id.to_string argname ^ "'") in
- let wfarg len = (Name argid', None,
- mkSubset (Name argid') argtyp
- (wf_rel_fun (mkRel 1) (mkRel (len + 1))))
+ let wfarg len = LocalAssum (Name argid',
+ mkSubset (Name argid') argtyp
+ (wf_rel_fun (mkRel 1) (mkRel (len + 1))))
in
let intern_bl = wfarg 1 :: [arg] in
let _intern_env = push_rel_context intern_bl env in
@@ -967,22 +998,22 @@ let build_wellfounded (recname,n,bl,arityc,body) r measure notation =
(* 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 = (Name (add_suffix recname "'"), None, intern_fun_arity_prod) in
+ let intern_fun_binder = LocalAssum (Name (add_suffix recname "'"), intern_fun_arity_prod) in
let curry_fun =
let wfpred = mkLambda (Name argid', argtyp, wf_rel_fun (mkRel 1) (mkRel (2 * len + 4))) in
let intro = (*FIXME*)Universes.constr_of_global (delayed_force build_sigma).Coqlib.intro in
let arg = mkApp (intro, [| argtyp; wfpred; lift 1 make; mkRel 1 |]) in
let app = mkApp (mkRel (2 * len + 2 (* recproof + orig binders + current binders *)), [| arg |]) in
let rcurry = mkApp (rel, [| measure; lift len measure |]) in
- let lam = (Name (Id.of_string "recproof"), None, rcurry) 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
- (Name recname, Some body, ty)
+ LocalDef (Name recname, body, ty)
in
let fun_bl = intern_fun_binder :: [arg] in
let lift_lets = Termops.lift_rel_context 1 letbinders in
let intern_body =
- let ctx = (Name recname, None, pi3 curry_fun) :: binders_rel in
+ let ctx = LocalAssum (Name recname, get_type curry_fun) :: binders_rel in
let (r, l, impls, scopes) =
Constrintern.compute_internalization_data env
Constrintern.Recursive full_arity impls
@@ -1002,7 +1033,7 @@ let build_wellfounded (recname,n,bl,arityc,body) r measure notation =
~src:(Loc.ghost, Evar_kinds.QuestionMark (Evar_kinds.Define false)) wf_proof;
prop |])
in
- let def = Typing.solve_evars env evdref def 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
@@ -1014,9 +1045,9 @@ let build_wellfounded (recname,n,bl,arityc,body) r measure notation =
let hook l gr _ =
let body = it_mkLambda_or_LetIn (mkApp (Universes.constr_of_global gr, [|make|])) binders_rel in
let ty = it_mkProd_or_LetIn top_arity binders_rel in
- let pl, univs = Evd.universe_context !evdref in
+ let pl, univs = Evd.universe_context ?names:pl !evdref in
(*FIXME poly? *)
- let ce = definition_entry ~types:ty ~univs (Evarutil.nf_evar !evdref body) in
+ let ce = definition_entry ~poly ~types:ty ~univs (Evarutil.nf_evar !evdref body) in
(** FIXME: include locality *)
let c = Declare.declare_constant recname (DefinitionEntry ce, IsDefinition Definition) in
let gr = ConstRef c in
@@ -1040,10 +1071,11 @@ let build_wellfounded (recname,n,bl,arityc,body) r measure notation =
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
+ ignore(Obligations.add_definition recname ~term:evars_def ?pl
evars_typ ctx evars ~hook)
let interp_recursive isfix fixl notations =
+ let open Context.Named.Declaration in
let env = Global.env() in
let fixnames = List.map (fun fix -> fix.fix_name) fixl in
@@ -1056,7 +1088,7 @@ let interp_recursive isfix fixl notations =
| Some ls , Some us ->
if not (CList.for_all2eq (fun x y -> Id.equal (snd x) (snd y)) ls us) then
error "(co)-recursive definitions should all have the same universe binders";
- Some (ls @ us)) fixl None in
+ Some us) fixl None in
let ctx = Evd.make_evar_universe_context env all_universes in
let evdref = ref (Evd.from_ctx ctx) in
let fixctxs, fiximppairs, fixannots =
@@ -1076,11 +1108,11 @@ let interp_recursive isfix fixl notations =
let fixprot =
try
let app = mkApp (delayed_force fix_proto, [|sort; t|]) in
- Typing.solve_evars env evdref app
- with e when Errors.noncritical e -> t
+ Typing.e_solve_evars env evdref app
+ with e when CErrors.noncritical e -> t
in
- (id,None,fixprot) :: env'
- else (id,None,t) :: env')
+ LocalAssum (id,fixprot) :: env'
+ else LocalAssum (id,t) :: env')
[] fixnames fixtypes
in
let env_rec = push_named_context rec_sign env in
@@ -1098,11 +1130,11 @@ let interp_recursive isfix fixl notations =
() in
(* Instantiate evars and check all are resolved *)
- let evd = consider_remaining_unif_problems env_rec !evdref in
+ let evd = solve_unif_constraints_with_heuristics env_rec !evdref in
let evd, nf = nf_evars_and_universes evd in
let fixdefs = List.map (Option.map nf) fixdefs in
let fixtypes = List.map nf fixtypes in
- let fixctxnames = List.map (fun (_,ctx) -> List.map pi1 ctx) fixctxs in
+ let fixctxnames = List.map (fun (_,ctx) -> List.map get_name ctx) fixctxs in
(* Build the fix declaration block *)
(env,rec_sign,all_universes,evd), (fixnames,fixdefs,fixtypes), List.combine3 fixctxnames fiximps fixannots
@@ -1220,6 +1252,11 @@ let out_def = function
| Some def -> def
| None -> error "Program Fixpoint needs defined bodies."
+let collect_evars_of_term evd c ty =
+ let evars = Evar.Set.union (Evd.evars_of_term c) (Evd.evars_of_term ty) in
+ Evar.Set.fold (fun ev acc -> Evd.add acc ev (Evd.find_undefined evd ev))
+ evars (Evd.from_ctx (Evd.evar_universe_context evd))
+
let do_program_recursive local p fixkind fixl ntns =
let isfix = fixkind != Obligations.IsCoFixpoint in
let (env, rec_sign, pl, evd), fix, info =
@@ -1237,8 +1274,9 @@ let do_program_recursive local p fixkind fixl ntns =
and typ =
nf_evar evd (Termops.it_mkNamedProd_or_LetIn typ rec_sign)
in
+ let evm = collect_evars_of_term evd def typ in
let evars, _, def, typ =
- Obligations.eterm_obligations env id evd
+ Obligations.eterm_obligations env id evm
(List.length rec_sign) def typ
in (id, def, typ, imps, evars)
in
@@ -1253,30 +1291,34 @@ let do_program_recursive local p fixkind fixl ntns =
Array.of_list (List.map (subst_vars (List.rev fixnames)) fixdefs)
in
let indexes =
- Pretyping.search_guard Loc.ghost (Global.env ()) possible_indexes fixdecls in
- List.iteri (fun i _ -> Inductive.check_fix env ((indexes,i),fixdecls)) fixl
+ Pretyping.search_guard
+ Loc.ghost (Global.env ()) possible_indexes fixdecls in
+ List.iteri (fun i _ ->
+ Inductive.check_fix env
+ ((indexes,i),fixdecls))
+ fixl
end in
let ctx = Evd.evar_universe_context evd in
let kind = match fixkind with
| Obligations.IsFixpoint _ -> (local, p, Fixpoint)
| Obligations.IsCoFixpoint -> (local, p, CoFixpoint)
in
- Obligations.add_mutual_definitions defs ~kind ctx ntns fixkind
+ Obligations.add_mutual_definitions defs ~kind ?pl ctx ntns fixkind
let do_program_fixpoint local poly l =
let g = List.map (fun ((_,wf,_,_,_),_) -> wf) l in
match g, l with
- | [(n, CWfRec r)], [((((_,id),_),_,bl,typ,def),ntn)] ->
+ | [(n, CWfRec r)], [((((_,id),pl),_,bl,typ,def),ntn)] ->
let recarg =
match n with
| Some n -> mkIdentC (snd n)
| None ->
errorlabstrm "do_program_fixpoint"
(str "Recursive argument required for well-founded fixpoints")
- in build_wellfounded (id, n, bl, typ, out_def def) r recarg ntn
+ in build_wellfounded (id, pl, n, bl, typ, out_def def) poly r recarg ntn
- | [(n, CMeasureRec (m, r))], [((((_,id),_),_,bl,typ,def),ntn)] ->
- build_wellfounded (id, n, bl, typ, out_def def)
+ | [(n, CMeasureRec (m, r))], [((((_,id),pl),_,bl,typ,def),ntn)] ->
+ build_wellfounded (id, pl, n, bl, typ, out_def def) poly
(Option.default (CRef (lt_ref,None)) r) m ntn
| _, _ when List.for_all (fun (n, ro) -> ro == CStructRec) g ->
@@ -1288,6 +1330,11 @@ let do_program_fixpoint local poly l =
errorlabstrm "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
@@ -1295,7 +1342,8 @@ let do_fixpoint local poly l =
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
+ 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
@@ -1303,4 +1351,5 @@ let do_cofixpoint local poly l =
do_program_recursive local poly Obligations.IsCoFixpoint fixl ntns
else
let cofix = interp_cofixpoint fixl ntns in
- declare_cofixpoint local poly cofix ntns
+ declare_cofixpoint local poly cofix ntns;
+ if not (check_safe ()) then Feedback.feedback Feedback.AddedAxiom else ()
diff --git a/toplevel/command.mli b/toplevel/command.mli
index b97cb487..616afb91 100644
--- a/toplevel/command.mli
+++ b/toplevel/command.mli
@@ -22,7 +22,7 @@ open Pfedit
val do_universe : polymorphic -> Id.t Loc.located list -> unit
val do_constraint : polymorphic ->
- (Id.t Loc.located * Univ.constraint_type * Id.t Loc.located) list -> unit
+ (Misctypes.glob_level * Univ.constraint_type * Misctypes.glob_level) list -> unit
(** {6 Hooks for Pcoq} *)
@@ -161,9 +161,11 @@ val declare_cofixpoint : locality -> polymorphic ->
(** 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 *)
diff --git a/toplevel/coqinit.ml b/toplevel/coqinit.ml
index 91cec4bb..acbf909c 100644
--- a/toplevel/coqinit.ml
+++ b/toplevel/coqinit.ml
@@ -36,7 +36,7 @@ let load_rcfile() =
else raise (Sys_error ("Cannot read rcfile: "^ !rcfile))
else
try
- let warn x = msg_warning (str x) in
+ let warn x = Feedback.msg_warning (str x) in
let inferedrc = List.find CUnix.file_readable_p [
Envars.xdg_config_home warn / rcdefaultname^"."^Coq_config.version;
Envars.xdg_config_home warn / rcdefaultname;
@@ -51,21 +51,20 @@ let load_rcfile() =
" found. Skipping rcfile loading."))
*)
with reraise ->
- let reraise = Errors.push reraise in
- let () = msg_info (str"Load of rcfile failed.") in
+ let reraise = CErrors.push reraise in
+ let () = Feedback.msg_info (str"Load of rcfile failed.") in
iraise reraise
else
- Flags.if_verbose msg_info (str"Skipping rcfile loading.")
+ Flags.if_verbose Feedback.msg_info (str"Skipping rcfile loading.")
(* Recursively puts dir in the LoadPath if -nois was not passed *)
let add_stdlib_path ~unix_path ~coq_root ~with_ml =
- Mltop.add_rec_path ~unix_path ~coq_root ~implicit:(!Flags.load_init);
- if with_ml then
- Mltop.add_rec_ml_dir unix_path
+ let add_ml = if with_ml then Mltop.AddRecML else Mltop.AddNoML in
+ Mltop.add_rec_path add_ml ~unix_path ~coq_root ~implicit:(!Flags.load_init)
let add_userlib_path ~unix_path =
- Mltop.add_rec_path ~unix_path ~coq_root:Nameops.default_root_prefix ~implicit:false;
- Mltop.add_rec_ml_dir unix_path
+ Mltop.add_rec_path Mltop.AddRecML ~unix_path
+ ~coq_root:Nameops.default_root_prefix ~implicit:false
(* Options -I, -I-as, and -R of the command line *)
let includes = ref []
@@ -78,7 +77,7 @@ let push_ml_include s = ml_includes := s :: !ml_includes
let init_load_path () =
let coqlib = Envars.coqlib () in
let user_contrib = coqlib/"user-contrib" in
- let xdg_dirs = Envars.xdg_dirs ~warn:(fun x -> msg_warning (str x)) in
+ let xdg_dirs = Envars.xdg_dirs ~warn:(fun x -> Feedback.msg_warning (str x)) in
let coqpath = Envars.coqpath in
let coq_root = Names.DirPath.make [Nameops.coq_root] in
(* NOTE: These directories are searched from last to first *)
@@ -90,7 +89,8 @@ let init_load_path () =
Mltop.add_ml_dir (coqlib/"stm");
Mltop.add_ml_dir (coqlib/"ide")
end;
- Mltop.add_ml_dir (coqlib/"toploop");
+ if System.exists_dir (coqlib/"toploop") then
+ Mltop.add_ml_dir (coqlib/"toploop");
(* then standard library *)
add_stdlib_path ~unix_path:(coqlib/"theories") ~coq_root ~with_ml:false;
(* then plugins *)
@@ -108,7 +108,7 @@ let init_load_path () =
(* additional loadpath, given with options -Q and -R *)
List.iter
(fun (unix_path, coq_root, implicit) ->
- Mltop.add_rec_path ~unix_path ~coq_root ~implicit)
+ Mltop.add_rec_path Mltop.AddNoML ~unix_path ~coq_root ~implicit)
(List.rev !includes);
(* additional ml directories, given with option -I *)
List.iter Mltop.add_ml_dir (List.rev !ml_includes)
@@ -125,16 +125,18 @@ let init_ocaml_path () =
Mltop.add_ml_dir (Envars.coqlib ());
List.iter add_subdir
[ [ "config" ]; [ "dev" ]; [ "lib" ]; [ "kernel" ]; [ "library" ];
- [ "pretyping" ]; [ "interp" ]; [ "parsing" ]; [ "proofs" ];
+ [ "engine" ]; [ "pretyping" ]; [ "interp" ]; [ "parsing" ]; [ "proofs" ];
[ "tactics" ]; [ "toplevel" ]; [ "printing" ]; [ "intf" ];
- [ "grammar" ]; [ "ide" ] ]
+ [ "grammar" ]; [ "ide" ]; [ "ltac" ]; ]
let get_compat_version = function
- | "8.5" -> Flags.Current
+ | "8.6" -> Flags.Current
+ | "8.5" -> Flags.V8_5
| "8.4" -> Flags.V8_4
| "8.3" -> Flags.V8_3
| "8.2" -> Flags.V8_2
| ("8.1" | "8.0") as s ->
- msg_warning (str "Compatibility with version " ++ str s ++ str " not supported.");
- Flags.V8_2
- | s -> Errors.errorlabstrm "get_compat_version" (str "Unknown compatibility version \"" ++ str s ++ str "\".")
+ CErrors.errorlabstrm "get_compat_version"
+ (str "Compatibility with version " ++ str s ++ str " not supported.")
+ | s -> CErrors.errorlabstrm "get_compat_version"
+ (str "Unknown compatibility version \"" ++ str s ++ str "\".")
diff --git a/toplevel/coqloop.ml b/toplevel/coqloop.ml
index 063ed896..e9771cfa 100644
--- a/toplevel/coqloop.ml
+++ b/toplevel/coqloop.ml
@@ -7,12 +7,14 @@
(************************************************************************)
open Pp
-open Errors
+open CErrors
open Util
open Flags
open Vernac
open Pcoq
+let top_stderr x = msg_with ~pp_tag:Ppstyle.pp_tag !Pp_control.err_ft x
+
(* A buffer for the character read from a channel. We store the command
* entered to be able to report errors without pretty-printing. *)
@@ -21,7 +23,7 @@ type input_buffer = {
mutable str : string; (* buffer of already read characters *)
mutable len : int; (* number of chars in the buffer *)
mutable bols : int list; (* offsets in str of beginning of lines *)
- mutable tokens : Gram.parsable; (* stream of tokens *)
+ mutable tokens : Gram.coq_parsable; (* stream of tokens *)
mutable start : int } (* stream count of the first char of the buffer *)
(* Double the size of the buffer. *)
@@ -32,7 +34,7 @@ let resize_buffer ibuf =
ibuf.str <- nstr
(* Delete all irrelevant lines of the input buffer. Keep the last line
- in the buffer (useful when there are several commands on the same line. *)
+ in the buffer (useful when there are several commands on the same line). *)
let resynch_buffer ibuf =
match ibuf.bols with
@@ -59,7 +61,7 @@ let prompt_char ic ibuf count =
| ll::_ -> Int.equal ibuf.len ll
| [] -> Int.equal ibuf.len 0
in
- if bol && not !print_emacs then msgerr (str (ibuf.prompt()));
+ if bol && not !print_emacs then top_stderr (str (ibuf.prompt()));
try
let c = input_char ic in
if c == '\n' then ibuf.bols <- (ibuf.len+1) :: ibuf.bols;
@@ -144,44 +146,23 @@ let print_highlight_location ib loc =
str sn ++ str dn) in
(l1 ++ li ++ ln)
in
- let loc = Loc.make_loc (bp,ep) in
- (str"Toplevel input, characters " ++ Cerrors.print_loc loc ++ str":" ++ fnl () ++
- highlight_lines ++ fnl ())
+ let loc = Loc.make_loc (bp,ep) in
+ (Pp.pr_loc loc ++ highlight_lines ++ fnl ())
(* Functions to report located errors in a file. *)
-let print_location_in_file {outer=s;inner=fname} loc =
- let errstrm = str"Error while reading " ++ str s in
+let print_location_in_file loc =
+ let fname = loc.Loc.fname in
+ let errstrm = str"Error while reading " ++ str fname in
if Loc.is_ghost loc then
hov 1 (errstrm ++ spc() ++ str" (unknown location):") ++ fnl ()
else
- let errstrm =
- if String.equal s fname then mt() else errstrm ++ str":" ++ fnl()
+ let errstrm = mt ()
+ (* if String.equal outer_fname fname then mt() else errstrm ++ str":" ++ fnl() *)
in
- let (bp,ep) = Loc.unloc loc in
- let line_of_pos lin bol cnt =
- try
- let ic = open_in fname in
- let rec line_of_pos lin bol cnt =
- if cnt < bp then
- if input_char ic == '\n'
- then line_of_pos (lin + 1) (cnt +1) (cnt+1)
- else line_of_pos lin bol (cnt+1)
- else (lin, bol)
- in
- let rc = line_of_pos lin bol cnt in
- close_in ic;
- rc
- with Sys_error _ -> 0, 0 in
- try
- let (line, bol) = line_of_pos 1 0 0 in
- hov 0 (* No line break so as to follow emacs error message format *)
- (errstrm ++ str"File " ++ str "\"" ++ str fname ++ str "\"" ++
- str", line " ++ int line ++ str", characters " ++
- Cerrors.print_loc (Loc.make_loc (bp-bol,ep-bol))) ++ str":" ++
- fnl ()
- with e when Errors.noncritical e ->
- hov 1 (errstrm ++ spc() ++ str"(invalid location):") ++ fnl ()
+ let open Loc in
+ hov 0 (* No line break so as to follow emacs error message format *)
+ (errstrm ++ Pp.pr_loc loc)
let valid_buffer_loc ib loc =
not (Loc.is_ghost loc) &&
@@ -262,14 +243,15 @@ let locate_exn = function
let print_toplevel_error (e, info) =
let loc = Option.default Loc.ghost (Loc.get_loc info) in
- let locmsg = match Vernac.get_exn_files info with
- | Some files -> print_location_in_file files loc
- | None ->
+ let fname = loc.Loc.fname in
+ let locmsg =
+ if Loc.is_ghost loc || String.equal fname "" then
if locate_exn e && valid_buffer_loc top_buffer loc then
- print_highlight_location top_buffer loc
+ print_highlight_location top_buffer loc
else mt ()
+ else print_location_in_file loc
in
- locmsg ++ Errors.iprint (e, info)
+ locmsg ++ CErrors.iprint (e, info)
(* Read the input stream until a dot is encountered *)
let parse_to_dot =
@@ -288,14 +270,16 @@ let rec discard_to_dot () =
try
Gram.entry_parse parse_to_dot top_buffer.tokens
with
- | Compat.Token.Error _ | Lexer.Error.E _ -> discard_to_dot ()
+ | Compat.Token.Error _ | CLexer.Error.E _ -> discard_to_dot ()
| End_of_input -> raise End_of_input
- | e when Errors.noncritical e -> ()
+ | e when CErrors.noncritical e -> ()
-let read_sentence () =
- try Vernac.parse_sentence (top_buffer.tokens, None)
+let read_sentence input =
+ try
+ let (loc, _ as r) = Vernac.parse_sentence input in
+ CWarnings.set_current_loc loc; r
with reraise ->
- let reraise = Errors.push reraise in
+ let reraise = CErrors.push reraise in
discard_to_dot ();
iraise reraise
@@ -310,23 +294,23 @@ let read_sentence () =
*)
let do_vernac () =
- msgerrnl (mt ());
- if !print_emacs then msgerr (str (top_buffer.prompt()));
+ top_stderr (fnl());
+ if !print_emacs then top_stderr (str (top_buffer.prompt()));
resynch_buffer top_buffer;
try
- Vernac.eval_expr (read_sentence ())
+ let input = (top_buffer.tokens, None) in
+ Vernac.process_expr top_buffer.tokens (read_sentence input)
with
- | End_of_input | Errors.Quit ->
- msgerrnl (mt ()); pp_flush(); raise Errors.Quit
- | Errors.Drop -> (* Last chance *)
- if Mltop.is_ocaml_top() then raise Errors.Drop
- else ppnl (str"Error: There is no ML toplevel." ++ fnl ())
+ | End_of_input | CErrors.Quit ->
+ top_stderr (fnl ()); raise CErrors.Quit
+ | CErrors.Drop -> (* Last chance *)
+ if Mltop.is_ocaml_top() then raise CErrors.Drop
+ else Feedback.msg_error (str"There is no ML toplevel.")
| any ->
- let any = Errors.push any in
- Format.set_formatter_out_channel stdout;
+ let any = CErrors.push any in
let msg = print_toplevel_error any ++ fnl () in
pp_with ~pp_tag:Ppstyle.pp_tag !Pp_control.std_ft msg;
- pp_flush ()
+ Format.pp_print_flush !Pp_control.std_ft ()
(** Main coq loop : read vernacular expressions until Drop is entered.
Ctrl-C is handled internally as Sys.Break instead of aborting Coq.
@@ -343,18 +327,28 @@ let feed_emacs = function
| _ -> ()
*)
+(* Flush in a compatible order with 8.5 *)
+(* This mimics the semantics of the old Pp.flush_all *)
+let loop_flush_all () =
+ Pervasives.flush stderr;
+ Pervasives.flush stdout;
+ Format.pp_print_flush !Pp_control.std_ft ();
+ Format.pp_print_flush !Pp_control.err_ft ()
+
let rec loop () =
Sys.catch_break true;
if !Flags.print_emacs then Vernacentries.qed_display_script := false;
Flags.coqtop_ui := true;
try
reset_input_buffer stdin top_buffer;
- while true do do_vernac(); flush_all() done
+ while true do do_vernac(); loop_flush_all () done
with
- | Errors.Drop -> ()
- | Errors.Quit -> exit 0
+ | CErrors.Drop -> ()
+ | CErrors.Quit -> exit 0
| any ->
- msgerrnl (str"Anomaly: main loop exited with exception: " ++
+ Feedback.msg_error (str"Anomaly: main loop exited with exception: " ++
str (Printexc.to_string any) ++
- fnl() ++ str"Please report.");
+ fnl() ++
+ str"Please report" ++
+ strbrk" at " ++ str Coq_config.wwwbugtracker ++ str ".");
loop ()
diff --git a/toplevel/coqloop.mli b/toplevel/coqloop.mli
index 00554da3..e40353e0 100644
--- a/toplevel/coqloop.mli
+++ b/toplevel/coqloop.mli
@@ -18,7 +18,7 @@ type input_buffer = {
mutable str : string; (** buffer of already read characters *)
mutable len : int; (** number of chars in the buffer *)
mutable bols : int list; (** offsets in str of begining of lines *)
- mutable tokens : Pcoq.Gram.parsable; (** stream of tokens *)
+ mutable tokens : Pcoq.Gram.coq_parsable; (** stream of tokens *)
mutable start : int } (** stream count of the first char of the buffer *)
(** The input buffer of stdin. *)
diff --git a/toplevel/coqtop.ml b/toplevel/coqtop.ml
index afd4ef40..d9f8ed88 100644
--- a/toplevel/coqtop.ml
+++ b/toplevel/coqtop.ml
@@ -7,9 +7,8 @@
(************************************************************************)
open Pp
-open Errors
+open CErrors
open Util
-open System
open Flags
open Names
open Libnames
@@ -27,15 +26,15 @@ let get_version_date () =
let rev = input_line ch in
let () = close_in ch in
(ver,rev)
- with e when Errors.noncritical e ->
+ with e when CErrors.noncritical e ->
(Coq_config.version,Coq_config.date)
let print_header () =
let (ver,rev) = get_version_date () in
- ppnl (str "Welcome to Coq " ++ str ver ++ str " (" ++ str rev ++ str ")");
- pp_flush ()
+ Feedback.msg_notice (str "Welcome to Coq " ++ str ver ++ str " (" ++ str rev ++ str ")");
+ flush_all ()
-let warning s = msg_warning (strbrk s)
+let warning s = with_option Flags.warn Feedback.msg_warning (strbrk s)
let toploop = ref None
@@ -55,14 +54,14 @@ let init_color () =
Terminal.has_style Unix.stderr &&
(* emacs compilation buffer does not support colors by default,
its TERM variable is set to "dumb". *)
- Unix.getenv "TERM" <> "dumb"
+ try Sys.getenv "TERM" <> "dumb" with Not_found -> false
in
if has_color then begin
let colors = try Some (Sys.getenv "COQ_COLORS") with Not_found -> None in
match colors with
| None ->
(** Default colors *)
- Ppstyle.init_color_output ()
+ Feedback.init_color_output ()
| Some "" ->
(** No color output *)
()
@@ -70,7 +69,7 @@ let init_color () =
(** Overwrite all colors *)
Ppstyle.clear_styles ();
Ppstyle.parse_config s;
- Ppstyle.init_color_output ()
+ Feedback.init_color_output ()
end
let toploop_init = ref begin fun x ->
@@ -96,8 +95,8 @@ let memory_stat = ref false
let print_memory_stat () =
begin (* -m|--memory from the command-line *)
if !memory_stat then
- ppnl
- (str "total heap size = " ++ int (CObj.heap_size_kb ()) ++ str " kbytes");
+ Feedback.msg_notice
+ (str "total heap size = " ++ int (CObj.heap_size_kb ()) ++ str " kbytes" ++ fnl ());
end;
begin
(* operf-macro interface:
@@ -114,23 +113,14 @@ let _ = at_exit print_memory_stat
let impredicative_set = ref Declarations.PredicativeSet
let set_impredicative_set c = impredicative_set := Declarations.ImpredicativeSet
-let type_in_type = ref Declarations.StratifiedType
-let set_type_in_type () = type_in_type := Declarations.TypeInType
+let set_type_in_type () =
+ let typing_flags = Environ.typing_flags (Global.env ()) in
+ Global.set_typing_flags { typing_flags with Declarations.check_universes = false }
let engage () =
- Global.set_engagement (!impredicative_set,!type_in_type)
+ Global.set_engagement !impredicative_set
let set_batch_mode () = batch_mode := true
-let user_warning = ref false
-(** User explicitly set warning *)
-
-let set_warning p =
- let () = user_warning := true in
- match p with
- | "all" -> make_warn true
- | "none" -> make_warn false
- | _ -> prerr_endline ("Error: all/none expected after option w"); exit 1
-
let toplevel_default_name = DirPath.make [Id.of_string "Top"]
let toplevel_name = ref (Some toplevel_default_name)
let set_toplevel_name dir =
@@ -140,18 +130,27 @@ let unset_toplevel_name () = toplevel_name := None
let remove_top_ml () = Mltop.remove ()
+let warn_deprecated_inputstate =
+ CWarnings.create ~name:"deprecated-inputstate" ~category:"deprecated"
+ (fun () -> strbrk "The inputstate option is deprecated and discouraged.")
+
let inputstate = ref ""
let set_inputstate s =
- let () = msg_warning (str "The inputstate option is deprecated and discouraged.") in
+ warn_deprecated_inputstate ();
inputstate:=s
let inputstate () =
if not (String.is_empty !inputstate) then
let fname = Loadpath.locate_file (CUnix.make_suffix !inputstate ".coq") in
intern_state fname
+let warn_deprecated_outputstate =
+ CWarnings.create ~name:"deprecated-outputstate" ~category:"deprecated"
+ (fun () ->
+ strbrk "The outputstate option is deprecated and discouraged.")
+
let outputstate = ref ""
let set_outputstate s =
- let () = msg_warning (str "The outputstate option is deprecated and discouraged.") in
+ warn_deprecated_outputstate ();
outputstate:=s
let outputstate () =
if not (String.is_empty !outputstate) then
@@ -169,7 +168,7 @@ let load_vernacular () =
List.iter
(fun (s,b) ->
let s = Loadpath.locate_file s in
- if Flags.do_beautify () then
+ if !Flags.beautify then
with_option beautify_file (Vernac.load_vernac b) s
else
Vernac.load_vernac b s)
@@ -199,6 +198,7 @@ let require () =
let add_compat_require v =
match v with
| Flags.V8_4 -> add_require "Coq.Compat.Coq84"
+ | Flags.V8_5 -> add_require "Coq.Compat.Coq85"
| _ -> ()
let compile_list = ref ([] : (bool * string) list)
@@ -219,33 +219,44 @@ let add_compile verbose s =
compile_list := (verbose,s) :: !compile_list
let compile_file (v,f) =
- if Flags.do_beautify () then
+ if !Flags.beautify then
with_option beautify_file (Vernac.compile v) f
else
Vernac.compile v f
let compile_files () =
- match !compile_list with
- | [] -> ()
- | [vf] -> compile_file vf (* One compilation : no need to save init state *)
- | l ->
- let init_state = States.freeze ~marshallable:`No in
- let coqdoc_init_state = Lexer.location_table () in
- List.iter
- (fun vf ->
- States.unfreeze init_state;
- Lexer.restore_location_table coqdoc_init_state;
- compile_file vf)
- (List.rev l)
+ if !compile_list == [] then ()
+ else
+ let init_state = States.freeze ~marshallable:`No in
+ Feedback.(add_feeder debug_feeder);
+ List.iter (fun vf ->
+ States.unfreeze init_state;
+ compile_file vf)
+ (List.rev !compile_list)
(** Options for proof general *)
let set_emacs () =
+ if not (Option.is_empty !toploop) then
+ error "Flag -emacs is incompatible with a custom toplevel loop";
Flags.print_emacs := true;
- Pp.make_pp_emacs ();
+ Feedback.(set_logger emacs_logger);
Vernacentries.qed_display_script := false;
color := `OFF
+(** Options for CoqIDE *)
+
+let set_ideslave () =
+ if !Flags.print_emacs then error "Flags -ideslave and -emacs are incompatible";
+ toploop := Some "coqidetop";
+ Flags.ide_slave := true
+
+(** Options for slaves *)
+
+let set_toploop name =
+ if !Flags.print_emacs then error "Flags -toploop and -emacs are incompatible";
+ toploop := Some name
+
(** GC tweaking *)
(** Coq is a heavy user of persistent data structures and symbolic ASTs, so the
@@ -257,25 +268,26 @@ let set_emacs () =
*)
let init_gc () =
- let param =
- try ignore (Sys.getenv "OCAMLRUNPARAM"); true
- with Not_found -> false
- in
- let control = Gc.get () in
- let tweaked_control = { control with
- Gc.minor_heap_size = 33554432; (** 4M *)
-(* Gc.major_heap_increment = 268435456; (** 32M *) *)
- Gc.space_overhead = 120;
- } in
- if param then ()
- else Gc.set tweaked_control
+ try
+ (* OCAMLRUNPARAM environment variable is set.
+ * In that case, we let ocamlrun to use the values provided by the user.
+ *)
+ ignore (Sys.getenv "OCAMLRUNPARAM")
+
+ with Not_found ->
+ (* OCAMLRUNPARAM environment variable is not set.
+ * In this case, we put in place our preferred configuration.
+ *)
+ Gc.set { (Gc.get ()) with
+ Gc.minor_heap_size = 33554432; (** 4M *)
+ Gc.space_overhead = 120}
(*s Parsing of the command line.
We no longer use [Arg.parse], in order to use share [Usage.print_usage]
between coqtop and coqc. *)
let usage () =
- Envars.set_coqlib Errors.error;
+ Envars.set_coqlib CErrors.error;
init_load_path ();
if !batch_mode then Usage.print_usage_coqc ()
else begin
@@ -316,9 +328,6 @@ let error_missing_arg s =
let filter_opts = ref false
let exitcode () = if !filter_opts then 2 else 0
-let verb_compat_ntn = ref false
-let no_compat_ntn = ref false
-
let print_where = ref false
let print_config = ref false
let print_tags = ref false
@@ -353,15 +362,25 @@ let get_int opt n =
with Failure _ ->
prerr_endline ("Error: integer expected after option "^opt); exit 1
+let get_float opt n =
+ try float_of_string n
+ with Failure _ ->
+ prerr_endline ("Error: float expected after option "^opt); exit 1
+
let get_host_port opt s =
match CString.split ':' s with
| [host; portr; portw] ->
Some (Spawned.Socket(host, int_of_string portr, int_of_string portw))
| ["stdfds"] -> Some Spawned.AnonPipe
| _ ->
- prerr_endline ("Error: host:port or stdfds expected after option "^opt);
+ prerr_endline ("Error: host:portr:portw or stdfds expected after option "^opt);
exit 1
+let get_error_resilience opt = function
+ | "on" | "all" | "yes" -> `All
+ | "off" | "no" -> `None
+ | s -> `Only (String.split ',' s)
+
let get_task_list s = List.map int_of_string (Str.split (Str.regexp ",") s)
let vio_tasks = ref []
@@ -442,10 +461,6 @@ let parse_args arglist =
end
|"-R" ->
begin match rem with
- | d :: "-as" :: [] -> error_missing_arg opt
- | d :: "-as" :: p :: rem ->
- warning "option -R * -as * deprecated, remove the -as";
- set_include d p true; args := rem
| d :: p :: rem -> set_include d p true; args := rem
| _ -> error_missing_arg opt
end
@@ -479,13 +494,19 @@ let parse_args arglist =
Flags.async_proofs_worker_priority := get_priority opt (next ())
|"-async-proofs-private-flags" ->
Flags.async_proofs_private_flags := Some (next ());
+ |"-async-proofs-tactic-error-resilience" ->
+ Flags.async_proofs_tac_error_resilience := get_error_resilience opt (next ())
+ |"-async-proofs-command-error-resilience" ->
+ Flags.async_proofs_cmd_error_resilience := get_bool opt (next ())
+ |"-async-proofs-delegation-threshold" ->
+ Flags.async_proofs_delegation_threshold:= get_float opt (next ())
|"-worker-id" -> set_worker_id opt (next ())
|"-compat" -> let v = get_compat_version (next ()) in Flags.compat_version := v; add_compat_require v
|"-compile" -> add_compile false (next ())
|"-compile-verbose" -> add_compile true (next ())
|"-dump-glob" -> Dumpglob.dump_into_file (next ()); glob_opt := true
|"-feedback-glob" -> Dumpglob.feedback_glob ()
- |"-exclude-dir" -> exclude_search_in_dirname (next ())
+ |"-exclude-dir" -> System.exclude_directory (next ())
|"-init-file" -> set_rcfile (next ())
|"-inputstate"|"-is" -> set_inputstate (next ())
|"-load-ml-object" -> Mltop.dir_ml_load (next ())
@@ -495,14 +516,16 @@ let parse_args arglist =
|"-load-vernac-source-verbose"|"-lv" -> add_load_vernacular true (next ())
|"-outputstate" -> set_outputstate (next ())
|"-print-mod-uid" -> let s = String.concat " " (List.map get_native_name rem) in print_endline s; exit 0
+ |"-profile-ltac-cutoff" -> Flags.profile_ltac := true; Flags.profile_ltac_cutoff := get_float opt (next ())
|"-require" -> add_require (next ())
|"-top" -> set_toplevel_name (dirpath_of_string (next ()))
|"-with-geoproof" -> Coq_config.with_geoproof := get_bool opt (next ())
|"-main-channel" -> Spawned.main_channel := get_host_port opt (next())
|"-control-channel" -> Spawned.control_channel := get_host_port opt (next())
|"-vio2vo" -> add_compile false (next ()); Flags.compilation_mode := Vio2Vo
- |"-toploop" -> toploop := Some (next ())
- |"-w" -> set_warning (next ())
+ |"-toploop" -> set_toploop (next ())
+ |"-w" | "-W" -> CWarnings.set_flags (CWarnings.normalize_flags_string (next ()))
+ |"-o" -> Flags.compilation_output_name := Some (next())
(* Options with zero arg *)
|"-async-queries-always-delegate"
@@ -513,7 +536,7 @@ let parse_args arglist =
Flags.async_proofs_never_reopen_branch := true;
|"-batch" -> set_batch_mode ()
|"-test-mode" -> test_mode := true
- |"-beautify" -> make_beautify true
+ |"-beautify" -> beautify := true
|"-boot" -> boot := true; no_load_rc ()
|"-bt" -> Backtrace.record_backtrace true
|"-color" -> set_color (next ())
@@ -522,13 +545,12 @@ let parse_args arglist =
|"-emacs" -> set_emacs ()
|"-filteropts" -> filter_opts := true
|"-h"|"-H"|"-?"|"-help"|"--help" -> usage ()
- |"-ideslave" -> toploop := Some "coqidetop"; Flags.ide_slave := true
+ |"-ideslave" -> set_ideslave ()
|"-impredicative-set" -> set_impredicative_set ()
|"-indices-matter" -> Indtypes.enforce_indices_matter ()
- |"-just-parsing" -> Vernac.just_parsing := true
+ |"-just-parsing" -> warning "-just-parsing option has been removed in 8.6"
|"-m"|"--memory" -> memory_stat := true
|"-noinit"|"-nois" -> load_init := false
- |"-no-compat-notations" -> no_compat_ntn := true
|"-no-glob"|"-noglob" -> Dumpglob.noglob (); glob_opt := true
|"-native-compiler" ->
if Coq_config.no_native_compiler then
@@ -536,6 +558,7 @@ let parse_args arglist =
else native_compiler := true
|"-notop" -> unset_toplevel_name ()
|"-output-context" -> output_context := true
+ |"-profile-ltac" -> Flags.profile_ltac := true
|"-q" -> no_load_rc ()
|"-quiet"|"-silent" -> Flags.make_silent true; Flags.make_warn false
|"-quick" -> Flags.compilation_mode := BuildVio
@@ -544,7 +567,7 @@ let parse_args arglist =
|"-type-in-type" -> set_type_in_type ()
|"-unicode" -> add_require "Utf8_core"
|"-v"|"--version" -> Usage.version (exitcode ())
- |"-verbose-compat-notations" -> verb_compat_ntn := true
+ |"--print-version" -> Usage.machine_readable_version (exitcode ())
|"-where" -> print_where := true
|"-xml" -> Flags.xml_export := true
@@ -573,10 +596,10 @@ let parse_args arglist =
with
| UserError(_, s) as e ->
if is_empty s then exit 1
- else fatal_error (Errors.print e) false
- | any -> fatal_error (Errors.print any) (Errors.is_anomaly any)
+ else fatal_error (CErrors.print e) false
+ | any -> fatal_error (CErrors.print any) (CErrors.is_anomaly any)
-let init arglist =
+let init_toplevel arglist =
init_gc ();
Sys.catch_break false; (* Ctrl-C is fatal during the initialisation *)
Lib.init();
@@ -588,7 +611,7 @@ let init arglist =
(* If we have been spawned by the Spawn module, this has to be done
* early since the master waits us to connect back *)
Spawned.init_channels ();
- Envars.set_coqlib Errors.error;
+ Envars.set_coqlib CErrors.error;
if !print_where then (print_endline(Envars.coqlib ()); exit(exitcode ()));
if !print_config then (Usage.print_config (); exit (exitcode ()));
if !print_tags then (print_style_tags (); exit (exitcode ()));
@@ -605,9 +628,6 @@ let init arglist =
inputstate ();
Mltop.init_known_plugins ();
engage ();
- (* Be careful to set these variables after the inputstate *)
- Syntax_def.set_verbose_compat_notations !verb_compat_ntn;
- Syntax_def.set_compat_notations (not !no_compat_ntn);
if (not !batch_mode || List.is_empty !compile_list)
&& Global.env_is_initial ()
then Option.iter Declaremods.start_library !toplevel_name;
@@ -623,30 +643,27 @@ let init arglist =
check_vio_tasks ();
outputstate ()
with any ->
- let any = Errors.push any in
+ let any = CErrors.push any in
flush_all();
let msg =
if !batch_mode then mt ()
else str "Error during initialization:" ++ fnl ()
in
- let is_anomaly e = Errors.is_anomaly e || not (Errors.handled e) in
+ let is_anomaly e = CErrors.is_anomaly e || not (CErrors.handled e) in
fatal_error (msg ++ Coqloop.print_toplevel_error any) (is_anomaly (fst any))
end;
if !batch_mode then begin
flush_all();
if !output_context then
- Pp.ppnl (with_option raw_print Prettyp.print_full_pure_context ());
+ Feedback.msg_notice (with_option raw_print Prettyp.print_full_pure_context () ++ fnl ());
Profile.print_profile ();
exit 0
end
-let init_toplevel = init
-
let start () =
let () = init_toplevel (List.tl (Array.to_list Sys.argv)) in
(* In batch mode, Coqtop has already exited at this point. In interactive one,
dump glob is nothing but garbage ... *)
- if not !user_warning then make_warn true;
!toploop_run ();
exit 1
diff --git a/toplevel/discharge.ml b/toplevel/discharge.ml
index 61573091..e24d5e74 100644
--- a/toplevel/discharge.ml
+++ b/toplevel/discharge.ml
@@ -7,22 +7,23 @@
(************************************************************************)
open Names
-open Errors
+open CErrors
open Util
-open Context
open Term
open Vars
-open Entries
open Declarations
open Cooking
+open Entries
+open Context.Rel.Declaration
(********************************)
(* Discharging mutual inductive *)
-let detype_param = function
- | (Name id,None,p) -> id, LocalAssum p
- | (Name id,Some p,_) -> id, LocalDef p
- | (Anonymous,_,_) -> anomaly (Pp.str "Unnamed inductive local variable")
+let detype_param =
+ function
+ | LocalAssum (Name id, p) -> id, LocalAssumEntry p
+ | LocalDef (Name id, p,_) -> id, LocalDefEntry p
+ | _ -> anomaly (Pp.str "Unnamed inductive local variable")
(* Replace
@@ -37,8 +38,8 @@ let detype_param = function
let abstract_inductive hyps nparams inds =
let ntyp = List.length inds in
- let nhyp = named_context_length hyps in
- let args = instance_from_named_context (List.rev hyps) in
+ let nhyp = Context.Named.length hyps in
+ let args = Context.Named.to_instance (List.rev hyps) in
let args = Array.of_list args in
let subs = List.init ntyp (fun k -> lift nhyp (mkApp(mkRel (k+1),args))) in
let inds' =
@@ -53,7 +54,7 @@ let abstract_inductive hyps nparams inds =
(* To be sure to be the same as before, should probably be moved to process_inductive *)
let params' = let (_,arity,_,_,_) = List.hd inds' in
let (params,_) = decompose_prod_n_assum nparams' arity in
- List.map detype_param params
+ List.map detype_param params
in
let ind'' =
List.map
@@ -100,7 +101,7 @@ let process_inductive (sechyps,abs_ctx) modlist mib =
Array.to_list mip.mind_consnames,
Array.to_list lc))
mib.mind_packets in
- let sechyps' = map_named_context (expmod_constr modlist) sechyps in
+ let sechyps' = Context.Named.map (expmod_constr modlist) sechyps in
let (params',inds') = abstract_inductive sechyps' nparams inds in
let abs_ctx = Univ.instantiate_univ_context abs_ctx in
let univs = Univ.UContext.union abs_ctx univs in
@@ -115,5 +116,5 @@ let process_inductive (sechyps,abs_ctx) modlist mib =
mind_entry_inds = inds';
mind_entry_polymorphic = mib.mind_polymorphic;
mind_entry_private = mib.mind_private;
- mind_entry_universes = univs
+ mind_entry_universes = univs;
}
diff --git a/toplevel/discharge.mli b/toplevel/discharge.mli
index 59140157..18d1b677 100644
--- a/toplevel/discharge.mli
+++ b/toplevel/discharge.mli
@@ -6,10 +6,9 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Context
open Declarations
open Entries
open Opaqueproof
val process_inductive :
- named_context Univ.in_universe_context -> work_list -> mutual_inductive_body -> mutual_inductive_entry
+ Context.Named.t Univ.in_universe_context -> work_list -> mutual_inductive_body -> mutual_inductive_entry
diff --git a/toplevel/cerrors.ml b/toplevel/explainErr.ml
index 600683d3..17897460 100644
--- a/toplevel/cerrors.ml
+++ b/toplevel/explainErr.ml
@@ -7,19 +7,12 @@
(************************************************************************)
open Pp
-open Errors
+open CErrors
open Indtypes
open Type_errors
open Pretype_errors
open Indrec
-let print_loc loc =
- if Loc.is_ghost loc then
- (str"<unknown>")
- else
- let loc = Loc.unloc loc in
- (int (fst loc) ++ str"-" ++ int (snd loc))
-
let guill s = str "\"" ++ str s ++ str "\""
(** Invariant : exceptions embedded in EvaluatedError satisfy
@@ -35,7 +28,7 @@ let explain_exn_default = function
(* Basic interaction exceptions *)
| Stream.Error txt -> hov 0 (str "Syntax error: " ++ str txt ++ str ".")
| Compat.Token.Error txt -> hov 0 (str "Syntax error: " ++ str txt ++ str ".")
- | Lexer.Error.E err -> hov 0 (str (Lexer.Error.to_string err))
+ | CLexer.Error.E err -> hov 0 (str (CLexer.Error.to_string err))
| Sys_error msg -> hov 0 (str "System error: " ++ guill msg)
| Out_of_memory -> hov 0 (str "Out of memory.")
| Stack_overflow -> hov 0 (str "Stack overflow.")
@@ -43,11 +36,11 @@ let explain_exn_default = function
| Sys.Break -> hov 0 (fnl () ++ str "User interrupt.")
(* Exceptions with pre-evaluated error messages *)
| EvaluatedError (msg,None) -> msg
- | EvaluatedError (msg,Some reraise) -> msg ++ Errors.print reraise
+ | EvaluatedError (msg,Some reraise) -> msg ++ CErrors.print reraise
(* Otherwise, not handled here *)
- | _ -> raise Errors.Unhandled
+ | _ -> raise CErrors.Unhandled
-let _ = Errors.register_handler explain_exn_default
+let _ = CErrors.register_handler explain_exn_default
(** Pre-explain a vernac interpretation error *)
@@ -110,29 +103,27 @@ let rec strip_wrapping_exceptions = function
strip_wrapping_exceptions e
| exc -> exc
+let additional_error_info = ref []
+
+let register_additional_error_info f =
+ additional_error_info := f :: !additional_error_info
+
let process_vernac_interp_error ?(allow_uncaught=true) ?(with_header=true) (exc, info) =
let exc = strip_wrapping_exceptions exc in
let e = process_vernac_interp_error with_header (exc, info) in
let () =
- if not allow_uncaught && not (Errors.handled (fst e)) then
+ if not allow_uncaught && not (CErrors.handled (fst e)) then
let (e, info) = e in
let msg = str "Uncaught exception " ++ str (Printexc.to_string e) in
- let err = Errors.make_anomaly msg in
+ let err = CErrors.make_anomaly msg in
Util.iraise (err, info)
in
- let ltac_trace = Exninfo.get info Proof_type.ltac_trace_info in
- let loc = Option.default Loc.ghost (Loc.get_loc info) in
- match ltac_trace with
+ let e' =
+ try Some (CList.find_map (fun f -> f e) !additional_error_info)
+ with _ -> None
+ in
+ match e' with
| None -> e
- | Some trace ->
- let (e, info) = e in
- match Himsg.extract_ltac_trace trace loc with
- | None, loc -> (e, Loc.add_loc info loc)
- | Some msg, loc ->
- (EvaluatedError (msg, Some e), Loc.add_loc info loc)
-
-let _ = Tactic_debug.explain_logic_error :=
- (fun e -> Errors.print (fst (process_vernac_interp_error (e, Exninfo.null))))
-
-let _ = Tactic_debug.explain_logic_error_no_anomaly :=
- (fun e -> Errors.print_no_report (fst (process_vernac_interp_error (e, Exninfo.null))))
+ | Some (None, loc) -> (fst e, Loc.add_loc (snd e) loc)
+ | Some (Some msg, loc) ->
+ (EvaluatedError (msg, Some (fst e)), Loc.add_loc (snd e) loc)
diff --git a/toplevel/cerrors.mli b/toplevel/explainErr.mli
index 68c46010..a67c887a 100644
--- a/toplevel/cerrors.mli
+++ b/toplevel/explainErr.mli
@@ -6,9 +6,8 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(** Error report. *)
-
-val print_loc : Loc.t -> Pp.std_ppcmds
+(** Toplevel Exception *)
+exception EvaluatedError of Pp.std_ppcmds * exn option
(** Pre-explain a vernac interpretation error *)
@@ -19,3 +18,4 @@ val process_vernac_interp_error : ?allow_uncaught:bool -> ?with_header:bool -> U
val explain_exn_default : exn -> Pp.std_ppcmds
+val register_additional_error_info : (Util.iexn -> (Pp.std_ppcmds option * Loc.t) option) -> unit
diff --git a/toplevel/g_obligations.ml4 b/toplevel/g_obligations.ml4
deleted file mode 100644
index d620febb..00000000
--- a/toplevel/g_obligations.ml4
+++ /dev/null
@@ -1,135 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(*i camlp4deps: "grammar/grammar.cma" i*)
-
-(*
- Syntax for the subtac terms and types.
- Elaborated from correctness/psyntax.ml4 by Jean-Christophe Filliâtre *)
-
-
-open Libnames
-open Constrexpr
-open Constrexpr_ops
-
-(* We define new entries for programs, with the use of this module
- * Subtac. These entries are named Subtac.<foo>
- *)
-
-module Gram = Pcoq.Gram
-module Vernac = Pcoq.Vernac_
-module Tactic = Pcoq.Tactic
-
-open Pcoq
-
-let sigref = mkRefC (Qualid (Loc.ghost, Libnames.qualid_of_string "Coq.Init.Specif.sig"))
-
-type 'a withtac_argtype = (Tacexpr.raw_tactic_expr option, 'a) Genarg.abstract_argument_type
-
-let wit_withtac : Tacexpr.raw_tactic_expr option Genarg.uniform_genarg_type =
- Genarg.create_arg None "withtac"
-
-let withtac = Pcoq.create_generic_entry "withtac" (Genarg.rawwit wit_withtac)
-
-GEXTEND Gram
- GLOBAL: withtac;
-
- withtac:
- [ [ "with"; t = Tactic.tactic -> Some t
- | -> None ] ]
- ;
-
- Constr.closed_binder:
- [[ "("; id=Prim.name; ":"; t=Constr.lconstr; "|"; c=Constr.lconstr; ")" ->
- let typ = mkAppC (sigref, [mkLambdaC ([id], default_binder_kind, t, c)]) in
- [LocalRawAssum ([id], default_binder_kind, typ)]
- ] ];
-
- END
-
-open Obligations
-
-let classify_obbl _ = Vernacexpr.(VtStartProof ("Classic",Doesn'tGuaranteeOpacity,[]), VtLater)
-
-VERNAC COMMAND EXTEND Obligations CLASSIFIED BY classify_obbl
-| [ "Obligation" integer(num) "of" ident(name) ":" lconstr(t) withtac(tac) ] ->
- [ obligation (num, Some name, Some t) tac ]
-| [ "Obligation" integer(num) "of" ident(name) withtac(tac) ] ->
- [ obligation (num, Some name, None) tac ]
-| [ "Obligation" integer(num) ":" lconstr(t) withtac(tac) ] ->
- [ obligation (num, None, Some t) tac ]
-| [ "Obligation" integer(num) withtac(tac) ] ->
- [ obligation (num, None, None) tac ]
-| [ "Next" "Obligation" "of" ident(name) withtac(tac) ] ->
- [ next_obligation (Some name) tac ]
-| [ "Next" "Obligation" withtac(tac) ] -> [ next_obligation None tac ]
-END
-
-VERNAC COMMAND EXTEND Solve_Obligation CLASSIFIED AS SIDEFF
-| [ "Solve" "Obligation" integer(num) "of" ident(name) "with" tactic(t) ] ->
- [ try_solve_obligation num (Some name) (Some (Tacinterp.interp t)) ]
-| [ "Solve" "Obligation" integer(num) "with" tactic(t) ] ->
- [ try_solve_obligation num None (Some (Tacinterp.interp t)) ]
-END
-
-VERNAC COMMAND EXTEND Solve_Obligations CLASSIFIED AS SIDEFF
-| [ "Solve" "Obligations" "of" ident(name) "with" tactic(t) ] ->
- [ try_solve_obligations (Some name) (Some (Tacinterp.interp t)) ]
-| [ "Solve" "Obligations" "with" tactic(t) ] ->
- [ try_solve_obligations None (Some (Tacinterp.interp t)) ]
-| [ "Solve" "Obligations" ] ->
- [ try_solve_obligations None None ]
-END
-
-VERNAC COMMAND EXTEND Solve_All_Obligations CLASSIFIED AS SIDEFF
-| [ "Solve" "All" "Obligations" "with" tactic(t) ] ->
- [ solve_all_obligations (Some (Tacinterp.interp t)) ]
-| [ "Solve" "All" "Obligations" ] ->
- [ solve_all_obligations None ]
-END
-
-VERNAC COMMAND EXTEND Admit_Obligations CLASSIFIED AS SIDEFF
-| [ "Admit" "Obligations" "of" ident(name) ] -> [ admit_obligations (Some name) ]
-| [ "Admit" "Obligations" ] -> [ admit_obligations None ]
-END
-
-VERNAC COMMAND EXTEND Set_Solver CLASSIFIED AS SIDEFF
-| [ "Obligation" "Tactic" ":=" tactic(t) ] -> [
- set_default_tactic
- (Locality.make_section_locality (Locality.LocalityFixme.consume ()))
- (Tacintern.glob_tactic t) ]
-END
-
-open Pp
-
-VERNAC COMMAND EXTEND Show_Solver CLASSIFIED AS QUERY
-| [ "Show" "Obligation" "Tactic" ] -> [
- msg_info (str"Program obligation tactic is " ++ print_default_tactic ()) ]
-END
-
-VERNAC COMMAND EXTEND Show_Obligations CLASSIFIED AS QUERY
-| [ "Obligations" "of" ident(name) ] -> [ show_obligations (Some name) ]
-| [ "Obligations" ] -> [ show_obligations None ]
-END
-
-VERNAC COMMAND EXTEND Show_Preterm CLASSIFIED AS QUERY
-| [ "Preterm" "of" ident(name) ] -> [ msg_info (show_term (Some name)) ]
-| [ "Preterm" ] -> [ msg_info (show_term None) ]
-END
-
-open Pp
-
-(* Declare a printer for the content of Program tactics *)
-let () =
- let printer _ _ _ = function
- | None -> mt ()
- | Some tac -> str "with" ++ spc () ++ Pptactic.pr_raw_tactic tac
- in
- (* should not happen *)
- let dummy _ _ _ expr = assert false in
- Pptactic.declare_extra_genarg_pprule wit_withtac printer dummy dummy
diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml
index 7ddfd46c..f98505c3 100644
--- a/toplevel/himsg.ml
+++ b/toplevel/himsg.ml
@@ -23,22 +23,26 @@ open Cases
open Logic
open Printer
open Evd
+open Context.Rel.Declaration
(* This simplifies the typing context of Cases clauses *)
(* hope it does not disturb other typing contexts *)
let contract env lc =
let l = ref [] in
- let contract_context (na,c,t) env =
- match c with
- | Some c' when isRel c' ->
+ let contract_context decl env =
+ match decl with
+ | LocalDef (_,c',_) when isRel c' ->
l := (Vars.substl !l c') :: !l;
env
| _ ->
- let t' = Vars.substl !l t in
- let c' = Option.map (Vars.substl !l) c in
- let na' = named_hd env t' na in
+ let t' = Vars.substl !l (get_type decl) in
+ let c' = Option.map (Vars.substl !l) (get_value decl) in
+ let na' = named_hd env t' (get_name decl) in
l := (mkRel 1) :: List.map (Vars.lift 1) !l;
- push_rel (na',c',t') env in
+ match c' with
+ | None -> push_rel (LocalAssum (na',t')) env
+ | Some c' -> push_rel (LocalDef (na',c',t')) env
+ in
let env = process_rel_context contract_context env in
(env, List.map (Vars.substl !l) lc)
@@ -65,12 +69,21 @@ let rec contract3' env a b c = function
let (env',t1,t2) = contract2 env' t1 t2 in
contract3 env a b c, ConversionFailed (env',t1,t2)
| NotSameArgSize | NotSameHead | NoCanonicalStructure
- | MetaOccurInBody _ | InstanceNotSameType _
+ | MetaOccurInBody _ | InstanceNotSameType _ | ProblemBeyondCapabilities
| UnifUnivInconsistency _ as x -> contract3 env a b c, x
- | CannotSolveConstraint ((pb,env,t,u),x) ->
- let env,t,u = contract2 env t u in
+ | CannotSolveConstraint ((pb,env',t,u),x) ->
+ let env',t,u = contract2 env' t u in
let y,x = contract3' env a b c x in
- y,CannotSolveConstraint ((pb,env,t,u),x)
+ y,CannotSolveConstraint ((pb,env',t,u),x)
+
+(** Ad-hoc reductions *)
+
+let j_nf_betaiotaevar sigma j =
+ { uj_val = Evarutil.nf_evar sigma j.uj_val;
+ uj_type = Reductionops.nf_betaiota sigma j.uj_type }
+
+let jv_nf_betaiotaevar sigma jl =
+ Array.map (j_nf_betaiotaevar sigma) jl
(** Printers *)
@@ -136,9 +149,9 @@ let pr_explicit env sigma t1 t2 = pr_explicit_aux env sigma t1 t2 explicit_flags
let pr_db env i =
try
- match lookup_rel i env with
- Name id, _, _ -> pr_id id
- | Anonymous, _, _ -> str "<>"
+ match lookup_rel i env |> get_name with
+ | Name id -> pr_id id
+ | Anonymous -> str "<>"
with Not_found -> str "UNBOUND_REL_" ++ int i
let explain_unbound_rel env sigma n =
@@ -260,7 +273,7 @@ let explain_generalization env sigma (name,var) j =
str "it has type" ++ spc () ++ pt ++
spc () ++ str "which should be Set, Prop or Type."
-let rec explain_unification_error env sigma p1 p2 = function
+let explain_unification_error env sigma p1 p2 = function
| None -> mt()
| Some e ->
let rec aux p1 p2 = function
@@ -291,7 +304,8 @@ let rec explain_unification_error env sigma p1 p2 = function
else []
| MetaOccurInBody evk ->
[str "instance for " ++ quote (pr_existential_key sigma evk) ++
- strbrk " refers to a metavariable - please report your example"]
+ strbrk " refers to a metavariable - please report your example" ++
+ strbrk "at " ++ str Coq_config.wwwbugtracker ++ str "."]
| InstanceNotSameType (evk,env,t,u) ->
let t, u = pr_explicit env sigma t u in
[str "unable to find a well-typed instantiation for " ++
@@ -307,9 +321,12 @@ let rec explain_unification_error env sigma p1 p2 = function
| CannotSolveConstraint ((pb,env,t,u),e) ->
let t = Evarutil.nf_evar sigma t in
let u = Evarutil.nf_evar sigma u in
+ let env = make_all_name_different env in
(strbrk "cannot satisfy constraint " ++ pr_lconstr_env env sigma t ++
str " == " ++ pr_lconstr_env env sigma u)
:: aux t u e
+ | ProblemBeyondCapabilities ->
+ []
in
match aux p1 p2 e with
| [] -> mt ()
@@ -318,7 +335,7 @@ let rec explain_unification_error env sigma p1 p2 = function
let explain_actual_type env sigma j t reason =
let env = make_all_name_different env in
- let j = Evarutil.j_nf_betaiotaevar sigma j in
+ let j = j_nf_betaiotaevar sigma j in
let t = Reductionops.nf_betaiota sigma t in
(** Actually print *)
let pe = pr_ne_context_of (str "In environment") env sigma in
@@ -333,7 +350,7 @@ 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 = Evarutil.jv_nf_betaiotaevar sigma randl in
+ let randl = jv_nf_betaiotaevar sigma randl in
let exptyp = Evarutil.nf_evar sigma exptyp in
let actualtyp = Reductionops.nf_betaiota sigma actualtyp in
let rator = Evarutil.j_nf_evar sigma rator in
@@ -481,7 +498,7 @@ let explain_ill_formed_rec_body env sigma err names i fixenv vdefj =
let fixenv = make_all_name_different fixenv in
let pvd = pr_lconstr_env fixenv sigma vdefj.(i).uj_val in
str"Recursive definition is:" ++ spc () ++ pvd ++ str "."
- with e when Errors.noncritical e -> mt ())
+ with e when CErrors.noncritical e -> mt ())
let explain_ill_typed_rec_body env sigma i names vdefj vargs =
let vdefj = Evarutil.jv_nf_evar sigma vdefj in
@@ -499,7 +516,8 @@ let explain_cant_find_case_type env sigma c =
let c = Evarutil.nf_evar sigma c in
let env = make_all_name_different env in
let pe = pr_lconstr_env env sigma c in
- str "Cannot infer type of pattern-matching on" ++ ws 1 ++ pe ++ str "."
+ str "Cannot infer the return type of pattern-matching on" ++ ws 1 ++
+ pe ++ str "."
let explain_occur_check env sigma ev rhs =
let rhs = Evarutil.nf_evar sigma rhs in
@@ -738,7 +756,7 @@ let pr_constraints printenv env sigma evars cstrs =
let evs =
prlist
(fun (ev, evi) -> fnl () ++ pr_existential_key sigma ev ++
- str " : " ++ pr_lconstr_env env' sigma evi.evar_concl) l
+ str " : " ++ pr_lconstr_env env' sigma evi.evar_concl ++ fnl ()) l
in
h 0 (pe ++ evs ++ pr_evar_constraints cstrs)
else
@@ -775,7 +793,7 @@ let explain_unsatisfiable_constraints env sigma constr comp =
explain_typeclass_resolution env sigma info k ++ fnl () ++ cstr
let explain_pretype_error env sigma err =
- let env = Evarutil.env_nf_betaiotaevar sigma env in
+ let env = Evardefine.env_nf_betaiotaevar sigma env in
let env = make_all_name_different env in
match err with
| CantFindCaseType c -> explain_cant_find_case_type env sigma c
@@ -822,7 +840,7 @@ let explain_not_match_error = function
| ModuleTypeFieldExpected ->
strbrk "a module type is expected"
| NotConvertibleInductiveField id | NotConvertibleConstructorField id ->
- str "types given to " ++ str (Id.to_string id) ++ str " differ"
+ str "types given to " ++ pr_id id ++ str " differ"
| NotConvertibleBodyField ->
str "the body of definitions differs"
| NotConvertibleTypeField (env, typ1, typ2) ->
@@ -847,7 +865,7 @@ let explain_not_match_error = function
| RecordProjectionsExpected nal ->
(if List.length nal >= 2 then str "expected projection names are "
else str "expected projection name is ") ++
- pr_enum (function Name id -> str (Id.to_string id) | _ -> str "_") nal
+ pr_enum (function Name id -> pr_id id | _ -> str "_") nal
| NotEqualInductiveAliases ->
str "Aliases to inductive types do not match"
| NoTypeConstraintExpected ->
@@ -890,17 +908,27 @@ let explain_is_a_functor () =
str "Illegal use of a functor."
let explain_incompatible_module_types mexpr1 mexpr2 =
- str "Incompatible module types."
+ let open Declarations in
+ let rec get_arg = function
+ | NoFunctor _ -> 0
+ | MoreFunctor (_, _, ty) -> succ (get_arg ty)
+ in
+ let len1 = get_arg mexpr1.mod_type in
+ let len2 = get_arg mexpr2.mod_type in
+ if len1 <> len2 then
+ str "Incompatible module types: module expects " ++ int len2 ++
+ str " arguments, found " ++ int len1 ++ str "."
+ else str "Incompatible module types."
let explain_not_equal_module_paths mp1 mp2 =
str "Non equal modules."
let explain_no_such_label l =
- str "No such label " ++ str (Label.to_string l) ++ str "."
+ str "No such label " ++ pr_label l ++ str "."
let explain_incompatible_labels l l' =
str "Opening and closing labels are not the same: " ++
- str (Label.to_string l) ++ str " <> " ++ str (Label.to_string l') ++ str "!"
+ pr_label l ++ str " <> " ++ pr_label l' ++ str "!"
let explain_not_a_module s =
quote (str s) ++ str " is not a module."
@@ -909,19 +937,19 @@ let explain_not_a_module_type s =
quote (str s) ++ str " is not a module type."
let explain_not_a_constant l =
- quote (Label.print l) ++ str " is not a constant."
+ quote (pr_label l) ++ str " is not a constant."
let explain_incorrect_label_constraint l =
str "Incorrect constraint for label " ++
- quote (Label.print l) ++ str "."
+ quote (pr_label l) ++ str "."
let explain_generative_module_expected l =
- str "The module " ++ str (Label.to_string l) ++ str " is not generative." ++
+ str "The module " ++ pr_label l ++ str " is not generative." ++
strbrk " Only components of generative modules can be changed" ++
strbrk " using the \"with\" construct."
let explain_label_missing l s =
- str "The field " ++ str (Label.to_string l) ++ str " is missing in "
+ str "The field " ++ pr_label l ++ str " is missing in "
++ str s ++ str "."
let explain_include_restricted_functor mp =
@@ -1122,6 +1150,11 @@ let error_not_allowed_case_analysis isrec kind i =
strbrk " is not allowed for inductive definition " ++
pr_inductive (Global.env()) (fst i) ++ str "."
+let error_not_allowed_dependent_analysis isrec i =
+ str "Dependent " ++ str (if isrec then "Induction" else "Case analysis") ++
+ strbrk " is not allowed for inductive definition " ++
+ pr_inductive (Global.env()) i ++ str "."
+
let error_not_mutual_in_scheme ind ind' =
if eq_ind ind ind' then
str "The inductive type " ++ pr_inductive (Global.env()) ind ++
@@ -1153,6 +1186,8 @@ let explain_recursion_scheme_error = function
| NotAllowedCaseAnalysis (isrec,k,i) ->
error_not_allowed_case_analysis isrec k i
| NotMutualInScheme (ind,ind')-> error_not_mutual_in_scheme ind ind'
+ | NotAllowedDependentAnalysis (isrec, i) ->
+ error_not_allowed_dependent_analysis isrec i
(* Pattern-matching errors *)
@@ -1197,7 +1232,7 @@ let explain_unused_clause env pats =
let explain_non_exhaustive env pats =
str "Non exhaustive pattern-matching: no clause found for " ++
str (String.plural (List.length pats) "pattern") ++
- spc () ++ hov 0 (pr_sequence pr_cases_pattern pats)
+ spc () ++ hov 0 (prlist_with_sep pr_comma pr_cases_pattern pats)
let explain_cannot_infer_predicate env sigma typs =
let env = make_all_name_different env in
@@ -1230,77 +1265,3 @@ let explain_reduction_tactic_error = function
quote (pr_goal_concl_style_env env sigma c) ++
spc () ++ str "is not well typed." ++ fnl () ++
explain_type_error env' Evd.empty e
-
-let is_defined_ltac trace =
- let rec aux = function
- | (_, Proof_type.LtacNameCall f) :: tail ->
- not (Tacenv.is_ltac_for_ml_tactic f)
- | (_, Proof_type.LtacAtomCall _) :: tail ->
- false
- | _ :: tail -> aux tail
- | [] -> false in
- aux (List.rev trace)
-
-let explain_ltac_call_trace last trace loc =
- let calls = last :: List.rev_map snd trace in
- let pr_call ck = match ck with
- | Proof_type.LtacNotationCall kn -> quote (KerName.print kn)
- | Proof_type.LtacNameCall cst -> quote (Pptactic.pr_ltac_constant cst)
- | Proof_type.LtacMLCall t ->
- quote (Pptactic.pr_glob_tactic (Global.env()) t)
- | Proof_type.LtacVarCall (id,t) ->
- quote (Nameops.pr_id id) ++ strbrk " (bound to " ++
- Pptactic.pr_glob_tactic (Global.env()) t ++ str ")"
- | Proof_type.LtacAtomCall te ->
- quote (Pptactic.pr_glob_tactic (Global.env())
- (Tacexpr.TacAtom (Loc.ghost,te)))
- | Proof_type.LtacConstrInterp (c, { Pretyping.ltac_constrs = vars }) ->
- quote (pr_glob_constr_env (Global.env()) c) ++
- (if not (Id.Map.is_empty vars) then
- strbrk " (with " ++
- prlist_with_sep pr_comma
- (fun (id,c) ->
- pr_id id ++ str ":=" ++ Printer.pr_lconstr_under_binders c)
- (List.rev (Id.Map.bindings vars)) ++ str ")"
- else mt())
- in
- match calls with
- | [] -> mt ()
- | _ ->
- let kind_of_last_call = match List.last calls with
- | Proof_type.LtacConstrInterp _ -> ", last term evaluation failed."
- | _ -> ", last call failed."
- in
- hov 0 (str "In nested Ltac calls to " ++
- pr_enum pr_call calls ++ strbrk kind_of_last_call)
-
-let skip_extensions trace =
- let rec aux = function
- | (_,Proof_type.LtacNameCall f as tac) :: _
- when Tacenv.is_ltac_for_ml_tactic f -> [tac]
- | (_,(Proof_type.LtacNotationCall _ | Proof_type.LtacMLCall _) as tac)
- :: _ -> [tac]
- | t :: tail -> t :: aux tail
- | [] -> [] in
- List.rev (aux (List.rev trace))
-
-let extract_ltac_trace trace eloc =
- let trace = skip_extensions trace in
- let (loc,c),tail = List.sep_last trace in
- if is_defined_ltac trace then
- (* We entered a user-defined tactic,
- we display the trace with location of the call *)
- let msg = hov 0 (explain_ltac_call_trace c tail eloc ++ fnl()) in
- Some msg, loc
- else
- (* We entered a primitive tactic, we don't display trace but
- report on the finest location *)
- let best_loc =
- if not (Loc.is_ghost eloc) then eloc else
- (* trace is with innermost call coming first *)
- let rec aux = function
- | (loc,_)::tail when not (Loc.is_ghost loc) -> loc
- | _::tail -> aux tail
- | [] -> Loc.ghost in
- aux trace in
- None, best_loc
diff --git a/toplevel/himsg.mli b/toplevel/himsg.mli
index 3ef98380..ced54fd2 100644
--- a/toplevel/himsg.mli
+++ b/toplevel/himsg.mli
@@ -36,9 +36,6 @@ val explain_pattern_matching_error :
val explain_reduction_tactic_error :
Tacred.reduction_tactic_error -> std_ppcmds
-val extract_ltac_trace :
- Proof_type.ltac_trace -> Loc.t -> std_ppcmds option * Loc.t
-
val explain_module_error : Modops.module_typing_error -> std_ppcmds
val explain_module_internalization_error :
diff --git a/toplevel/ind_tables.ml b/toplevel/ind_tables.ml
index 35717ed6..6d57a21d 100644
--- a/toplevel/ind_tables.ml
+++ b/toplevel/ind_tables.ml
@@ -18,7 +18,7 @@ open Libobject
open Nameops
open Declarations
open Term
-open Errors
+open CErrors
open Util
open Declare
open Entries
diff --git a/toplevel/indschemes.ml b/toplevel/indschemes.ml
index c4ac0e41..e8ea617f 100644
--- a/toplevel/indschemes.ml
+++ b/toplevel/indschemes.ml
@@ -15,7 +15,7 @@
declaring new schemes *)
open Pp
-open Errors
+open CErrors
open Util
open Names
open Declarations
@@ -38,6 +38,7 @@ open Ind_tables
open Auto_ind_decl
open Eqschemes
open Elimschemes
+open Context.Rel.Declaration
(* Flags governing automatic synthesis of schemes *)
@@ -149,16 +150,18 @@ let alarm what internal msg =
| UserAutomaticRequest
| InternalTacticRequest ->
(if debug then
- msg_warning
- (hov 0 msg ++ fnl () ++ what ++ str " not defined."))
- | _ -> errorlabstrm "" msg
+ Feedback.msg_debug
+ (hov 0 msg ++ fnl () ++ what ++ str " not defined.")); None
+ | _ -> Some msg
let try_declare_scheme what f internal names kn =
try f internal names kn
- with
+ with e ->
+ let e = CErrors.push e in
+ let msg = match fst e with
| ParameterWithoutEquality cst ->
alarm what internal
- (str "Boolean equality not found for parameter " ++ pr_con cst ++
+ (str "Boolean equality not found for parameter " ++ Printer.pr_global cst ++
str".")
| InductiveWithProduct ->
alarm what internal
@@ -183,9 +186,14 @@ let try_declare_scheme what f internal names kn =
| DecidabilityMutualNotSupported ->
alarm what internal
(str "Decidability lemma for mutual inductive types not supported.")
- | e when Errors.noncritical e ->
+ | e when CErrors.noncritical e ->
alarm what internal
- (str "Unexpected error during scheme creation: " ++ Errors.print e)
+ (str "Unexpected error during scheme creation: " ++ CErrors.print e)
+ | _ -> iraise e
+ in
+ match msg with
+ | None -> ()
+ | Some msg -> iraise (UserError ("", msg), snd e)
let beq_scheme_msg mind =
let mib = Global.lookup_mind mind in
@@ -209,7 +217,11 @@ let declare_beq_scheme = declare_beq_scheme_with []
let declare_one_case_analysis_scheme ind =
let (mib,mip) = Global.lookup_inductive ind in
let kind = inductive_sort_family mip in
- let dep = if kind == InProp then case_scheme_kind_from_prop else case_dep_scheme_kind_from_type in
+ let dep =
+ if kind == InProp then case_scheme_kind_from_prop
+ else if not (Inductiveops.has_dependent_elim mib) then
+ case_scheme_kind_from_type
+ else case_dep_scheme_kind_from_type in
let kelim = elim_sorts (mib,mip) in
(* in case the inductive has a type elimination, generates only one
induction scheme, the other ones share the same code with the
@@ -229,15 +241,23 @@ let kinds_from_type =
InProp,ind_dep_scheme_kind_from_type;
InSet,rec_dep_scheme_kind_from_type]
+let nondep_kinds_from_type =
+ [InType,rect_scheme_kind_from_type;
+ InProp,ind_scheme_kind_from_type;
+ InSet,rec_scheme_kind_from_type]
+
let declare_one_induction_scheme ind =
let (mib,mip) = Global.lookup_inductive ind in
let kind = inductive_sort_family mip in
let from_prop = kind == InProp in
+ let depelim = Inductiveops.has_dependent_elim mib in
let kelim = elim_sorts (mib,mip) in
let elims =
List.map_filter (fun (sort,kind) ->
if Sorts.List.mem sort kelim then Some kind else None)
- (if from_prop then kinds_from_prop else kinds_from_type) in
+ (if from_prop then kinds_from_prop
+ else if depelim then kinds_from_type
+ else nondep_kinds_from_type) in
List.iter (fun kind -> ignore (define_individual_scheme kind UserAutomaticRequest None ind))
elims
@@ -270,7 +290,7 @@ let try_declare_eq_decidability kn =
let declare_eq_decidability = declare_eq_decidability_scheme_with []
let ignore_error f x =
- try ignore (f x) with e when Errors.noncritical e -> ()
+ try ignore (f x) with e when CErrors.noncritical e -> ()
let declare_rewriting_schemes ind =
if Hipattern.is_inductive_equality ind then begin
@@ -287,15 +307,20 @@ let declare_rewriting_schemes ind =
(define_individual_scheme rew_l2r_forward_dep_scheme_kind UserAutomaticRequest None) ind
end
+let warn_cannot_build_congruence =
+ CWarnings.create ~name:"cannot-build-congruence" ~category:"schemes"
+ (fun () ->
+ strbrk "Cannot build congruence scheme because eq is not found")
+
let declare_congr_scheme ind =
if Hipattern.is_equality_type (mkInd ind) then begin
if
try Coqlib.check_required_library Coqlib.logic_module_name; true
- with e when Errors.noncritical e -> false
+ with e when CErrors.noncritical e -> false
then
ignore (define_individual_scheme congr_scheme_kind UserAutomaticRequest None ind)
else
- msg_warning (strbrk "Cannot build congruence scheme because eq is not found")
+ warn_cannot_build_congruence ()
end
let declare_sym_scheme ind =
@@ -379,7 +404,6 @@ let do_mutual_induction_scheme lnamedepindsort =
let sigma, listdecl = Indrec.build_mutual_induction_scheme env0 sigma lrecspec in
let declare decl fi lrecref =
let decltype = Retyping.get_type_of env0 sigma decl in
- (* let decltype = refresh_universes decltype in *)
let proof_output = Future.from_val ((decl,Univ.ContextSet.empty),Safe_typing.empty_private_constants) in
let cst = define fi UserIndividualRequest sigma proof_output (Some decltype) in
ConstRef cst :: lrecref
@@ -463,7 +487,7 @@ let build_combined_scheme env schemes =
in
let ctx, _ =
list_split_rev_at prods
- (List.rev_map (fun (x, y) -> x, None, y) ctx) in
+ (List.rev_map (fun (x, y) -> LocalAssum (x, y)) ctx) in
let typ = it_mkProd_wo_LetIn concl_typ ctx in
let body = it_mkLambda_or_LetIn concl_bod ctx in
(body, typ)
@@ -489,7 +513,8 @@ 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) then
+ if !elim_flag && (mib.mind_finite <> BiFinite || !bifinite_elim_flag)
+ && mib.mind_typing_flags.check_guarded then
declare_induction_schemes kn;
if !case_flag then map_inductive_block declare_one_case_analysis_scheme kn n;
if is_eq_flag() then try_declare_beq_scheme kn;
diff --git a/toplevel/locality.ml b/toplevel/locality.ml
index ef789aa5..154f787e 100644
--- a/toplevel/locality.ml
+++ b/toplevel/locality.ml
@@ -18,7 +18,7 @@ let check_locality locality_flag =
match locality_flag with
| Some b ->
let s = if b then "Local" else "Global" in
- Errors.errorlabstrm "Locality.check_locality"
+ CErrors.errorlabstrm "Locality.check_locality"
(str "This command does not support the \"" ++ str s ++ str "\" prefix.")
| None -> ()
@@ -26,17 +26,22 @@ let check_locality 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 ->
- Errors.error "Cannot be simultaneously Local and Global."
+ CErrors.error "Cannot be simultaneously Local and Global."
| Some true when local ->
- Errors.error "Use only prefix \"Local\"."
+ CErrors.error "Use only prefix \"Local\"."
| None ->
if local then begin
- Pp.msg_warning (Pp.str "Obsolete syntax: use \"Local\" as a prefix.");
- Some true
+ warn_deprecated_local_syntax ();
+ Some true
end else
None
| Some b -> Some b in
@@ -61,7 +66,7 @@ let enforce_locality_exp locality_flag local =
| None, Some local -> local
| Some b, None -> local_of_bool b
| None, None -> Decl_kinds.Global
- | Some _, Some _ -> Errors.error "Local non allowed in this case"
+ | Some _, Some _ -> CErrors.error "Local non allowed in this case"
(* For commands whose default is to not discharge but to export:
Global in sections forces discharge, Global not in section is the default;
@@ -82,7 +87,7 @@ let enforce_section_locality locality_flag local =
let make_module_locality = function
| Some false ->
if Lib.sections_are_opened () then
- Errors.error
+ CErrors.error
"This command does not support the Global option in sections.";
false
| Some true -> true
diff --git a/toplevel/metasyntax.ml b/toplevel/metasyntax.ml
index ae82b64e..008d5cf9 100644
--- a/toplevel/metasyntax.ml
+++ b/toplevel/metasyntax.ml
@@ -8,7 +8,7 @@
open Pp
open Flags
-open Errors
+open CErrors
open Util
open Names
open Constrexpr
@@ -23,7 +23,6 @@ open Vernacexpr
open Pcoq
open Libnames
open Tok
-open Egramml
open Egramcoq
open Notation
open Nameops
@@ -31,7 +30,7 @@ open Nameops
(**********************************************************************)
(* Tokens *)
-let cache_token (_,s) = Lexer.add_keyword s
+let cache_token (_,s) = CLexer.add_keyword s
let inToken : string -> obj =
declare_object {(default_object "TOKEN") with
@@ -43,161 +42,6 @@ let inToken : string -> obj =
let add_token_obj s = Lib.add_anonymous_leaf (inToken s)
(**********************************************************************)
-(* Tactic Notation *)
-
-let interp_prod_item lev = function
- | TacTerm s -> GramTerminal s
- | TacNonTerm (loc, nt, po) ->
- let sep = match po with Some (_,sep) -> sep | _ -> "" in
- let (etyp, e) = interp_entry_name true (Some lev) nt sep in
- GramNonTerminal (loc, etyp, e, Option.map fst po)
-
-let make_terminal_status = function
- | GramTerminal s -> Some s
- | GramNonTerminal _ -> None
-
-let rec make_tags = function
- | GramTerminal s :: l -> make_tags l
- | GramNonTerminal (loc, etyp, _, po) :: l -> etyp :: make_tags l
- | [] -> []
-
-let make_fresh_key =
- let id = Summary.ref ~name:"TACTIC-NOTATION-COUNTER" 0 in
- fun () ->
- let cur = incr id; !id in
- let lbl = Id.of_string ("_" ^ string_of_int cur) in
- let kn = Lib.make_kn lbl in
- let (mp, dir, _) = KerName.repr kn in
- (** We embed the full path of the kernel name in the label so that the
- identifier should be unique. This ensures that including two modules
- together won't confuse the corresponding labels. *)
- let lbl = Id.of_string_soft (Printf.sprintf "%s#%s#%i"
- (ModPath.to_string mp) (DirPath.to_string dir) cur)
- in
- KerName.make mp dir (Label.of_id lbl)
-
-type tactic_grammar_obj = {
- tacobj_key : KerName.t;
- tacobj_local : locality_flag;
- tacobj_tacgram : tactic_grammar;
- tacobj_tacpp : Pptactic.pp_tactic;
- tacobj_body : Tacexpr.glob_tactic_expr
-}
-
-let check_key key =
- if Tacenv.check_alias key then
- error "Conflicting tactic notations keys. This can happen when including \
- twice the same module."
-
-let cache_tactic_notation (_, tobj) =
- let key = tobj.tacobj_key in
- let () = check_key key in
- Tacenv.register_alias key tobj.tacobj_body;
- Egramcoq.extend_tactic_grammar key tobj.tacobj_tacgram;
- Pptactic.declare_notation_tactic_pprule key tobj.tacobj_tacpp
-
-let open_tactic_notation i (_, tobj) =
- let key = tobj.tacobj_key in
- if Int.equal i 1 && not tobj.tacobj_local then
- Egramcoq.extend_tactic_grammar key tobj.tacobj_tacgram
-
-let load_tactic_notation i (_, tobj) =
- let key = tobj.tacobj_key in
- let () = check_key key in
- (** Only add the printing and interpretation rules. *)
- Tacenv.register_alias key tobj.tacobj_body;
- Pptactic.declare_notation_tactic_pprule key tobj.tacobj_tacpp;
- if Int.equal i 1 && not tobj.tacobj_local then
- Egramcoq.extend_tactic_grammar key tobj.tacobj_tacgram
-
-let subst_tactic_notation (subst, tobj) =
- { tobj with
- tacobj_key = Mod_subst.subst_kn subst tobj.tacobj_key;
- tacobj_body = Tacsubst.subst_tactic subst tobj.tacobj_body;
- }
-
-let classify_tactic_notation tacobj = Substitute tacobj
-
-let inTacticGrammar : tactic_grammar_obj -> obj =
- declare_object {(default_object "TacticGrammar") with
- open_function = open_tactic_notation;
- load_function = load_tactic_notation;
- cache_function = cache_tactic_notation;
- subst_function = subst_tactic_notation;
- classify_function = classify_tactic_notation}
-
-let cons_production_parameter l = function
- | GramTerminal _ -> l
- | GramNonTerminal (_,_,_,ido) -> Option.List.cons ido l
-
-let add_tactic_notation (local,n,prods,e) =
- let prods = List.map (interp_prod_item n) prods in
- let tags = make_tags prods in
- let pprule = {
- Pptactic.pptac_args = tags;
- pptac_prods = (n, List.map make_terminal_status prods);
- } in
- let ids = List.fold_left cons_production_parameter [] prods in
- let tac = Tacintern.glob_tactic_env ids (Global.env()) e in
- let parule = {
- tacgram_level = n;
- tacgram_prods = prods;
- } in
- let tacobj = {
- tacobj_key = make_fresh_key ();
- tacobj_local = local;
- tacobj_tacgram = parule;
- tacobj_tacpp = pprule;
- tacobj_body = tac;
- } in
- Lib.add_anonymous_leaf (inTacticGrammar tacobj)
-
-(**********************************************************************)
-(* ML Tactic entries *)
-
-type atomic_entry = string * Genarg.glob_generic_argument list option
-
-type ml_tactic_grammar_obj = {
- mltacobj_name : Tacexpr.ml_tactic_name;
- (** ML-side unique name *)
- mltacobj_prod : grammar_prod_item list list;
- (** Grammar rules generating the ML tactic. *)
-}
-
-(** ML tactic notations whose use can be restricted to an identifier are added
- as true Ltac entries. *)
-let extend_atomic_tactic name entries =
- let add_atomic (id, args) = match args with
- | None -> ()
- | Some args ->
- let body = Tacexpr.TacML (Loc.ghost, name, args) in
- Tacenv.register_ltac false false (Names.Id.of_string id) body
- in
- List.iter add_atomic entries
-
-let cache_ml_tactic_notation (_, obj) =
- extend_ml_tactic_grammar obj.mltacobj_name obj.mltacobj_prod
-
-let open_ml_tactic_notation i obj =
- if Int.equal i 1 then cache_ml_tactic_notation obj
-
-let inMLTacticGrammar : ml_tactic_grammar_obj -> obj =
- declare_object { (default_object "MLTacticGrammar") with
- open_function = open_ml_tactic_notation;
- cache_function = cache_ml_tactic_notation;
- classify_function = (fun o -> Substitute o);
- subst_function = (fun (_, o) -> o);
- }
-
-let add_ml_tactic_notation name prods atomic =
- let obj = {
- mltacobj_name = name;
- mltacobj_prod = prods;
- } in
- Lib.add_anonymous_leaf (inMLTacticGrammar obj);
- extend_atomic_tactic name atomic
-
-(**********************************************************************)
(* Printing grammar entries *)
let entry_buf = Buffer.create 64
@@ -340,7 +184,7 @@ let parse_format ((loc, str) : lstring) =
else
error "Empty format."
with reraise ->
- let (e, info) = Errors.push reraise in
+ let (e, info) = CErrors.push reraise in
let info = Loc.add_loc info loc in
iraise (e, info)
@@ -428,7 +272,7 @@ let rec interp_list_parser hd = function
(* To protect alphabetic tokens and quotes from being seen as variables *)
let quote_notation_token x =
let n = String.length x in
- let norm = Lexer.is_ident x in
+ let norm = CLexer.is_ident x in
if (n > 0 && norm) || (n > 2 && x.[0] == '\'') then "'"^x^"'"
else x
@@ -436,7 +280,7 @@ let rec raw_analyze_notation_tokens = function
| [] -> []
| String ".." :: sl -> NonTerminal ldots_var :: raw_analyze_notation_tokens sl
| String "_" :: _ -> error "_ must be quoted."
- | String x :: sl when Lexer.is_ident x ->
+ | 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
@@ -540,10 +384,15 @@ let add_break_if_none n = function
| l -> UnpCut (PpBrk(n,0)) :: l
let check_open_binder isopen sl m =
+ let pr_token = function
+ | Terminal s -> str s
+ | Break n -> str "â£"
+ | _ -> assert false
+ in
if isopen && not (List.is_empty sl) then
errorlabstrm "" (str "as " ++ pr_id m ++
str " is a non-closed binder, no such \"" ++
- prlist_with_sep spc (function Terminal s -> str s | _ -> assert false) sl
+ prlist_with_sep spc pr_token sl
++ strbrk "\" is allowed to occur.")
(* Heuristics for building default printing rules *)
@@ -719,8 +568,8 @@ let is_not_small_constr = function
let rec define_keywords_aux = function
| GramConstrNonTerminal(e,Some _) as n1 :: GramConstrTerminal(IDENT k) :: l
when is_not_small_constr e ->
- Flags.if_verbose msg_info (str "Identifier '" ++ str k ++ str "' now a keyword");
- Lexer.add_keyword k;
+ Flags.if_verbose Feedback.msg_info (str "Identifier '" ++ str k ++ str "' now a keyword");
+ CLexer.add_keyword k;
n1 :: GramConstrTerminal(KEYWORD k) :: define_keywords_aux l
| n :: l -> n :: define_keywords_aux l
| [] -> []
@@ -728,8 +577,8 @@ let rec define_keywords_aux = function
(* Ensure that IDENT articulation terminal symbols are keywords *)
let define_keywords = function
| GramConstrTerminal(IDENT k)::l ->
- Flags.if_verbose msg_info (str "Identifier '" ++ str k ++ str "' now a keyword");
- Lexer.add_keyword k;
+ Flags.if_verbose Feedback.msg_info (str "Identifier '" ++ str k ++ str "' now a keyword");
+ CLexer.add_keyword k;
GramConstrTerminal(KEYWORD k) :: define_keywords_aux l
| l -> define_keywords_aux l
@@ -758,12 +607,12 @@ let make_production etyps symbols =
let typ = List.assoc m etyps in
distribute [GramConstrNonTerminal (typ, Some m)] ll
| Terminal s ->
- distribute [GramConstrTerminal (Lexer.terminal s)] ll
+ distribute [GramConstrTerminal (CLexer.terminal s)] ll
| Break _ ->
ll
| SProdList (x,sl) ->
let tkl = List.flatten
- (List.map (function Terminal s -> [Lexer.terminal s]
+ (List.map (function Terminal s -> [CLexer.terminal s]
| Break _ -> []
| _ -> anomaly (Pp.str "Found a non terminal token in recursive notation separator")) sl) in
match List.assoc x etyps with
@@ -824,24 +673,32 @@ type syntax_extension = {
synext_notgram : notation_grammar;
synext_unparsing : unparsing list;
synext_extra : (string * string) list;
+ synext_compat : Flags.compat_version option;
}
+let is_active_compat = function
+| None -> true
+| Some v -> 0 <= Flags.version_compare v !Flags.compat_version
+
type syntax_extension_obj = locality_flag * syntax_extension list
let cache_one_syntax_extension se =
let ntn = se.synext_notation in
let prec = se.synext_level in
+ let onlyprint = se.synext_notgram.notgram_onlyprinting in
try
let oldprec = Notation.level_of_notation ntn in
if not (Notation.level_eq prec oldprec) then error_incompatible_level ntn oldprec prec
with Not_found ->
- (* Reserve the notation level *)
- Notation.declare_notation_level ntn prec;
- (* Declare the parsing rule *)
- Egramcoq.extend_constr_grammar prec se.synext_notgram;
- (* Declare the printing rule *)
- Notation.declare_notation_printing_rule ntn
- ~extra:se.synext_extra (se.synext_unparsing, fst prec)
+ if is_active_compat se.synext_compat then begin
+ (* Reserve the notation level *)
+ Notation.declare_notation_level ntn prec;
+ (* Declare the parsing rule *)
+ if not onlyprint then Egramcoq.extend_constr_grammar prec se.synext_notgram;
+ (* Declare the notation rule *)
+ Notation.declare_notation_rule ntn
+ ~extra:se.synext_extra (se.synext_unparsing, fst prec) se.synext_notgram
+ end
let cache_syntax_extension (_, (_, sy)) =
List.iter cache_one_syntax_extension sy
@@ -874,9 +731,11 @@ let inSyntaxExtension : syntax_extension_obj -> obj =
let interp_modifiers modl =
let onlyparsing = ref false in
+ let onlyprinting = ref false in
+ let compat = ref None in
let rec interp assoc level etyps format extra = function
| [] ->
- (assoc,level,etyps,!onlyparsing,format,extra)
+ (assoc,level,etyps,!onlyparsing,!onlyprinting,!compat,format,extra)
| SetEntryType (s,typ) :: l ->
let id = Id.of_string s in
if Id.List.mem_assoc id etyps then
@@ -898,9 +757,15 @@ let interp_modifiers modl =
| SetAssoc a :: l ->
if not (Option.is_empty assoc) then error"An associativity is given more than once.";
interp (Some a) level etyps format extra l
- | SetOnlyParsing _ :: l ->
+ | SetOnlyParsing :: l ->
onlyparsing := true;
interp assoc level etyps format extra l
+ | SetOnlyPrinting :: l ->
+ onlyprinting := true;
+ interp assoc level etyps format extra l
+ | SetCompatVersion v :: l ->
+ compat := Some v;
+ interp assoc level etyps format extra l
| SetFormat ("text",s) :: l ->
if not (Option.is_empty format) then error "A format is given more than once.";
interp assoc level etyps (Some s) extra l
@@ -909,7 +774,7 @@ let interp_modifiers modl =
in interp None None [] None [] modl
let check_infix_modifiers modifiers =
- let (assoc,level,t,b,fmt,extra) = interp_modifiers modifiers in
+ let (_, _, t, _, _, _, _, _) = interp_modifiers modifiers in
if not (List.is_empty t) then
error "Explicit entry level or type unexpected in infix notation."
@@ -920,13 +785,25 @@ let check_useless_entry_types recvars mainvars etyps =
(pr_id x ++ str " is unbound in the notation.")
| _ -> ()
-let no_syntax_modifiers = function
- | [] | [SetOnlyParsing _] -> true
- | _ -> false
+let not_a_syntax_modifier = function
+| SetOnlyParsing -> true
+| SetOnlyPrinting -> true
+| SetCompatVersion _ -> true
+| _ -> false
-let is_only_parsing = function
- | [SetOnlyParsing _] -> true
- | _ -> false
+let no_syntax_modifiers mods = List.for_all not_a_syntax_modifier mods
+
+let is_only_parsing mods =
+ let test = function SetOnlyParsing -> true | _ -> false in
+ List.exists test mods
+
+let is_only_printing mods =
+ let test = function SetOnlyPrinting -> true | _ -> false in
+ List.exists test mods
+
+let get_compat_version mods =
+ let test = function SetCompatVersion v -> Some v | _ -> None in
+ try Some (List.find_map test mods) with Not_found -> None
(* Compute precedences from modifiers (or find default ones) *)
@@ -973,11 +850,12 @@ let make_internalization_vars recvars mainvars typs =
let extratyps = List.map (fun (x,y) -> (y,List.assoc x maintyps)) recvars in
maintyps @ extratyps
-let make_interpretation_type isrec = function
+let make_interpretation_type isrec isonlybinding = function
| NtnInternTypeConstr when isrec -> NtnTypeConstrList
- | NtnInternTypeConstr | NtnInternTypeIdent -> NtnTypeConstr
+ | NtnInternTypeConstr | NtnInternTypeIdent ->
+ if isonlybinding then NtnTypeOnlyBinder else NtnTypeConstr
| NtnInternTypeBinder when isrec -> NtnTypeBinderList
- | NtnInternTypeBinder -> error "Type not allowed in recursive notation."
+ | NtnInternTypeBinder -> error "Type binder is only for use in recursive notations for binders."
let make_interpretation_vars recvars allvars =
let eq_subscope (sc1, l1) (sc2, l2) =
@@ -985,45 +863,65 @@ let make_interpretation_vars recvars allvars =
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 (sc, typ) ->
- (sc, make_interpretation_type (Id.List.mem_assoc x recvars) typ)) mainvars
+ Id.Map.mapi (fun x (isonlybinding, sc, typ) ->
+ (sc, make_interpretation_type (Id.List.mem_assoc x recvars) isonlybinding typ)) mainvars
let check_rule_productivity l =
- if List.for_all (function NonTerminal _ -> true | _ -> false) l then
+ if List.for_all (function NonTerminal _ | Break _ -> true | _ -> false) l then
error "A notation must include at least one symbol.";
if (match l with SProdList _ :: _ -> true | _ -> false) then
error "A recursive notation must start with at least one symbol."
-let is_not_printable onlyparse noninjective = function
+let warn_notation_bound_to_variable =
+ CWarnings.create ~name:"notation-bound-to-variable" ~category:"parsing"
+ (fun () ->
+ strbrk "This notation will not be used for printing as it is bound to a single variable.")
+
+let warn_non_reversible_notation =
+ CWarnings.create ~name:"non-reversible-notation" ~category:"parsing"
+ (fun () ->
+ strbrk "This notation will not be used for printing as it is not reversible.")
+
+let is_not_printable onlyparse nonreversible = function
| NVar _ ->
- let () = if not onlyparse then
- msg_warning (strbrk "This notation will not be used for printing as it is bound to a single variable.")
- in
+ if not onlyparse then warn_notation_bound_to_variable ();
true
| _ ->
- if not onlyparse && noninjective then
- let () = msg_warning (strbrk "This notation will not be used for printing as it is not reversible.") in
- true
+ if not onlyparse && nonreversible then
+ (warn_non_reversible_notation (); true)
else onlyparse
let find_precedence lev etyps symbols =
- match symbols with
- | NonTerminal x :: _ ->
+ let first_symbol =
+ let rec aux = function
+ | Break _ :: t -> aux t
+ | h :: t -> h
+ | [] -> assert false (* rule is known to be productive *) in
+ aux symbols in
+ let last_is_terminal () =
+ let rec aux b = function
+ | Break _ :: t -> aux b t
+ | Terminal _ :: t -> aux true t
+ | _ :: t -> aux false t
+ | [] -> b in
+ aux false symbols in
+ match first_symbol with
+ | NonTerminal x ->
(try match List.assoc x etyps with
| ETConstr _ ->
error "The level of the leftmost non-terminal cannot be changed."
| ETName | ETBigint | ETReference ->
begin match lev with
| None ->
- ([msg_info,strbrk "Setting notation at level 0."],0)
+ ([Feedback.msg_info ?loc:None ,strbrk "Setting notation at level 0."],0)
| Some 0 ->
([],0)
| _ ->
@@ -1039,11 +937,9 @@ let find_precedence lev etyps symbols =
if Option.is_empty lev then
error "A left-recursive notation must have an explicit level."
else [],Option.get lev)
- | Terminal _ ::l when
- (match List.last symbols with Terminal _ -> true |_ -> false)
- ->
+ | Terminal _ when last_is_terminal () ->
if Option.is_empty lev then
- ([msg_info,strbrk "Setting notation at level 0."], 0)
+ ([Feedback.msg_info ?loc:None ,strbrk "Setting notation at level 0."], 0)
else [],Option.get lev
| _ ->
if Option.is_empty lev then error "Cannot determine the level.";
@@ -1055,6 +951,10 @@ let check_curly_brackets_notation_exists () =
error "Notations involving patterns of the form \"{ _ }\" are treated \n\
specially and require that the notation \"{ _ }\" is already reserved."
+let warn_skip_spaces_curly =
+ CWarnings.create ~name:"skip-spaces-curly" ~category:"parsing"
+ (fun () ->strbrk "Skipping spaces inside curly brackets")
+
(* Remove patterns of the form "{ _ }", unless it is the "{ _ }" notation *)
let remove_curly_brackets l =
let rec skip_break acc = function
@@ -1069,9 +969,10 @@ let remove_curly_brackets l =
let br',next' = skip_break [] l' in
(match next' with
| Terminal "}" as t2 :: l'' as l1 ->
- if not (List.equal Notation.symbol_eq l l0) || not (List.equal Notation.symbol_eq l' l1) then
- msg_warning (strbrk "Skipping spaces inside curly brackets");
- if deb && List.is_empty l'' then [t1;x;t2] else begin
+ if not (List.equal Notation.symbol_eq l l0) ||
+ not (List.equal Notation.symbol_eq l' l1) then
+ warn_skip_spaces_curly ();
+ if deb && List.is_empty l'' then [t1;x;t2] else begin
check_curly_brackets_notation_exists ();
x :: aux false l''
end
@@ -1081,7 +982,7 @@ let remove_curly_brackets l =
in aux true l
let compute_syntax_data df modifiers =
- let (assoc,n,etyps,onlyparse,fmt,extra) = interp_modifiers modifiers in
+ let (assoc,n,etyps,onlyparse,onlyprint,compat,fmt,extra) = interp_modifiers modifiers in
let assoc = match assoc with None -> (* default *) Some NonA | a -> a in
let toks = split_notation_string df in
let (recvars,mainvars,symbols) = analyze_notation_tokens toks in
@@ -1107,18 +1008,18 @@ let compute_syntax_data df modifiers =
let sy_data = (n,sy_typs,symbols',fmt) in
let sy_fulldata = (i_typs,ntn_for_grammar,prec,need_squash,sy_data) in
let df' = ((Lib.library_dp(),Lib.current_dirpath true),df) in
- let i_data = (onlyparse,recvars,mainvars,(ntn_for_interp,df')) in
+ let i_data = (onlyparse,onlyprint,compat,recvars,mainvars,(ntn_for_interp,df')) in
(* Return relevant data for interpretation and for parsing/printing *)
(msgs,i_data,i_typs,sy_fulldata,extra)
let compute_pure_syntax_data df mods =
- let (msgs,(onlyparse,_,_,_),_,sy_data,extra) = compute_syntax_data df mods in
+ let (msgs,(onlyparse,onlyprint,_,_,_,_),_,sy_data,extra) = compute_syntax_data df mods in
let msgs =
if onlyparse then
- (msg_warning,
+ (Feedback.msg_warning ?loc:None,
strbrk "The only parsing modifier has no effect in Reserved Notation.")::msgs
else msgs in
- msgs, sy_data, extra
+ msgs, sy_data, extra, onlyprint
(**********************************************************************)
(* Registration of notations interpretation *)
@@ -1128,6 +1029,8 @@ type notation_obj = {
notobj_scope : scope_name option;
notobj_interp : interpretation;
notobj_onlyparse : bool;
+ notobj_onlyprint : bool;
+ notobj_compat : Flags.compat_version option;
notobj_notation : notation * notation_location;
}
@@ -1138,9 +1041,12 @@ let open_notation i (_, nobj) =
let scope = nobj.notobj_scope in
let (ntn, df) = nobj.notobj_notation in
let pat = nobj.notobj_interp in
- if Int.equal i 1 && not (Notation.exists_notation_in_scope scope ntn pat) then begin
+ let fresh = not (Notation.exists_notation_in_scope scope ntn pat) in
+ let active = is_active_compat nobj.notobj_compat in
+ if Int.equal i 1 && fresh && active then begin
(* Declare the interpretation *)
- Notation.declare_notation_interpretation ntn scope pat df;
+ let onlyprint = nobj.notobj_onlyprint in
+ let () = Notation.declare_notation_interpretation ntn scope pat df ~onlyprint in
(* Declare the uninterpretation *)
if not nobj.notobj_onlyparse then
Notation.declare_uninterpretation (NotationRule (scope, ntn)) pat
@@ -1170,7 +1076,7 @@ let with_lib_stk_protection f x =
let fs = Lib.freeze `No in
try let a = f x in Lib.unfreeze fs; a
with reraise ->
- let reraise = Errors.push reraise in
+ let reraise = CErrors.push reraise in
let () = Lib.unfreeze fs in
iraise reraise
@@ -1187,7 +1093,10 @@ let contract_notation ntn =
let rec aux ntn i =
if i <= String.length ntn - 5 then
let ntn' =
- if String.is_sub "{ _ }" ntn i then
+ if String.is_sub "{ _ }" ntn i &&
+ (i = 0 || ntn.[i-1] = ' ') &&
+ (i = String.length ntn - 5 || ntn.[i+5] = ' ')
+ then
String.sub ntn 0 i ^ "_" ^
String.sub ntn (i+5) (String.length ntn -i-5)
else ntn in
@@ -1202,12 +1111,14 @@ let recover_syntax ntn =
let prec = Notation.level_of_notation ntn in
let pp_rule,_ = Notation.find_notation_printing_rule ntn in
let pp_extra_rules = Notation.find_notation_extra_printing_rules ntn in
- let pa_rule = Egramcoq.recover_constr_grammar ntn prec in
+ let pa_rule = Notation.find_notation_parsing_rules ntn in
{ synext_level = prec;
synext_notation = ntn;
synext_notgram = pa_rule;
synext_unparsing = pp_rule;
- synext_extra = pp_extra_rules }
+ synext_extra = pp_extra_rules;
+ synext_compat = None;
+ }
with Not_found ->
raise NoSyntaxRule
@@ -1220,27 +1131,29 @@ let recover_notation_syntax rawntn =
let sy = recover_syntax ntn in
let need_squash = not (String.equal ntn rawntn) in
let rules = if need_squash then recover_squash_syntax sy else [sy] in
- sy.synext_notgram.notgram_typs, rules
+ sy.synext_notgram.notgram_typs, rules, sy.synext_notgram.notgram_onlyprinting
(**********************************************************************)
(* Main entry point for building parsing and printing rules *)
-let make_pa_rule i_typs (n,typs,symbols,_) ntn =
+let make_pa_rule i_typs (n,typs,symbols,_) ntn onlyprint =
let assoc = recompute_assoc typs in
let prod = make_production typs symbols in
{ notgram_level = n;
notgram_assoc = assoc;
notgram_notation = ntn;
notgram_prods = prod;
- notgram_typs = i_typs; }
+ notgram_typs = i_typs;
+ notgram_onlyprinting = onlyprint;
+ }
let make_pp_rule (n,typs,symbols,fmt) =
match fmt with
| None -> [UnpBox (PpHOVB 0, make_hunks typs symbols n)]
| Some fmt -> hunks_of_format (n, List.split typs) (symbols, parse_format fmt)
-let make_syntax_rules (i_typs,ntn,prec,need_squash,sy_data) extra =
- let pa_rule = make_pa_rule i_typs sy_data ntn in
+let make_syntax_rules (i_typs,ntn,prec,need_squash,sy_data) extra onlyprint compat =
+ let pa_rule = make_pa_rule i_typs sy_data ntn onlyprint in
let pp_rule = make_pp_rule sy_data in
let sy = {
synext_level = prec;
@@ -1248,6 +1161,7 @@ let make_syntax_rules (i_typs,ntn,prec,need_squash,sy_data) extra =
synext_notgram = pa_rule;
synext_unparsing = pp_rule;
synext_extra = extra;
+ synext_compat = compat;
} in
(* By construction, the rule for "{ _ }" is declared, but we need to
redeclare it because the file where it is declared needs not be open
@@ -1263,26 +1177,27 @@ let to_map l =
let add_notation_in_scope local df c mods scope =
let (msgs,i_data,i_typs,sy_data,extra) = compute_syntax_data df mods in
- (* Prepare the parsing and printing rules *)
- let sy_rules = make_syntax_rules sy_data extra in
(* Prepare the interpretation *)
- let (onlyparse, recvars,mainvars, df') = i_data in
+ let (onlyparse, onlyprint, compat, recvars,mainvars, df') = i_data in
+ (* Prepare the parsing and printing rules *)
+ let sy_rules = make_syntax_rules sy_data extra onlyprint compat in
let i_vars = make_internalization_vars recvars mainvars i_typs in
let nenv = {
ninterp_var_type = to_map i_vars;
ninterp_rec_vars = to_map recvars;
- ninterp_only_parse = false;
} in
- let (acvars, ac) = interp_notation_constr nenv c in
+ let (acvars, ac, reversible) = interp_notation_constr nenv c in
let interp = make_interpretation_vars recvars acvars in
let map (x, _) = try Some (x, Id.Map.find x interp) with Not_found -> None in
- let onlyparse = is_not_printable onlyparse nenv.ninterp_only_parse ac in
+ let onlyparse = is_not_printable onlyparse (not reversible) ac in
let notation = {
notobj_local = local;
notobj_scope = scope;
notobj_interp = (List.map_filter map i_vars, ac);
(** Order is important here! *)
notobj_onlyparse = onlyparse;
+ notobj_onlyprint = onlyprint;
+ notobj_compat = compat;
notobj_notation = df';
} in
(* Ready to change the global state *)
@@ -1291,14 +1206,17 @@ let add_notation_in_scope local df c mods scope =
Lib.add_anonymous_leaf (inNotation notation);
df'
-let add_notation_interpretation_core local df ?(impls=empty_internalization_env) c scope onlyparse =
+let add_notation_interpretation_core local df ?(impls=empty_internalization_env) c scope onlyparse onlyprint compat =
let dfs = split_notation_string df in
let (recvars,mainvars,symbs) = analyze_notation_tokens dfs in
(* Recover types of variables and pa/pp rules; redeclare them if needed *)
- let i_typs = if not (is_numeral symbs) then begin
- let i_typs,sy_rules = recover_notation_syntax (make_notation_key symbs) in
- Lib.add_anonymous_leaf (inSyntaxExtension (local,sy_rules)); i_typs
- end else [] in
+ let i_typs, onlyprint = if not (is_numeral symbs) then begin
+ let i_typs,sy_rules,onlyprint' = recover_notation_syntax (make_notation_key symbs) in
+ let () = Lib.add_anonymous_leaf (inSyntaxExtension (local,sy_rules)) in
+ (** If the only printing flag has been explicitly requested, put it back *)
+ let onlyprint = onlyprint || onlyprint' in
+ i_typs, onlyprint
+ end else [], false in
(* Declare interpretation *)
let path = (Lib.library_dp(),Lib.current_dirpath true) in
let df' = (make_notation_key symbs,(path,df)) in
@@ -1306,18 +1224,19 @@ let add_notation_interpretation_core local df ?(impls=empty_internalization_env)
let nenv = {
ninterp_var_type = to_map i_vars;
ninterp_rec_vars = to_map recvars;
- ninterp_only_parse = false;
} in
- let (acvars, ac) = interp_notation_constr ~impls nenv c in
+ let (acvars, ac, reversible) = interp_notation_constr ~impls nenv c in
let interp = make_interpretation_vars recvars acvars in
let map (x, _) = try Some (x, Id.Map.find x interp) with Not_found -> None in
- let onlyparse = is_not_printable onlyparse nenv.ninterp_only_parse ac in
+ let onlyparse = is_not_printable onlyparse (not reversible) ac in
let notation = {
notobj_local = local;
notobj_scope = scope;
notobj_interp = (List.map_filter map i_vars, ac);
(** Order is important here! *)
notobj_onlyparse = onlyparse;
+ notobj_onlyprint = onlyprint;
+ notobj_compat = compat;
notobj_notation = df';
} in
Lib.add_anonymous_leaf (inNotation notation);
@@ -1326,20 +1245,20 @@ let add_notation_interpretation_core local df ?(impls=empty_internalization_env)
(* Notations without interpretation (Reserved Notation) *)
let add_syntax_extension local ((loc,df),mods) =
- let msgs, sy_data, extra = compute_pure_syntax_data df mods in
- let sy_rules = make_syntax_rules sy_data extra in
+ let msgs, sy_data, extra, onlyprint = compute_pure_syntax_data df mods in
+ let sy_rules = make_syntax_rules sy_data extra onlyprint None in
Flags.if_verbose (List.iter (fun (f,x) -> f x)) msgs;
Lib.add_anonymous_leaf (inSyntaxExtension(local,sy_rules))
(* Notations with only interpretation *)
let add_notation_interpretation ((loc,df),c,sc) =
- let df' = add_notation_interpretation_core false df c sc false in
+ let df' = add_notation_interpretation_core false df c sc false false None in
Dumpglob.dump_notation (loc,df') sc true
let set_notation_for_interpretation impls ((_,df),c,sc) =
(try ignore
- (silently (add_notation_interpretation_core false df ~impls c sc) false);
+ (silently (fun () -> add_notation_interpretation_core false df ~impls c sc false false None) ());
with NoSyntaxRule ->
error "Parsing rule for this notation has to be previously declared.");
Option.iter (fun sc -> Notation.open_close_scope (false,true,sc)) sc
@@ -1351,7 +1270,9 @@ let add_notation local c ((loc,df),modifiers) sc =
if no_syntax_modifiers modifiers then
(* No syntax data: try to rely on a previously declared rule *)
let onlyparse = is_only_parsing modifiers in
- try add_notation_interpretation_core local df c sc onlyparse
+ let onlyprint = is_only_printing modifiers in
+ let compat = get_compat_version modifiers in
+ try add_notation_interpretation_core local df c sc onlyparse onlyprint compat
with NoSyntaxRule ->
(* Try to determine a default syntax rule *)
add_notation_in_scope local df c modifiers sc
@@ -1444,11 +1365,10 @@ let add_syntactic_definition ident (vars,c) local onlyparse =
let nenv = {
ninterp_var_type = i_vars;
ninterp_rec_vars = Id.Map.empty;
- ninterp_only_parse = false;
} in
- let nvars, pat = interp_notation_constr nenv c in
- let () = nonprintable := nenv.ninterp_only_parse in
- let map id = let (sc, _) = Id.Map.find id nvars in (id, sc) in
+ let nvars, pat, reversible = interp_notation_constr nenv c in
+ let () = nonprintable := not reversible in
+ let map id = let (_,sc,_) = Id.Map.find id nvars in (id, sc) in
List.map map vars, pat
in
let onlyparse = match onlyparse with
diff --git a/toplevel/metasyntax.mli b/toplevel/metasyntax.mli
index ffebd07d..085cc87c 100644
--- a/toplevel/metasyntax.mli
+++ b/toplevel/metasyntax.mli
@@ -15,17 +15,6 @@ open Notation_term
val add_token_obj : string -> unit
-(** Adding a tactic notation in the environment *)
-
-val add_tactic_notation :
- locality_flag * int * grammar_tactic_prod_item_expr list * raw_tactic_expr ->
- unit
-
-type atomic_entry = string * Genarg.glob_generic_argument list option
-
-val add_ml_tactic_notation : ml_tactic_name ->
- Egramml.grammar_prod_item list list -> atomic_entry list -> unit
-
(** Adding a (constr) notation in the environment*)
val add_infix : locality_flag -> (lstring * syntax_modifier list) ->
diff --git a/toplevel/mltop.ml b/toplevel/mltop.ml
index 0b6d93d6..b6690fe4 100644
--- a/toplevel/mltop.ml
+++ b/toplevel/mltop.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Errors
+open CErrors
open Util
open Pp
open Flags
@@ -45,12 +45,12 @@ open System
to build a dummy dynlink.cmxa, cf. dev/dynlink.ml. *)
(* This path is where we look for .cmo *)
-let coq_mlpath_copy = ref ["."]
+let coq_mlpath_copy = ref [Sys.getcwd ()]
let keep_copy_mlpath path =
let cpath = CUnix.canonical_path_name path in
- let filter path' = not (String.equal cpath (CUnix.canonical_path_name path'))
+ let filter path' = not (String.equal cpath path')
in
- coq_mlpath_copy := path :: List.filter filter !coq_mlpath_copy
+ coq_mlpath_copy := cpath :: List.filter filter !coq_mlpath_copy
(* If there is a toplevel under Coq *)
type toplevel = {
@@ -118,8 +118,8 @@ let ml_load s =
| WithTop t ->
(try t.load_obj s; s
with
- | e when Errors.noncritical e ->
- let e = Errors.push e in
+ | e when CErrors.noncritical e ->
+ let e = CErrors.push e in
match fst e with
| (UserError _ | Failure _ | Not_found as u) -> Exninfo.iraise (u, snd e)
| exc ->
@@ -146,7 +146,12 @@ let dir_ml_load s =
let dir_ml_use s =
match !load with
| WithTop t -> t.use_file s
- | _ -> msg_warning (str "Cannot access the ML compiler")
+ | _ ->
+ let moreinfo =
+ if Dynlink.is_native then " Loading ML code works only in bytecode."
+ else ""
+ in
+ errorlabstrm "Mltop.dir_ml_use" (str "Could not load ML code." ++ str moreinfo)
(* Adds a path to the ML paths *)
let add_ml_dir s =
@@ -155,19 +160,31 @@ let add_ml_dir s =
| WithoutTop when has_dynlink -> keep_copy_mlpath s
| _ -> ()
-(* For Rec Add ML Path *)
+(* For Rec Add ML Path (-R) *)
let add_rec_ml_dir unix_path =
List.iter (fun (lp,_) -> add_ml_dir lp) (all_subdirs ~unix_path)
(* Adding files to Coq and ML loadpath *)
+let warn_cannot_use_directory =
+ CWarnings.create ~name:"cannot-use-directory" ~category:"filesystem"
+ (fun d ->
+ str "Directory " ++ str d ++
+ strbrk " cannot be used as a Coq identifier (skipped)")
+
let convert_string d =
try Names.Id.of_string d
with UserError _ ->
- msg_warning (str "Directory " ++ str d ++ str " cannot be used as a Coq identifier (skipped)");
+ warn_cannot_use_directory d;
raise Exit
-let add_rec_path ~unix_path ~coq_root ~implicit =
+let warn_cannot_open_path =
+ CWarnings.create ~name:"cannot-open-path" ~category:"filesystem"
+ (fun unix_path -> str "Cannot open " ++ str unix_path)
+
+type add_ml = AddNoML | AddTopML | AddRecML
+
+let add_rec_path add_ml ~unix_path ~coq_root ~implicit =
if exists_dir unix_path then
let dirs = all_subdirs ~unix_path in
let prefix = Names.DirPath.repr coq_root in
@@ -178,13 +195,16 @@ let add_rec_path ~unix_path ~coq_root ~implicit =
with Exit -> None
in
let dirs = List.map_filter convert_dirs dirs in
- let () = add_ml_dir unix_path in
+ let () = match add_ml with
+ | AddNoML -> ()
+ | AddTopML -> add_ml_dir unix_path
+ | AddRecML -> List.iter (fun (lp,_) -> add_ml_dir lp) dirs in
let add (path, dir) =
Loadpath.add_load_path path ~implicit dir in
let () = List.iter add dirs in
Loadpath.add_load_path unix_path ~implicit coq_root
else
- msg_warning (str "Cannot open " ++ str unix_path)
+ warn_cannot_open_path unix_path
(* convertit un nom quelconque en nom de fichier ou de module *)
let mod_of_name name =
@@ -324,10 +344,10 @@ let if_verbose_load verb f name ?path fname =
let info = str "[Loading ML file " ++ str fname ++ str " ..." in
try
let path = f name ?path fname in
- msg_info (info ++ str " done]");
+ Feedback.msg_info (info ++ str " done]");
path
with reraise ->
- msg_info (info ++ str " failed]");
+ Feedback.msg_info (info ++ str " failed]");
raise reraise
(** Load a module for the first time (i.e. dynlink it)
@@ -391,7 +411,7 @@ let inMLModule : ml_module_object -> obj =
let declare_ml_modules local l =
let l = List.map mod_of_name l in
- Lib.add_anonymous_leaf (inMLModule {mlocal=local; mnames=l})
+ Lib.add_anonymous_leaf ~cache_first:false (inMLModule {mlocal=local; mnames=l})
let print_ml_path () =
let l = !coq_mlpath_copy in
diff --git a/toplevel/mltop.mli b/toplevel/mltop.mli
index 5d054682..6633cb93 100644
--- a/toplevel/mltop.mli
+++ b/toplevel/mltop.mli
@@ -46,8 +46,10 @@ val dir_ml_use : string -> unit
val add_ml_dir : string -> unit
val add_rec_ml_dir : string -> unit
+type add_ml = AddNoML | AddTopML | AddRecML
+
(** Adds a path to the Coq and ML paths *)
-val add_rec_path : unix_path:string -> coq_root:Names.DirPath.t -> implicit:bool -> unit
+val add_rec_path : add_ml -> unix_path:string -> coq_root:Names.DirPath.t -> implicit:bool -> unit
(** List of modules linked to the toplevel *)
val add_known_module : string -> unit
diff --git a/toplevel/obligations.ml b/toplevel/obligations.ml
index 314789ce..29d74573 100644
--- a/toplevel/obligations.ml
+++ b/toplevel/obligations.ml
@@ -13,21 +13,16 @@ open Declare
*)
open Term
-open Context
open Vars
open Names
open Evd
open Pp
-open Errors
+open CErrors
open Util
let declare_fix_ref = ref (fun ?opaque _ _ _ _ _ _ -> assert false)
let declare_definition_ref = ref (fun _ _ _ _ _ -> assert false)
-let trace s =
- if !Flags.debug then msg_debug s
- else ()
-
let succfix (depth, fixrels) =
(succ depth, List.map succ fixrels)
@@ -44,21 +39,19 @@ let check_evars env evm =
type oblinfo =
{ ev_name: int * Id.t;
- ev_hyps: named_context;
- ev_status: Evar_kinds.obligation_definition_status;
+ ev_hyps: Context.Named.t;
+ ev_status: bool * Evar_kinds.obligation_definition_status;
ev_chop: int option;
ev_src: Evar_kinds.t Loc.located;
ev_typ: types;
ev_tac: unit Proofview.tactic option;
ev_deps: Int.Set.t }
-(* spiwack: Store field for internalizing ev_tac in evar_infos' evar_extra. *)
-let evar_tactic = Store.field ()
-
(** Substitute evar references in t using De Bruijn indices,
where n binders were passed through. *)
-let subst_evar_constr evs n idf t =
+let subst_evar_constr evs n idf t =
+ let open Context.Named.Declaration in
let seen = ref Int.Set.empty in
let transparent = ref Id.Set.empty in
let evar_info id = List.assoc_f Evar.equal id evs in
@@ -82,9 +75,9 @@ let subst_evar_constr evs n idf t =
let args =
let rec aux hyps args acc =
match hyps, args with
- ((_, None, _) :: tlh), (c :: tla) ->
+ (LocalAssum _ :: tlh), (c :: tla) ->
aux tlh tla ((substrec (depth, fixrels) c) :: acc)
- | ((_, Some _, _) :: tlh), (_ :: tla) ->
+ | (LocalDef _ :: tlh), (_ :: tla) ->
aux tlh tla acc
| [], [] -> acc
| _, _ -> acc (*failwith "subst_evars: invalid argument"*)
@@ -120,22 +113,23 @@ let subst_vars acc n t =
Changes evars and hypothesis references to variable references.
*)
let etype_of_evar evs hyps concl =
+ let open Context.Named.Declaration in
let rec aux acc n = function
- (id, copt, t) :: tl ->
- let t', s, trans = subst_evar_constr evs n mkVar t in
+ decl :: tl ->
+ let t', s, trans = subst_evar_constr evs n mkVar (get_type decl) in
let t'' = subst_vars acc 0 t' in
- let rest, s', trans' = aux (id :: acc) (succ n) tl in
+ let rest, s', trans' = aux (get_id decl :: acc) (succ n) tl in
let s' = Int.Set.union s s' in
let trans' = Id.Set.union trans trans' in
- (match copt with
- Some c ->
+ (match decl with
+ | LocalDef (id,c,_) ->
let c', s'', trans'' = subst_evar_constr evs n mkVar c in
let c' = subst_vars acc 0 c' in
- mkNamedProd_or_LetIn (id, Some c', t'') rest,
+ mkNamedProd_or_LetIn (LocalDef (id, c', t'')) rest,
Int.Set.union s'' s',
Id.Set.union trans'' trans'
- | None ->
- mkNamedProd_or_LetIn (id, None, t'') rest, s', trans')
+ | LocalAssum (id,_) ->
+ mkNamedProd_or_LetIn (LocalAssum (id, t'')) rest, s', trans')
| [] ->
let t', s, trans = subst_evar_constr evs n mkVar concl in
subst_vars acc 0 t', s, trans
@@ -194,7 +188,7 @@ open Environ
let eterm_obligations env name evm fs ?status t ty =
(* 'Serialize' the evars *)
let nc = Environ.named_context env in
- let nc_len = Context.named_context_length nc in
+ let nc_len = Context.Named.length nc in
let evm = Evarutil.nf_evar_map_undefined evm in
let evl = Evarutil.non_instantiated evm in
let evl = Evar.Map.bindings evl in
@@ -221,25 +215,21 @@ let eterm_obligations env name evm fs ?status t ty =
| None -> evtyp, hyps, 0
in
let loc, k = evar_source id evm in
- let status = match k with Evar_kinds.QuestionMark o -> Some o | _ -> status in
- let status, chop = match status with
- | Some (Evar_kinds.Define true as stat) ->
- if not (Int.equal chop fs) then Evar_kinds.Define false, None
- else stat, Some chop
- | Some s -> s, None
- | None -> Evar_kinds.Define true, None
- in
- let tac = match Store.get ev.evar_extra evar_tactic with
- | Some t ->
- if Dyn.has_tag t "tactic" then
- Some (Tacinterp.interp
- (Tacinterp.globTacticIn (Tacinterp.tactic_out t)))
- else None
- | None -> None
+ let status = match k with
+ | Evar_kinds.QuestionMark o -> o
+ | _ -> match status with
+ | Some o -> o
+ | None -> Evar_kinds.Define (not (Program.get_proofs_transparency ()))
+ in
+ let force_status, status, chop = match status with
+ | Evar_kinds.Define true as stat ->
+ if not (Int.equal chop fs) then true, Evar_kinds.Define false, None
+ else false, stat, Some chop
+ | s -> false, s, None
in
let info = { ev_name = (n, nstr);
- ev_hyps = hyps; ev_status = status; ev_chop = chop;
- ev_src = loc, k; ev_typ = evtyp ; ev_deps = deps; ev_tac = tac }
+ ev_hyps = hyps; ev_status = force_status, status; ev_chop = chop;
+ ev_src = loc, k; ev_typ = evtyp ; ev_deps = deps; ev_tac = None }
in (id, info) :: l)
evn []
in
@@ -249,14 +239,14 @@ let eterm_obligations env name evm fs ?status t ty =
let ty, _, _ = subst_evar_constr evts 0 mkVar ty in
let evars =
List.map (fun (ev, info) ->
- let { ev_name = (_, name); ev_status = status;
+ let { ev_name = (_, name); ev_status = force_status, status;
ev_src = src; ev_typ = typ; ev_deps = deps; ev_tac = tac } = info
in
- let status = match status with
+ let force_status, status = match status with
| Evar_kinds.Define true when Id.Set.mem name transparent ->
- Evar_kinds.Define false
- | _ -> status
- in name, typ, src, status, deps, tac) evts
+ true, Evar_kinds.Define false
+ | _ -> force_status, status
+ in name, typ, src, (force_status, status), deps, tac) evts
in
let evnames = List.map (fun (ev, info) -> ev, snd info.ev_name) evts in
let evmap f c = pi1 (subst_evar_constr evts 0 f c) in
@@ -268,21 +258,22 @@ let safe_init_constant md name () =
Coqlib.gen_constant "Obligations" md name
let hide_obligation = safe_init_constant tactics_module "obligation"
-let pperror cmd = Errors.errorlabstrm "Program" cmd
+let pperror cmd = CErrors.errorlabstrm "Program" cmd
let error s = pperror (str s)
let reduce c =
- Reductionops.clos_norm_flags Closure.betaiota (Global.env ()) Evd.empty c
+ Reductionops.clos_norm_flags CClosure.betaiota (Global.env ()) Evd.empty c
exception NoObligations of Id.t option
let explain_no_obligations = function
- Some ident -> str "No obligations for program " ++ str (Id.to_string ident)
+ Some ident -> str "No obligations for program " ++ Id.print ident
| None -> str "No obligations remaining"
type obligation_info =
(Names.Id.t * Term.types * Evar_kinds.t Loc.located *
- Evar_kinds.obligation_definition_status * Int.Set.t * unit Proofview.tactic option) array
+ (bool * Evar_kinds.obligation_definition_status)
+ * Int.Set.t * unit Proofview.tactic option) array
type 'a obligation_body =
| DefinedObl of 'a
@@ -293,7 +284,7 @@ type obligation =
obl_type : types;
obl_location : Evar_kinds.t Loc.located;
obl_body : constant obligation_body option;
- obl_status : Evar_kinds.obligation_definition_status;
+ obl_status : bool * Evar_kinds.obligation_definition_status;
obl_deps : Int.Set.t;
obl_tac : unit Proofview.tactic option;
}
@@ -311,6 +302,7 @@ type program_info_aux = {
prg_body: constr;
prg_type: constr;
prg_ctx: Evd.evar_universe_context;
+ prg_pl: Id.t Loc.located list option;
prg_obligations: obligations;
prg_deps : Id.t list;
prg_fixkind : fixpoint_kind option ;
@@ -323,34 +315,16 @@ type program_info_aux = {
prg_sign: named_context_val;
}
-type program_info = program_info_aux Ephemeron.key
+type program_info = program_info_aux CEphemeron.key
let get_info x =
- try Ephemeron.get x
- with Ephemeron.InvalidKey ->
- Errors.anomaly Pp.(str "Program obligation can't be accessed by a worker")
+ try CEphemeron.get x
+ with CEphemeron.InvalidKey ->
+ CErrors.anomaly Pp.(str "Program obligation can't be accessed by a worker")
let assumption_message = Declare.assumption_message
-let (set_default_tactic, get_default_tactic, print_default_tactic) =
- Tactic_option.declare_tactic_option "Program tactic"
-
-(* true = All transparent, false = Opaque if possible *)
-let proofs_transparency = ref true
-
-let set_proofs_transparency = (:=) proofs_transparency
-let get_proofs_transparency () = !proofs_transparency
-
-open Goptions
-
-let _ =
- declare_bool_option
- { optsync = true;
- optdepr = false;
- optname = "transparency of Program obligations";
- optkey = ["Transparent";"Obligations"];
- optread = get_proofs_transparency;
- optwrite = set_proofs_transparency; }
+let default_tactic = ref (Proofview.tclUNIT ())
(* true = hide obligations *)
let hide_obligations = ref false
@@ -358,6 +332,7 @@ let hide_obligations = ref false
let set_hide_obligations = (:=) hide_obligations
let get_hide_obligations () = !hide_obligations
+open Goptions
let _ =
declare_bool_option
{ optsync = true;
@@ -367,7 +342,7 @@ let _ =
optread = get_hide_obligations;
optwrite = set_hide_obligations; }
-let shrink_obligations = ref false
+let shrink_obligations = ref true
let set_shrink_obligations = (:=) shrink_obligations
let get_shrink_obligations () = !shrink_obligations
@@ -375,7 +350,7 @@ let get_shrink_obligations () = !shrink_obligations
let _ =
declare_bool_option
{ optsync = true;
- optdepr = false;
+ optdepr = true;
optname = "Shrinking of Program obligations";
optkey = ["Shrink";"Obligations"];
optread = get_shrink_obligations;
@@ -383,36 +358,35 @@ let _ =
let evar_of_obligation o = make_evar (Global.named_context_val ()) o.obl_type
-let get_body obl =
+let get_body obl =
match obl.obl_body with
- | None -> assert false
+ | None -> None
| Some (DefinedObl c) ->
let ctx = Environ.constant_context (Global.env ()) c in
let pc = (c, Univ.UContext.instance ctx) in
- DefinedObl pc
- | Some (TermObl c) ->
- TermObl c
+ Some (DefinedObl pc)
+ | Some (TermObl c) ->
+ Some (TermObl c)
let get_obligation_body expand obl =
- let c = get_body obl in
- let c' =
- if expand && obl.obl_status == Evar_kinds.Expand then
- (match c with
- | DefinedObl pc -> constant_value_in (Global.env ()) pc
- | TermObl c -> c)
- else (match c with
- | DefinedObl pc -> mkConstU pc
- | TermObl c -> c)
- in c'
+ match get_body obl with
+ | None -> None
+ | Some c ->
+ if expand && snd obl.obl_status == Evar_kinds.Expand then
+ match c with
+ | DefinedObl pc -> Some (constant_value_in (Global.env ()) pc)
+ | TermObl c -> Some c
+ else (match c with
+ | DefinedObl pc -> Some (mkConstU pc)
+ | TermObl c -> Some c)
let obl_substitution expand obls deps =
Int.Set.fold
(fun x acc ->
let xobl = obls.(x) in
- let oblb =
- try get_obligation_body expand xobl
- with e when Errors.noncritical e -> assert false
- in (xobl.obl_name, (xobl.obl_type, oblb)) :: acc)
+ match get_obligation_body expand xobl with
+ | None -> acc
+ | Some oblb -> (xobl.obl_name, (xobl.obl_type, oblb)) :: acc)
deps []
let subst_deps expand obls deps t =
@@ -458,9 +432,9 @@ let subst_deps_obl obls obl =
let t' = subst_deps true obls obl.obl_deps obl.obl_type in
{ obl with obl_type = t' }
-module ProgMap = Map.Make(Id)
+module ProgMap = Id.Map
-let map_replace k v m = ProgMap.add k (Ephemeron.create v) (ProgMap.remove k m)
+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 []
@@ -510,15 +484,21 @@ let declare_definition prg =
(Evd.evar_universe_context_subst prg.prg_ctx) in
let opaque = prg.prg_opaque in
let fix_exn = Stm.get_fix_exn () in
+ let pl, ctx =
+ Evd.universe_context ?names:prg.prg_pl (Evd.from_ctx prg.prg_ctx) in
let ce =
definition_entry ~fix_exn
~opaque ~types:(nf typ) ~poly:(pi2 prg.prg_kind)
~univs:(Evd.evar_context_universe_context prg.prg_ctx) (nf body)
in
- progmap_remove prg;
+ let () = progmap_remove prg in
+ let cst =
!declare_definition_ref prg.prg_name
- prg.prg_kind ce prg.prg_implicits
- (Lemmas.mk_hook (fun l r -> Lemmas.call_hook fix_exn prg.prg_hook l r prg.prg_ctx; r))
+ prg.prg_kind ce prg.prg_implicits
+ (Lemmas.mk_hook (fun l r -> Lemmas.call_hook fix_exn prg.prg_hook l r prg.prg_ctx; r))
+ in
+ Universes.register_universe_binders cst pl;
+ cst
open Pp
@@ -539,7 +519,7 @@ let compute_possible_guardness_evidences (n,_) fixbody fixtype =
but doing it properly involves delta-reduction, and it finally
doesn't seem to worth the effort (except for huge mutual
fixpoints ?) *)
- let m = nb_prod fixtype in
+ let m = Termops.nb_prod fixtype in
let ctx = fst (decompose_prod_n_assum m fixtype) in
List.map_i (fun i _ -> i) 0 ctx
@@ -571,7 +551,8 @@ let declare_mutual_definition l =
List.map3 compute_possible_guardness_evidences
wfl fixdefs fixtypes in
let indexes =
- Pretyping.search_guard Loc.ghost (Global.env())
+ Pretyping.search_guard
+ Loc.ghost (Global.env())
possible_indexes fixdecls in
Some indexes,
List.map_i (fun i _ ->
@@ -594,15 +575,49 @@ let declare_mutual_definition l =
Lemmas.call_hook fix_exn first.prg_hook local gr first.prg_ctx;
List.iter progmap_remove l; kn
-let shrink_body c =
- let ctx, b = decompose_lam c in
- let b', n, args =
- List.fold_left (fun (b, i, args) (n,t) ->
- if noccurn 1 b then
- subst1 mkProp b, succ i, args
- else mkLambda (n,t,b), succ i, mkRel i :: args)
- (b, 1, []) ctx
- in List.map (fun (c,t) -> (c,None,t)) ctx, b', Array.of_list args
+let decompose_lam_prod c ty =
+ let open Context.Rel.Declaration in
+ let rec aux ctx c ty =
+ match kind_of_term c, kind_of_term ty with
+ | LetIn (x, b, t, c), LetIn (x', b', t', ty)
+ when eq_constr b b' && eq_constr t t' ->
+ let ctx' = Context.Rel.add (LocalDef (x,b',t')) ctx in
+ aux ctx' c ty
+ | _, LetIn (x', b', t', ty) ->
+ let ctx' = Context.Rel.add (LocalDef (x',b',t')) ctx in
+ aux ctx' (lift 1 c) ty
+ | LetIn (x, b, t, c), _ ->
+ let ctx' = Context.Rel.add (LocalDef (x,b,t)) ctx in
+ aux ctx' c (lift 1 ty)
+ | Lambda (x, b, t), Prod (x', b', t')
+ (* By invariant, must be convertible *) ->
+ let ctx' = Context.Rel.add (LocalAssum (x,b')) ctx in
+ aux ctx' t t'
+ | Cast (c, _, _), _ -> aux ctx c ty
+ | _, _ -> ctx, c, ty
+ in aux Context.Rel.empty c ty
+
+let shrink_body c ty =
+ let open Context.Rel.Declaration in
+ let ctx, b, ty =
+ match ty with
+ | None ->
+ let ctx, b = decompose_lam_assum c in
+ ctx, b, None
+ | Some ty ->
+ let ctx, b, ty = decompose_lam_prod c ty in
+ ctx, b, Some ty
+ in
+ let b', ty', n, args =
+ List.fold_left (fun (b, ty, i, args) decl ->
+ if noccurn 1 b && Option.cata (noccurn 1) true ty then
+ subst1 mkProp b, Option.map (subst1 mkProp) ty, succ i, args
+ else
+ let args = if is_local_assum decl then mkRel i :: args else args in
+ mkLambda_or_LetIn decl b, Option.map (mkProd_or_LetIn decl) ty,
+ succ i, args)
+ (b, ty, 1, []) ctx
+ in ctx, b', ty', Array.of_list args
let unfold_entry cst = Hints.HintsUnfoldEntry [EvalConstRef cst]
@@ -613,46 +628,47 @@ let declare_obligation prg obl body ty uctx =
let body = prg.prg_reduce body in
let ty = Option.map prg.prg_reduce ty in
match obl.obl_status with
- | Evar_kinds.Expand -> false, { obl with obl_body = Some (TermObl body) }
- | Evar_kinds.Define opaque ->
- let opaque = if get_proofs_transparency () then false else opaque in
+ | _, Evar_kinds.Expand -> false, { obl with obl_body = Some (TermObl body) }
+ | force, Evar_kinds.Define opaque ->
+ let opaque = not force && opaque in
let poly = pi2 prg.prg_kind in
- let ctx, body, args =
+ let ctx, body, ty, args =
if get_shrink_obligations () && not poly then
- shrink_body body else [], body, [||]
+ shrink_body body ty else [], body, ty, [||]
in
let body = ((body,Univ.ContextSet.empty),Safe_typing.empty_private_constants) in
- let ce =
+ let ce =
{ const_entry_body = Future.from_val ~fix_exn:(fun x -> x) body;
const_entry_secctx = None;
- const_entry_type = if List.is_empty ctx then ty else None;
+ const_entry_type = ty;
const_entry_polymorphic = poly;
const_entry_universes = uctx;
const_entry_opaque = opaque;
const_entry_inline_code = false;
const_entry_feedback = None;
} in
- (** ppedrot: seems legit to have obligations as local *)
+ (** ppedrot: seems legit to have obligations as local *)
let constant = Declare.declare_constant obl.obl_name ~local:true
(DefinitionEntry ce,IsProof Property)
in
if not opaque then add_hint false prg constant;
definition_message obl.obl_name;
true, { obl with obl_body =
- if poly then
+ if poly then
Some (DefinedObl constant)
else
Some (TermObl (it_mkLambda_or_LetIn (mkApp (mkConst constant, args)) ctx)) }
-let init_prog_info ?(opaque = false) sign n b t ctx deps fixkind notations obls impls kind reduce hook =
- let obls', b =
+let init_prog_info ?(opaque = false) sign n pl b t ctx deps fixkind
+ notations obls impls kind reduce hook =
+ let obls', b =
match b with
| None ->
assert(Int.equal (Array.length obls) 0);
let n = Nameops.add_suffix n "_obligation" in
[| { obl_name = n; obl_body = None;
obl_location = Loc.ghost, Evar_kinds.InternalHole; obl_type = t;
- obl_status = Evar_kinds.Expand; obl_deps = Int.Set.empty;
+ obl_status = false, Evar_kinds.Expand; obl_deps = Int.Set.empty;
obl_tac = None } |],
mkVar n
| Some b ->
@@ -664,7 +680,7 @@ let init_prog_info ?(opaque = false) sign n b t ctx deps fixkind notations obls
obls, b
in
{ prg_name = n ; prg_body = b; prg_type = reduce t;
- prg_ctx = ctx;
+ prg_ctx = ctx; prg_pl = pl;
prg_obligations = (obls', Array.length obls');
prg_deps = deps; prg_fixkind = fixkind ; prg_notations = notations ;
prg_implicits = impls; prg_kind = kind; prg_reduce = reduce;
@@ -674,7 +690,7 @@ let init_prog_info ?(opaque = false) sign n b t ctx deps fixkind notations obls
let map_cardinal m =
let i = ref 0 in
ProgMap.iter (fun _ v ->
- if snd (Ephemeron.get v).prg_obligations > 0 then incr i) m;
+ if snd (CEphemeron.get v).prg_obligations > 0 then incr i) m;
!i
exception Found of program_info
@@ -682,7 +698,7 @@ exception Found of program_info
let map_first m =
try
ProgMap.iter (fun _ v ->
- if snd (Ephemeron.get v).prg_obligations > 0 then
+ if snd (CEphemeron.get v).prg_obligations > 0 then
raise (Found v)) m;
assert(false)
with Found x -> x
@@ -698,11 +714,13 @@ let get_prog name =
match n with
0 -> raise (NoObligations None)
| 1 -> get_info (map_first prg_infos)
- | _ ->
- error ("More than one program with unsolved obligations: "^
- String.concat ", "
- (List.map string_of_id
- (ProgMap.fold (fun k _ s -> k::s) prg_infos []))))
+ | _ ->
+ let progs = Id.Set.elements (ProgMap.domain prg_infos) in
+ let prog = List.hd progs in
+ let progs = prlist_with_sep pr_comma Nameops.pr_id progs in
+ errorlabstrm ""
+ (str "More than one program with unsolved obligations: " ++ progs
+ ++ str "; use the \"of\" clause to specify, as in \"Obligation 1 of " ++ Nameops.pr_id prog ++ str "\""))
let get_any_prog () =
let prg_infos = !from_prg in
@@ -729,11 +747,11 @@ type progress =
let obligations_message rem =
if rem > 0 then
if Int.equal rem 1 then
- Flags.if_verbose msg_info (int rem ++ str " obligation remaining")
+ Flags.if_verbose Feedback.msg_info (int rem ++ str " obligation remaining")
else
- Flags.if_verbose msg_info (int rem ++ str " obligations remaining")
+ Flags.if_verbose Feedback.msg_info (int rem ++ str " obligations remaining")
else
- Flags.if_verbose msg_info (str "No more obligations remaining")
+ Flags.if_verbose Feedback.msg_info (str "No more obligations remaining")
let update_obls prg obls rem =
let prg' = { prg with prg_obligations = (obls, rem) } in
@@ -807,14 +825,64 @@ let solve_by_tac name evi t poly ctx =
Inductiveops.control_only_guard (Global.env ()) (fst body);
(fst body), entry.const_entry_type, Evd.evar_universe_context ctx'
+let obligation_terminator name num guard hook auto pf =
+ let open Proof_global in
+ let term = Lemmas.universe_proof_terminator guard hook in
+ match pf with
+ | Admitted _ -> apply_terminator term pf
+ | Proved (opq, id, proof) ->
+ if not !shrink_obligations then apply_terminator term pf
+ else
+ let (_, (entry, uctx, _)) = Pfedit.cook_this_proof proof in
+ let env = Global.env () in
+ let entry = Safe_typing.inline_private_constants_in_definition_entry env entry in
+ let ty = entry.Entries.const_entry_type in
+ let (body, cstr), eff = Future.force entry.Entries.const_entry_body in
+ assert(Safe_typing.empty_private_constants = eff);
+ let sigma = Evd.from_ctx (fst uctx) in
+ let sigma = Evd.merge_context_set ~sideff:true Evd.univ_rigid sigma cstr in
+ Inductiveops.control_only_guard (Global.env ()) body;
+ (** Declare the obligation ourselves and drop the hook *)
+ let prg = get_info (ProgMap.find name !from_prg) in
+ let ctx = Evd.evar_universe_context sigma in
+ let prg = { prg with prg_ctx = ctx } in
+ let obls, rem = prg.prg_obligations in
+ let obl = obls.(num) in
+ let status =
+ match obl.obl_status, opq with
+ | (_, Evar_kinds.Expand), Vernacexpr.Opaque _ -> err_not_transp ()
+ | (true, _), Vernacexpr.Opaque _ -> err_not_transp ()
+ | (false, _), Vernacexpr.Opaque _ -> Evar_kinds.Define true
+ | (_, Evar_kinds.Define true), Vernacexpr.Transparent -> Evar_kinds.Define false
+ | (_, status), Vernacexpr.Transparent -> status
+ in
+ let obl = { obl with obl_status = false, status } in
+ let uctx = Evd.evar_context_universe_context ctx in
+ let (_, obl) = declare_obligation prg obl body ty uctx in
+ let obls = Array.copy obls in
+ let _ = obls.(num) <- obl in
+ try
+ ignore (update_obls prg obls (pred rem));
+ if pred rem > 0 then
+ begin
+ let deps = dependencies obls num in
+ if not (Int.Set.is_empty deps) then
+ ignore (auto (Some name) None deps)
+ end
+ with e when CErrors.noncritical e ->
+ let e = CErrors.push e in
+ pperror (CErrors.iprint (ExplainErr.process_vernac_interp_error e))
+
let obligation_hook prg obl num auto ctx' _ gr =
let obls, rem = prg.prg_obligations in
let cst = match gr with ConstRef cst -> cst | _ -> assert false in
let transparent = evaluable_constant cst (Global.env ()) in
let () = match obl.obl_status with
- | Evar_kinds.Expand -> if not transparent then err_not_transp ()
- | Evar_kinds.Define op -> if not op && not transparent then err_not_transp ()
- in
+ (true, Evar_kinds.Expand)
+ | (true, Evar_kinds.Define true) ->
+ if not transparent then err_not_transp ()
+ | _ -> ()
+in
let obl = { obl with obl_body = Some (DefinedObl cst) } in
let () = if transparent then add_hint true prg cst in
let obls = Array.copy obls in
@@ -832,9 +900,9 @@ let obligation_hook prg obl num auto ctx' _ gr =
let prg = { prg with prg_ctx = ctx' } in
let () =
try ignore (update_obls prg obls (pred rem))
- with e when Errors.noncritical e ->
- let e = Errors.push e in
- pperror (Errors.iprint (Cerrors.process_vernac_interp_error e))
+ with e when CErrors.noncritical e ->
+ let e = CErrors.push e in
+ pperror (CErrors.iprint (ExplainErr.process_vernac_interp_error e))
in
if pred rem > 0 then begin
let deps = dependencies obls num in
@@ -855,15 +923,16 @@ let rec solve_obligation prg num tac =
++ str (string_of_list ", " (fun x -> string_of_int (succ x)) remaining));
in
let obl = subst_deps_obl obls obl in
- let kind = kind_of_obligation (pi2 prg.prg_kind) obl.obl_status in
+ let kind = kind_of_obligation (pi2 prg.prg_kind) (snd obl.obl_status) in
let evd = Evd.from_ctx prg.prg_ctx in
let evd = Evd.update_sigma_env evd (Global.env ()) in
let auto n tac oblset = auto_solve_obligations n ~oblset tac in
+ let terminator guard hook =
+ Proof_global.make_terminator
+ (obligation_terminator prg.prg_name num guard hook auto) in
let hook ctx = Lemmas.mk_hook (obligation_hook prg obl num auto ctx) in
- let () = Lemmas.start_proof_univs ~sign:prg.prg_sign obl.obl_name kind evd obl.obl_type hook in
- let () = trace (str "Started obligation " ++ int user_num ++ str " proof: " ++
- Printer.pr_constr_env (Global.env ()) Evd.empty obl.obl_type) in
- let _ = Pfedit.by (snd (get_default_tactic ())) in
+ let () = Lemmas.start_proof_univs ~sign:prg.prg_sign obl.obl_name kind evd obl.obl_type ~terminator hook in
+ let _ = Pfedit.by !default_tactic in
Option.iter (fun tac -> Pfedit.set_end_tac tac) tac
and obligation (user_num, name, typ) tac =
@@ -881,7 +950,7 @@ and obligation (user_num, name, typ) tac =
and solve_obligation_by_tac prg obls i tac =
let obl = obls.(i) in
match obl.obl_body with
- | Some _ -> false
+ | Some _ -> None
| None ->
try
if List.is_empty (deps_remaining obls obl.obl_deps) then
@@ -892,32 +961,32 @@ and solve_obligation_by_tac prg obls i tac =
| None ->
match obl.obl_tac with
| Some t -> t
- | None -> snd (get_default_tactic ())
+ | None -> !default_tactic
in
- let evd = Evd.from_ctx !prg.prg_ctx in
+ let evd = Evd.from_ctx prg.prg_ctx in
let evd = Evd.update_sigma_env evd (Global.env ()) in
- let t, ty, ctx =
- solve_by_tac obl.obl_name (evar_of_obligation obl) tac
- (pi2 !prg.prg_kind) (Evd.evar_universe_context evd)
+ let t, ty, ctx =
+ solve_by_tac obl.obl_name (evar_of_obligation obl) tac
+ (pi2 prg.prg_kind) (Evd.evar_universe_context evd)
in
let uctx = Evd.evar_context_universe_context ctx in
- let () = prg := {!prg with prg_ctx = ctx} in
- let def, obl' = declare_obligation !prg obl t ty uctx in
+ let prg = {prg with prg_ctx = ctx} in
+ let def, obl' = declare_obligation prg obl t ty uctx in
obls.(i) <- obl';
- if def && not (pi2 !prg.prg_kind) then (
+ if def && not (pi2 prg.prg_kind) then (
(* Declare the term constraints with the first obligation only *)
let evd = Evd.from_env (Global.env ()) in
let evd = Evd.merge_universe_subst evd (Evd.universe_subst (Evd.from_ctx ctx)) in
let ctx' = Evd.evar_universe_context evd in
- prg := {!prg with prg_ctx = ctx'});
- true
- else false
- with e when Errors.noncritical e ->
- let (e, _) = Errors.push e in
+ Some {prg with prg_ctx = ctx'})
+ else Some prg
+ else None
+ with e when CErrors.noncritical e ->
+ let (e, _) = CErrors.push e in
match e with
| Refiner.FailError (_, s) ->
user_err_loc (fst obl.obl_location, "solve_obligation", Lazy.force s)
- | e -> false
+ | e -> None (* FIXME really ? *)
and solve_prg_obligations prg ?oblset tac =
let obls, rem = prg.prg_obligations in
@@ -929,16 +998,20 @@ and solve_prg_obligations prg ?oblset tac =
| Some s -> set := s;
(fun i -> Int.Set.mem i !set)
in
- let prg = ref prg in
+ let prgref = ref prg in
let _ =
Array.iteri (fun i x ->
- if p i && solve_obligation_by_tac prg obls' i tac then
- let deps = dependencies obls i in
- (set := Int.Set.union !set deps;
- decr rem))
+ if p i then
+ match solve_obligation_by_tac !prgref obls' i tac with
+ | None -> ()
+ | Some prg' ->
+ prgref := prg';
+ let deps = dependencies obls i in
+ (set := Int.Set.union !set deps;
+ decr rem))
obls'
in
- update_obls !prg obls' !rem
+ update_obls !prgref obls' !rem
and solve_obligations n tac =
let prg = get_prog_err n in
@@ -951,15 +1024,15 @@ and try_solve_obligation n prg tac =
let prg = get_prog prg in
let obls, rem = prg.prg_obligations in
let obls' = Array.copy obls in
- let prgref = ref prg in
- if solve_obligation_by_tac prgref obls' n tac then
- ignore(update_obls !prgref obls' (pred rem));
+ match solve_obligation_by_tac prg obls' n tac with
+ | Some prg' -> ignore(update_obls prg' obls' (pred rem))
+ | None -> ()
and try_solve_obligations n tac =
try ignore (solve_obligations n tac) with NoObligations _ -> ()
and auto_solve_obligations n ?oblset tac : progress =
- Flags.if_verbose msg_info (str "Solving obligations automatically...");
+ Flags.if_verbose Feedback.msg_info (str "Solving obligations automatically...");
try solve_prg_obligations (get_prog_err n) ?oblset tac with NoObligations _ -> Dependent
open Pp
@@ -967,14 +1040,15 @@ let show_obligations_of_prg ?(msg=true) prg =
let n = prg.prg_name in
let obls, rem = prg.prg_obligations in
let showed = ref 5 in
- if msg then msg_info (int rem ++ str " obligation(s) remaining: ");
+ if msg then Feedback.msg_info (int rem ++ str " obligation(s) remaining: ");
Array.iteri (fun i x ->
match x.obl_body with
| None ->
if !showed > 0 then (
- decr showed;
- msg_info (str "Obligation" ++ spc() ++ int (succ i) ++ spc () ++
- str "of" ++ spc() ++ str (Id.to_string n) ++ str ":" ++ spc () ++
+ decr showed;
+ let x = subst_deps_obl obls x in
+ Feedback.msg_info (str "Obligation" ++ spc() ++ int (succ i) ++ spc () ++
+ str "of" ++ spc() ++ Id.print n ++ str ":" ++ spc () ++
hov 1 (Printer.pr_constr_env (Global.env ()) Evd.empty x.obl_type ++
str "." ++ fnl ())))
| Some _ -> ())
@@ -991,38 +1065,38 @@ let show_obligations ?(msg=true) n =
let show_term n =
let prg = get_prog_err n in
let n = prg.prg_name in
- (str (Id.to_string n) ++ spc () ++ str":" ++ spc () ++
+ (Id.print n ++ spc () ++ str":" ++ spc () ++
Printer.pr_constr_env (Global.env ()) Evd.empty prg.prg_type ++ spc () ++ str ":=" ++ fnl ()
++ Printer.pr_constr_env (Global.env ()) Evd.empty prg.prg_body)
-let add_definition n ?term t ctx ?(implicits=[]) ?(kind=Global,false,Definition) ?tactic
+let add_definition n ?term t ctx ?pl ?(implicits=[]) ?(kind=Global,false,Definition) ?tactic
?(reduce=reduce) ?(hook=Lemmas.mk_hook (fun _ _ _ -> ())) ?(opaque = false) obls =
let sign = Decls.initialize_named_context_for_proof () in
- let info = str (Id.to_string n) ++ str " has type-checked" in
- let prg = init_prog_info sign ~opaque n term t ctx [] None [] obls implicits kind reduce hook in
+ let info = Id.print n ++ str " has type-checked" in
+ let prg = init_prog_info sign ~opaque n pl term t ctx [] None [] obls implicits kind reduce hook in
let obls,_ = prg.prg_obligations in
if Int.equal (Array.length obls) 0 then (
- Flags.if_verbose msg_info (info ++ str ".");
+ Flags.if_verbose Feedback.msg_info (info ++ str ".");
let cst = declare_definition prg in
Defined cst)
else (
let len = Array.length obls in
- let _ = Flags.if_verbose msg_info (info ++ str ", generating " ++ int len ++ str " obligation(s)") in
- progmap_add n (Ephemeron.create prg);
+ let _ = Flags.if_verbose Feedback.msg_info (info ++ str ", generating " ++ int len ++ str " obligation(s)") in
+ progmap_add n (CEphemeron.create prg);
let res = auto_solve_obligations (Some n) tactic in
match res with
| Remain rem -> Flags.if_verbose (fun () -> show_obligations ~msg:false (Some n)) (); res
| _ -> res)
-let add_mutual_definitions l ctx ?tactic ?(kind=Global,false,Definition) ?(reduce=reduce)
+let add_mutual_definitions l ctx ?pl ?tactic ?(kind=Global,false,Definition) ?(reduce=reduce)
?(hook=Lemmas.mk_hook (fun _ _ _ -> ())) ?(opaque = false) notations fixkind =
let sign = Decls.initialize_named_context_for_proof () in
let deps = List.map (fun (n, b, t, imps, obls) -> n) l in
List.iter
(fun (n, b, t, imps, obls) ->
- let prg = init_prog_info sign ~opaque n (Some b) t ctx deps (Some fixkind)
+ let prg = init_prog_info sign ~opaque n pl (Some b) t ctx deps (Some fixkind)
notations obls imps kind reduce hook
- in progmap_add n (Ephemeron.create prg)) l;
+ in progmap_add n (CEphemeron.create prg)) l;
let _defined =
List.fold_left (fun finished x ->
if finished then finished
diff --git a/toplevel/obligations.mli b/toplevel/obligations.mli
index b2320a57..69d20696 100644
--- a/toplevel/obligations.mli
+++ b/toplevel/obligations.mli
@@ -33,7 +33,8 @@ val sort_dependencies : (Evar.t * evar_info * Evar.Set.t) list -> (Evar.t * evar
evars contexts, object and type *)
val eterm_obligations : env -> Id.t -> evar_map -> int ->
?status:Evar_kinds.obligation_definition_status -> constr -> types ->
- (Id.t * types * Evar_kinds.t Loc.located * Evar_kinds.obligation_definition_status * Int.Set.t *
+ (Id.t * types * Evar_kinds.t Loc.located *
+ (bool * Evar_kinds.obligation_definition_status) * Int.Set.t *
unit Proofview.tactic option) array
(* Existential key, obl. name, type as product,
location of the original evar, associated tactic,
@@ -46,7 +47,7 @@ val eterm_obligations : env -> Id.t -> evar_map -> int ->
type obligation_info =
(Id.t * Term.types * Evar_kinds.t Loc.located *
- Evar_kinds.obligation_definition_status * Int.Set.t * unit Proofview.tactic option) array
+ (bool * Evar_kinds.obligation_definition_status) * Int.Set.t * unit Proofview.tactic option) array
(* ident, type, location, (opaque or transparent, expand or define),
dependencies, tactic to solve it *)
@@ -54,16 +55,12 @@ type progress = (* Resolution status of a program *)
| Remain of int (* n obligations remaining *)
| Dependent (* Dependent on other definitions *)
| Defined of global_reference (* Defined as id *)
-
-val set_default_tactic : bool -> Tacexpr.glob_tactic_expr -> unit
-val get_default_tactic : unit -> locality_flag * unit Proofview.tactic
-val print_default_tactic : unit -> Pp.std_ppcmds
-val set_proofs_transparency : bool -> unit (* true = All transparent, false = Opaque if possible *)
-val get_proofs_transparency : unit -> bool
+val default_tactic : unit Proofview.tactic ref
val add_definition : Names.Id.t -> ?term:Term.constr -> Term.types ->
Evd.evar_universe_context ->
+ ?pl:(Id.t Loc.located list) -> (* Universe binders *)
?implicits:(Constrexpr.explicitation * (bool * bool * bool)) list ->
?kind:Decl_kinds.definition_kind ->
?tactic:unit Proofview.tactic ->
@@ -81,6 +78,7 @@ val add_mutual_definitions :
(Names.Id.t * Term.constr * Term.types *
(Constrexpr.explicitation * (bool * bool * bool)) list * obligation_info) list ->
Evd.evar_universe_context ->
+ ?pl:(Id.t Loc.located list) -> (* Universe binders *)
?tactic:unit Proofview.tactic ->
?kind:Decl_kinds.definition_kind ->
?reduce:(Term.constr -> Term.constr) ->
diff --git a/toplevel/record.ml b/toplevel/record.ml
index 04da628c..ef09f6fa 100644
--- a/toplevel/record.ml
+++ b/toplevel/record.ml
@@ -7,13 +7,12 @@
(************************************************************************)
open Pp
-open Errors
+open CErrors
open Util
open Names
open Globnames
open Nameops
open Term
-open Context
open Vars
open Environ
open Declarations
@@ -25,6 +24,8 @@ open Type_errors
open Constrexpr
open Constrexpr_ops
open Goptions
+open Sigma.Notations
+open Context.Rel.Declaration
(********** definition d'un record (structure) **************)
@@ -69,16 +70,19 @@ let interp_fields_evars env evars impls_env nots l =
| Anonymous -> impls
| Name id -> Id.Map.add id (compute_internalization_data env Constrintern.Method t' impl) impls
in
- let d = (i,b',t') in
+ let d = match b' with
+ | None -> LocalAssum (i,t')
+ | Some b' -> LocalDef (i,b',t')
+ in
List.iter (Metasyntax.set_notation_for_interpretation impls) no;
(push_rel d env, impl :: uimpls, d::params, impls))
(env, [], [], impls_env) nots l
let compute_constructor_level evars env l =
- List.fold_right (fun (n,b,t as d) (env, univ) ->
+ List.fold_right (fun d (env, univ) ->
let univ =
- if b = None then
- let s = Retyping.get_sort_of env evars t in
+ if is_local_assum d then
+ let s = Retyping.get_sort_of env evars (get_type d) in
Univ.sup (univ_of_sort s) univ
else univ
in (push_rel d env, univ))
@@ -103,27 +107,33 @@ let typecheck_params_and_fields def id pl t ps nots fs =
in
List.iter
(function LocalRawDef (b, _) -> error default_binder_kind b
- | LocalRawAssum (ls, bk, ce) -> List.iter (error bk) ls) ps
+ | LocalRawAssum (ls, bk, ce) -> List.iter (error bk) ls
+ | LocalPattern _ -> assert false) ps
in
let impls_env, ((env1,newps), imps) = interp_context_evars env0 evars ps in
let t', template = match t with
| Some t ->
let env = push_rel_context newps env0 in
+ let poly =
+ match t with
+ | CSort (_, Misctypes.GType []) -> true | _ -> false in
let s = interp_type_evars env evars ~impls:empty_internalization_env t in
- let sred = Reductionops.whd_betadeltaiota env !evars s in
+ let sred = Reductionops.whd_all env !evars s in
(match kind_of_term sred with
| Sort s' ->
- (match Evd.is_sort_variable !evars s' with
- | Some l -> evars := Evd.make_flexible_variable !evars true (* (not def) *) l;
- sred, true
- | None -> s, false)
+ (if poly then
+ match Evd.is_sort_variable !evars s' with
+ | Some l -> evars := Evd.make_flexible_variable !evars true l;
+ sred, true
+ | None -> s, false
+ else s, false)
| _ -> user_err_loc (constr_loc t,"", str"Sort expected."))
| None ->
- let uvarkind = if (* not def *) true then Evd.univ_flexible_alg else Evd.univ_flexible in
- mkSort (Evarutil.evd_comb0 (Evd.new_sort_variable uvarkind) evars), false
+ let uvarkind = Evd.univ_flexible_alg in
+ mkSort (Evarutil.evd_comb0 (Evd.new_sort_variable uvarkind) evars), true
in
let fullarity = it_mkProd_or_LetIn t' newps in
- let env_ar = push_rel_context newps (push_rel (Name id,None,fullarity) env0) in
+ let env_ar = push_rel_context newps (push_rel (LocalAssum (Name id,fullarity)) env0) in
let env2,impls,newfs,data =
interp_fields_evars env_ar evars impls_env nots (binders_of_decls fs)
in
@@ -135,43 +145,51 @@ let typecheck_params_and_fields def id pl t ps nots fs =
let _, univ = compute_constructor_level evars env_ar newfs in
let ctx, aritysort = Reduction.dest_arity env0 arity in
assert(List.is_empty ctx); (* Ensured by above analysis *)
- if Sorts.is_prop aritysort ||
- (Sorts.is_set aritysort && is_impredicative_set env0) then
+ if not def && (Sorts.is_prop aritysort ||
+ (Sorts.is_set aritysort && is_impredicative_set env0)) then
arity, evars
else
let evars = Evd.set_leq_sort env_ar evars (Type univ) aritysort in
- if Univ.is_small_univ univ then
- (* We can assume that the level aritysort is not constrained
- and clear it. *)
- mkArity (ctx, Sorts.sort_of_univ univ),
- Evd.set_eq_sort env_ar evars (Prop Pos) aritysort
- else arity, evars
+ if Univ.is_small_univ univ &&
+ Option.cata (Evd.is_flexible_level evars) false (Evd.is_sort_variable evars aritysort) then
+ (* We can assume that the level in aritysort is not constrained
+ and clear it, if it is flexible *)
+ mkArity (ctx, Sorts.sort_of_univ univ),
+ Evd.set_eq_sort env_ar evars (Prop Pos) aritysort
+ else arity, evars
in
let evars, nf = Evarutil.nf_evars_and_universes evars in
- let newps = map_rel_context nf newps in
- let newfs = map_rel_context nf newfs in
- let ce t = Evarutil.check_evars env0 Evd.empty evars t in
- List.iter (fun (n, b, t) -> Option.iter ce b; ce t) (List.rev newps);
- List.iter (fun (n, b, t) -> Option.iter ce b; ce t) (List.rev newfs);
+ let newps = Context.Rel.map nf newps in
+ let newfs = Context.Rel.map nf newfs in
+ let ce t = Pretyping.check_evars env0 Evd.empty evars t in
+ List.iter (iter_constr ce) (List.rev newps);
+ List.iter (iter_constr ce) (List.rev newfs);
Evd.universe_context ?names:pl evars, nf arity, template, imps, newps, impls, newfs
-let degenerate_decl (na,b,t) =
- let id = match na with
+let degenerate_decl decl =
+ let id = match get_name decl with
| Name id -> id
| Anonymous -> anomaly (Pp.str "Unnamed record variable") in
- match b with
- | None -> (id, LocalAssum t)
- | Some b -> (id, LocalDef b)
+ match decl with
+ | LocalAssum (_,t) -> (id, LocalAssumEntry t)
+ | LocalDef (_,b,_) -> (id, LocalDefEntry b)
type record_error =
| MissingProj of Id.t * Id.t list
| BadTypedProj of Id.t * env * Type_errors.type_error
+let warn_cannot_define_projection =
+ CWarnings.create ~name:"cannot-define-projection" ~category:"records"
+ (fun msg -> hov 0 msg)
+
+(* If a projection is not definable, we throw an error if the user
+asked it to be a coercion. Otherwise, we just print an info
+message. The user might still want to name the field of the record. *)
let warning_or_error coe indsp err =
let st = match err with
| MissingProj (fi,projs) ->
let s,have = if List.length projs > 1 then "s","were" else "","was" in
- (str(Id.to_string fi) ++
+ (pr_id fi ++
strbrk" cannot be defined because the projection" ++ str s ++ spc () ++
prlist_with_sep pr_comma pr_id projs ++ spc () ++ str have ++
strbrk " not defined.")
@@ -191,7 +209,7 @@ let warning_or_error coe indsp err =
(pr_id fi ++ strbrk " cannot be defined because it is not typable.")
in
if coe then errorlabstrm "structure" st;
- Flags.if_verbose msg_warning (hov 0 st)
+ Flags.if_verbose Feedback.msg_info (hov 0 st)
type field_status =
| NoProjection of Name.t
@@ -234,6 +252,12 @@ let instantiate_possibly_recursive_type indu paramdecls fields =
let subst = List.map_i (fun i _ -> mkRel i) 1 paramdecls in
Termops.substl_rel_context (subst@[mkIndU indu]) fields
+let warn_non_primitive_record =
+ CWarnings.create ~name:"non-primitive-record" ~category:"record"
+ (fun (env,indsp) ->
+ (hov 0 (str "The record " ++ Printer.pr_inductive env indsp ++
+ strbrk" could not be defined as a primitive record")))
+
(* We build projections *)
let declare_projections indsp ?(kind=StructureComponent) binder_name coers fieldimpls fields =
let env = Global.env() in
@@ -244,8 +268,8 @@ let declare_projections indsp ?(kind=StructureComponent) binder_name coers field
let ctx = Univ.instantiate_univ_context mib.mind_universes in
let indu = indsp, u in
let r = mkIndU (indsp,u) in
- let rp = applist (r, Termops.extended_rel_list 0 paramdecls) in
- let paramargs = Termops.extended_rel_list 1 paramdecls in (*def in [[params;x:rp]]*)
+ let rp = applist (r, Context.Rel.to_extended_list 0 paramdecls) in
+ let paramargs = Context.Rel.to_extended_list 1 paramdecls in (*def in [[params;x:rp]]*)
let x = Name binder_name in
let fields = instantiate_possibly_recursive_type indu paramdecls fields in
let lifted_fields = Termops.lift_rel_context 1 fields in
@@ -257,31 +281,31 @@ let declare_projections indsp ?(kind=StructureComponent) binder_name coers field
| Some None | None -> false
in
if not is_primitive then
- Flags.if_verbose msg_warning
- (hov 0 (str "The record " ++ Printer.pr_inductive env indsp ++
- str" could not be defined as a primitive record"));
+ warn_non_primitive_record (env,indsp);
is_primitive
else false
in
let (_,_,kinds,sp_projs,_) =
List.fold_left3
- (fun (nfi,i,kinds,sp_projs,subst) coe (fi,optci,ti) impls ->
+ (fun (nfi,i,kinds,sp_projs,subst) coe decl impls ->
+ let fi = get_name decl in
+ let ti = get_type decl in
let (sp_projs,i,subst) =
match fi with
| Anonymous ->
(None::sp_projs,i,NoProjection fi::subst)
| Name fid -> try
let kn, term =
- if optci = None && primitive then
+ if is_local_assum decl && primitive then
(** Already defined in the kernel silently *)
let kn = destConstRef (Nametab.locate (Libnames.qualid_of_ident fid)) in
Declare.definition_message fid;
kn, mkProj (Projection.make kn false,mkRel 1)
else
let ccl = subst_projection fid subst ti in
- let body = match optci with
- | Some ci -> subst_projection fid subst ci
- | None ->
+ let body = match decl with
+ | LocalDef (_,ci,_) -> subst_projection fid subst ci
+ | LocalAssum _ ->
(* [ccl] is defined in context [params;x:rp] *)
(* [ccl'] is defined in context [params;x:rp;x:rp] *)
let ccl' = liftn 1 2 ccl in
@@ -323,28 +347,32 @@ let declare_projections indsp ?(kind=StructureComponent) binder_name coers field
let cl = Class.class_of_global (IndRef indsp) in
Class.try_add_new_coercion_with_source refi ~local:false poly ~source:cl
end;
- let i = if Option.is_empty optci then i+1 else i in
+ let i = if is_local_assum decl then i+1 else i in
(Some kn::sp_projs, i, Projection term::subst)
with NotDefinable why ->
warning_or_error coe indsp why;
(None::sp_projs,i,NoProjection fi::subst) in
- (nfi-1,i,(fi, Option.is_empty optci)::kinds,sp_projs,subst))
+ (nfi-1,i,(fi, is_local_assum decl)::kinds,sp_projs,subst))
(List.length fields,0,[],[],[]) coers (List.rev fields) (List.rev fieldimpls)
in (kinds,sp_projs)
let structure_signature ctx =
let rec deps_to_evar evm l =
match l with [] -> Evd.empty
- | [(_,_,typ)] ->
+ | [decl] ->
let env = Environ.empty_named_context_val in
- let (evm, _) = Evarutil.new_pure_evar env evm typ in
+ let evm = Sigma.Unsafe.of_evar_map evm in
+ let Sigma (_, evm, _) = Evarutil.new_pure_evar env evm (get_type decl) in
+ let evm = Sigma.to_evar_map evm in
evm
- | (_,_,typ)::tl ->
+ | decl::tl ->
let env = Environ.empty_named_context_val in
- let (evm, ev) = Evarutil.new_pure_evar env evm typ in
+ let evm = Sigma.Unsafe.of_evar_map evm in
+ let Sigma (ev, evm, _) = Evarutil.new_pure_evar env evm (get_type decl) in
+ let evm = Sigma.to_evar_map evm in
let new_tl = Util.List.map_i
- (fun pos (n,c,t) -> n,c,
- Termops.replace_term (mkRel pos) (mkEvar(ev,[||])) t) 1 tl in
+ (fun pos decl ->
+ map_type (fun t -> Termops.replace_term (mkRel pos) (mkEvar(ev,[||])) t) decl) 1 tl in
deps_to_evar evm new_tl in
deps_to_evar Evd.empty (List.rev ctx)
@@ -353,7 +381,7 @@ open Typeclasses
let declare_structure finite poly ctx id idbuild paramimpls params arity template
fieldimpls fields ?(kind=StructureComponent) ?name is_coe coers sign =
let nparams = List.length params and nfields = List.length fields in
- let args = Termops.extended_rel_list nfields params in
+ let args = Context.Rel.to_extended_list nfields params in
let ind = applist (mkRel (1+nparams+nfields), args) in
let type_constructor = it_mkProd_or_LetIn ind fields in
let binder_name =
@@ -375,7 +403,9 @@ let declare_structure finite poly ctx id idbuild paramimpls params arity templat
mind_entry_inds = [mie_ind];
mind_entry_polymorphic = poly;
mind_entry_private = None;
- mind_entry_universes = ctx } in
+ mind_entry_universes = ctx;
+ }
+ in
let kn = Command.declare_mutual_inductive_with_eliminations mie [] [(paramimpls,[])] in
let rsp = (kn,0) in (* This is ind path of idstruc *)
let cstr = (rsp,1) in
@@ -392,7 +422,7 @@ let implicits_of_context ctx =
| Name n -> Some n
| Anonymous -> None
in ExplByPos (i, explname), (true, true, true))
- 1 (List.rev (Anonymous :: (List.map pi1 ctx)))
+ 1 (List.rev (Anonymous :: (List.map get_name ctx)))
let declare_class finite def poly ctx id idbuild paramimpls params arity
template fieldimpls fields ?(kind=StructureComponent) is_coe coers priorities sign =
@@ -405,11 +435,11 @@ let declare_class finite def poly ctx id idbuild paramimpls params arity
let binder_name = Namegen.next_ident_away (snd id) (Termops.ids_of_context (Global.env())) in
let impl, projs =
match fields with
- | [(Name proj_name, _, field)] when def ->
+ | [LocalAssum (Name proj_name, field) | LocalDef (Name proj_name, _, field)] when def ->
let class_body = it_mkLambda_or_LetIn field params in
- let _class_type = it_mkProd_or_LetIn arity params in
+ let class_type = it_mkProd_or_LetIn arity params in
let class_entry =
- Declare.definition_entry (* ?types:class_type *) ~poly ~univs:ctx class_body in
+ Declare.definition_entry ~types:class_type ~poly ~univs:ctx class_body in
let cst = Declare.declare_constant (snd id)
(DefinitionEntry class_entry, IsDefinition Definition)
in
@@ -446,13 +476,13 @@ let declare_class finite def poly ctx id idbuild paramimpls params arity
if b then Backward, pri else Forward, pri) coe)
coers priorities
in
- let l = List.map3 (fun (id, _, _) b y -> (id, b, y))
+ let l = List.map3 (fun decl b y -> get_name decl, b, y)
(List.rev fields) coers (Recordops.lookup_projections ind)
in IndRef ind, l
in
let ctx_context =
- List.map (fun (na, b, t) ->
- match Typeclasses.class_of_constr t with
+ List.map (fun decl ->
+ match Typeclasses.class_of_constr (get_type decl) with
| Some (_, ((cl,_), _)) -> Some (cl.cl_impl, true)
| None -> None)
params, params
@@ -474,7 +504,7 @@ let add_constant_class cst =
let tc =
{ cl_impl = ConstRef cst;
cl_context = (List.map (const None) ctx, ctx);
- cl_props = [(Anonymous, None, arity)];
+ cl_props = [LocalAssum (Anonymous, arity)];
cl_projs = [];
cl_strict = !typeclasses_strict;
cl_unique = !typeclasses_unique
@@ -487,19 +517,13 @@ let add_inductive_class ind =
let k =
let ctx = oneind.mind_arity_ctxt in
let inst = Univ.UContext.instance mind.mind_universes in
- let map = function
- | (_, Some _, _) -> None
- | (_, None, t) -> Some (lazy t)
- in
- let args = List.map_filter map ctx in
- let ty = Inductive.type_of_inductive_knowing_parameters
+ let ty = Inductive.type_of_inductive
(push_rel_context ctx (Global.env ()))
((mind,oneind),inst)
- (Array.of_list args)
in
{ cl_impl = IndRef ind;
cl_context = List.map (const None) ctx, ctx;
- cl_props = [Anonymous, None, ty];
+ cl_props = [LocalAssum (Anonymous, ty)];
cl_projs = [];
cl_strict = !typeclasses_strict;
cl_unique = !typeclasses_unique }
@@ -516,7 +540,7 @@ open Vernacexpr
(* [fs] corresponds to fields and [ps] to parameters; [coers] is a
list telling if the corresponding fields must me declared as coercions
- or subinstances *)
+ or subinstances. *)
let definition_structure (kind,poly,finite,(is_coe,((loc,idstruc),pl)),ps,cfs,idbuild,s) =
let cfs,notations = List.split cfs in
let cfs,priorities = List.split cfs in
@@ -537,15 +561,16 @@ let definition_structure (kind,poly,finite,(is_coe,((loc,idstruc),pl)),ps,cfs,id
typecheck_params_and_fields (kind = Class true) idstruc pl s ps notations fs) () in
let sign = structure_signature (fields@params) in
let gr = match kind with
- | Class def ->
- let gr = declare_class finite def poly ctx (loc,idstruc) idbuild
+ | Class def ->
+ let priorities = List.map (fun id -> {hint_priority = id; hint_pattern = None}) priorities in
+ let gr = declare_class finite def poly ctx (loc,idstruc) idbuild
implpars params arity template implfs fields is_coe coers priorities sign in
gr
| _ ->
let implfs = List.map
(fun impls -> implpars @ Impargs.lift_implicits
(succ (List.length params)) impls) implfs in
- let ind = declare_structure finite poly ctx idstruc
+ let ind = declare_structure finite poly ctx idstruc
idbuild implpars params arity template implfs
fields is_coe (List.map (fun coe -> not (Option.is_empty coe)) coers) sign in
IndRef ind
diff --git a/toplevel/record.mli b/toplevel/record.mli
index 4ce27755..c50e5778 100644
--- a/toplevel/record.mli
+++ b/toplevel/record.mli
@@ -8,7 +8,6 @@
open Names
open Term
-open Context
open Vernacexpr
open Constrexpr
open Impargs
@@ -22,15 +21,16 @@ val primitive_flag : bool ref
val declare_projections :
inductive -> ?kind:Decl_kinds.definition_object_kind -> Id.t ->
- coercion_flag list -> manual_explicitation list list -> rel_context ->
+ coercion_flag list -> manual_explicitation list list -> Context.Rel.t ->
(Name.t * bool) list * constant option list
-val declare_structure : Decl_kinds.recursivity_kind ->
+val declare_structure :
+ Decl_kinds.recursivity_kind ->
bool (** polymorphic?*) -> Univ.universe_context ->
Id.t -> Id.t ->
- manual_explicitation list -> rel_context -> (** params *) constr -> (** arity *)
+ manual_explicitation list -> Context.Rel.t -> (** params *) constr -> (** arity *)
bool (** template arity ? *) ->
- Impargs.manual_explicitation list list -> rel_context -> (** fields *)
+ Impargs.manual_explicitation list list -> Context.Rel.t -> (** fields *)
?kind:Decl_kinds.definition_object_kind -> ?name:Id.t ->
bool -> (** coercion? *)
bool list -> (** field coercions *)
@@ -38,7 +38,8 @@ val declare_structure : Decl_kinds.recursivity_kind ->
inductive
val definition_structure :
- inductive_kind * Decl_kinds.polymorphic * Decl_kinds.recursivity_kind * plident with_coercion * local_binder list *
+ inductive_kind * Decl_kinds.polymorphic * Decl_kinds.recursivity_kind *
+ plident with_coercion * local_binder list *
(local_decl_expr with_instance with_priority with_notation) list *
Id.t * constr_expr option -> global_reference
diff --git a/toplevel/search.ml b/toplevel/search.ml
index d7a4cbe7..ff3c7a4f 100644
--- a/toplevel/search.ml
+++ b/toplevel/search.ml
@@ -18,10 +18,17 @@ open Printer
open Libnames
open Globnames
open Nametab
+open Goptions
type filter_function = global_reference -> env -> constr -> bool
type display_function = global_reference -> env -> constr -> unit
+(* This option restricts the output of [SearchPattern ...],
+[SearchAbout ...], etc. to the names of the symbols matching the
+query, separated by a newline. This type of output is useful for
+editors (like emacs), to generate a list of completion candidates
+without having to parse thorugh the types of all symbols. *)
+
type glob_search_about_item =
| GlobSearchSubPattern of constr_pattern
| GlobSearchString of string
@@ -49,7 +56,9 @@ let iter_constructors indsp u fn env nconstr =
fn (ConstructRef (indsp, i)) env typ
done
-let iter_named_context_name_type f = List.iter (fun (nme,_,typ) -> f nme typ)
+let iter_named_context_name_type f =
+ let open Context.Named.Declaration in
+ List.iter (fun decl -> f (get_id decl) (get_type decl))
(* General search over hypothesis of a goal *)
let iter_hypothesis glnum (fn : global_reference -> env -> constr -> unit) =
@@ -61,12 +70,13 @@ let iter_hypothesis glnum (fn : global_reference -> env -> constr -> unit) =
(* General search over declarations *)
let iter_declarations (fn : global_reference -> env -> constr -> unit) =
+ let open Context.Named.Declaration in
let env = Global.env () in
let iter_obj (sp, kn) lobj = match object_tag lobj with
| "VARIABLE" ->
begin try
- let (id, _, typ) = Global.lookup_named (basename sp) in
- fn (VarRef id) env typ
+ let decl = Global.lookup_named (basename sp) in
+ fn (VarRef (get_id decl)) env (get_type decl)
with Not_found -> (* we are in a section *) () end
| "CONSTANT" ->
let cst = Global.constant_of_delta_kn kn in
@@ -97,15 +107,6 @@ let generic_search glnumopt fn =
| Some glnum -> iter_hypothesis glnum fn);
iter_declarations fn
-(** Standard display *)
-
-let plain_display accu ref env c =
- let pc = pr_lconstr_env env Evd.empty c in
- let pr = pr_global ref in
- accu := hov 2 (pr ++ str":" ++ spc () ++ pc) :: !accu
-
-let format_display l = prlist_with_sep fnl (fun x -> x) (List.rev l)
-
(** Filters *)
(** This function tries to see whether the conclusion matches a pattern. *)
@@ -131,8 +132,9 @@ let full_name_of_reference ref =
DirPath.to_string dir ^ "." ^ Id.to_string id
(** Whether a reference is blacklisted *)
-let blacklist_filter ref env typ =
+let blacklist_filter_aux () =
let l = SearchBlacklist.elements () in
+ fun ref env typ ->
let name = full_name_of_reference ref in
let is_not_bl str = not (String.string_contains ~where:name ~what:str) in
List.for_all is_not_bl l
@@ -156,19 +158,17 @@ let search_about_filter query gr env typ = match query with
(** SearchPattern *)
-let search_pattern gopt pat mods =
- let ans = ref [] in
+let search_pattern gopt pat mods pr_search =
+ let blacklist_filter = blacklist_filter_aux () in
let filter ref env typ =
- let f_module = module_filter mods ref env typ in
- let f_blacklist = blacklist_filter ref env typ in
- let f_pattern () = pattern_filter pat ref env typ in
- f_module && f_pattern () && f_blacklist
+ module_filter mods ref env typ &&
+ pattern_filter pat ref env typ &&
+ blacklist_filter ref env typ
in
let iter ref env typ =
- if filter ref env typ then plain_display ans ref env typ
+ if filter ref env typ then pr_search ref env typ
in
- let () = generic_search gopt iter in
- format_display !ans
+ generic_search gopt iter
(** SearchRewrite *)
@@ -180,63 +180,56 @@ let rewrite_pat1 pat =
let rewrite_pat2 pat =
PApp (PRef eq, [| PMeta None; PMeta None; pat |])
-let search_rewrite gopt pat mods =
+let search_rewrite gopt pat mods pr_search =
let pat1 = rewrite_pat1 pat in
let pat2 = rewrite_pat2 pat in
- let ans = ref [] in
+ let blacklist_filter = blacklist_filter_aux () in
let filter ref env typ =
- let f_module = module_filter mods ref env typ in
- let f_blacklist = blacklist_filter ref env typ in
- let f_pattern () =
- pattern_filter pat1 ref env typ ||
- pattern_filter pat2 ref env typ
- in
- f_module && f_pattern () && f_blacklist
+ module_filter mods ref env typ &&
+ (pattern_filter pat1 ref env typ ||
+ pattern_filter pat2 ref env typ) &&
+ blacklist_filter ref env typ
in
let iter ref env typ =
- if filter ref env typ then plain_display ans ref env typ
+ if filter ref env typ then pr_search ref env typ
in
- let () = generic_search gopt iter in
- format_display !ans
+ generic_search gopt iter
(** Search *)
-let search_by_head gopt pat mods =
- let ans = ref [] in
+let search_by_head gopt pat mods pr_search =
+ let blacklist_filter = blacklist_filter_aux () in
let filter ref env typ =
- let f_module = module_filter mods ref env typ in
- let f_blacklist = blacklist_filter ref env typ in
- let f_pattern () = head_filter pat ref env typ in
- f_module && f_pattern () && f_blacklist
+ module_filter mods ref env typ &&
+ head_filter pat ref env typ &&
+ blacklist_filter ref env typ
in
let iter ref env typ =
- if filter ref env typ then plain_display ans ref env typ
+ if filter ref env typ then pr_search ref env typ
in
- let () = generic_search gopt iter in
- format_display !ans
+ generic_search gopt iter
(** SearchAbout *)
-let search_about gopt items mods =
- let ans = ref [] in
+let search_about gopt items mods pr_search =
+ let blacklist_filter = blacklist_filter_aux () in
let filter ref env typ =
let eqb b1 b2 = if b1 then b2 else not b2 in
- let f_module = module_filter mods ref env typ in
- let f_about (b, i) = eqb b (search_about_filter i ref env typ) in
- let f_blacklist = blacklist_filter ref env typ in
- f_module && List.for_all f_about items && f_blacklist
+ module_filter mods ref env typ &&
+ List.for_all
+ (fun (b,i) -> eqb b (search_about_filter i ref env typ)) items &&
+ blacklist_filter ref env typ
in
let iter ref env typ =
- if filter ref env typ then plain_display ans ref env typ
+ if filter ref env typ then pr_search ref env typ
in
- let () = generic_search gopt iter in
- format_display !ans
+ generic_search gopt iter
type search_constraint =
- | Name_Pattern of string
- | Type_Pattern of string
- | SubType_Pattern of string
- | In_Module of string list
+ | Name_Pattern of Str.regexp
+ | Type_Pattern of Pattern.constr_pattern
+ | SubType_Pattern of Pattern.constr_pattern
+ | In_Module of Names.DirPath.t
| Include_Blacklist
type 'a coq_object = {
@@ -245,43 +238,25 @@ type 'a coq_object = {
coq_object_object : 'a;
}
-let interface_search flags =
- let env = Global.env () in
+let interface_search =
let rec extract_flags name tpe subtpe mods blacklist = function
| [] -> (name, tpe, subtpe, mods, blacklist)
- | (Name_Pattern s, b) :: l ->
- let regexp =
- try Str.regexp s
- with e when Errors.noncritical e ->
- Errors.errorlabstrm "Search.interface_search"
- (str "Invalid regexp: " ++ str s)
- in
+ | (Name_Pattern regexp, b) :: l ->
extract_flags ((regexp, b) :: name) tpe subtpe mods blacklist l
- | (Type_Pattern s, b) :: l ->
- let constr = Pcoq.parse_string Pcoq.Constr.lconstr_pattern s in
- let (_, pat) = Constrintern.intern_constr_pattern env constr in
+ | (Type_Pattern pat, b) :: l ->
extract_flags name ((pat, b) :: tpe) subtpe mods blacklist l
- | (SubType_Pattern s, b) :: l ->
- let constr = Pcoq.parse_string Pcoq.Constr.lconstr_pattern s in
- let (_, pat) = Constrintern.intern_constr_pattern env constr in
+ | (SubType_Pattern pat, b) :: l ->
extract_flags name tpe ((pat, b) :: subtpe) mods blacklist l
- | (In_Module m, b) :: l ->
- let path = String.concat "." m in
- let m = Pcoq.parse_string Pcoq.Constr.global path in
- let (_, qid) = Libnames.qualid_of_reference m in
- let id =
- try Nametab.full_name_module qid
- with Not_found ->
- Errors.errorlabstrm "Search.interface_search"
- (str "Module " ++ str path ++ str " not found.")
- in
+ | (In_Module id, b) :: l ->
extract_flags name tpe subtpe ((id, b) :: mods) blacklist l
| (Include_Blacklist, b) :: l ->
extract_flags name tpe subtpe mods b l
in
+ fun ?glnum flags ->
let (name, tpe, subtpe, mods, blacklist) =
extract_flags [] [] [] [] false flags
in
+ let blacklist_filter = blacklist_filter_aux () in
let filter_function ref env constr =
let id = Names.Id.to_string (Nametab.basename_of_global ref) in
let path = Libnames.dirpath (Nametab.path_of_global ref) in
@@ -300,13 +275,11 @@ let interface_search flags =
let match_module (mdl, flag) =
toggle (Libnames.is_dirpath_prefix_of mdl path) flag
in
- let in_blacklist =
- blacklist || (blacklist_filter ref env constr)
- in
List.for_all match_name name &&
List.for_all match_type tpe &&
List.for_all match_subtype subtpe &&
- List.for_all match_module mods && in_blacklist
+ List.for_all match_module mods &&
+ (blacklist || blacklist_filter ref env constr)
in
let ans = ref [] in
let print_function ref env constr =
@@ -335,5 +308,8 @@ let interface_search flags =
let iter ref env typ =
if filter_function ref env typ then print_function ref env typ
in
- let () = generic_search None iter in (* TODO: chose a goal number? *)
+ let () = generic_search glnum iter in
!ans
+
+let blacklist_filter ref env typ =
+ blacklist_filter_aux () ref env typ
diff --git a/toplevel/search.mli b/toplevel/search.mli
index 78b0c45c..ba3d48ef 100644
--- a/toplevel/search.mli
+++ b/toplevel/search.mli
@@ -39,21 +39,24 @@ val search_about_filter : glob_search_about_item -> filter_function
goal and the global environment for things matching [pattern] and
satisfying module exclude/include clauses of [modinout]. *)
-val search_by_head : int option -> constr_pattern -> DirPath.t list * bool -> std_ppcmds
-val search_rewrite : int option -> constr_pattern -> DirPath.t list * bool -> std_ppcmds
-val search_pattern : int option -> constr_pattern -> DirPath.t list * bool -> std_ppcmds
+val search_by_head : int option -> constr_pattern -> DirPath.t list * bool
+ -> display_function -> unit
+val search_rewrite : int option -> constr_pattern -> DirPath.t list * bool
+ -> display_function -> unit
+val search_pattern : int option -> constr_pattern -> DirPath.t list * bool
+ -> display_function -> unit
val search_about : int option -> (bool * glob_search_about_item) list
- -> DirPath.t list * bool -> std_ppcmds
+ -> DirPath.t list * bool -> display_function -> unit
type search_constraint =
(** Whether the name satisfies a regexp (uses Ocaml Str syntax) *)
- | Name_Pattern of string
+ | Name_Pattern of Str.regexp
(** Whether the object type satisfies a pattern *)
- | Type_Pattern of string
+ | Type_Pattern of Pattern.constr_pattern
(** Whether some subtype of object type satisfies a pattern *)
- | SubType_Pattern of string
+ | SubType_Pattern of Pattern.constr_pattern
(** Whether the object pertains to a module *)
- | In_Module of string list
+ | In_Module of Names.DirPath.t
(** Bypass the Search blacklist *)
| Include_Blacklist
@@ -63,7 +66,7 @@ type 'a coq_object = {
coq_object_object : 'a;
}
-val interface_search : (search_constraint * bool) list ->
+val interface_search : ?glnum:int -> (search_constraint * bool) list ->
string coq_object list
(** {6 Generic search function} *)
diff --git a/toplevel/toplevel.mllib b/toplevel/toplevel.mllib
index 5aa7d428..d6892236 100644
--- a/toplevel/toplevel.mllib
+++ b/toplevel/toplevel.mllib
@@ -1,5 +1,5 @@
Himsg
-Cerrors
+ExplainErr
Class
Locality
Metasyntax
diff --git a/toplevel/usage.ml b/toplevel/usage.ml
index 4280006b..38ceacf5 100644
--- a/toplevel/usage.ml
+++ b/toplevel/usage.ml
@@ -11,6 +11,10 @@ let version ret =
Coq_config.version Coq_config.date;
Printf.printf "compiled on %s with OCaml %s\n" Coq_config.compile_date Coq_config.caml_version;
exit ret
+let machine_readable_version ret =
+ Printf.printf "%s %s\n"
+ Coq_config.version Coq_config.caml_version;
+ exit ret
(* print the usage of coqtop (or coqc) on channel co *)
@@ -32,8 +36,6 @@ let print_usage_channel co command =
\n -noinit start without loading the Init library\
\n -nois (idem)\
\n -compat X.Y provides compatibility support for Coq version X.Y\
-\n -verbose-compat-notations be warned when using compatibility notations\
-\n -no-compat-notations get an error when using compatibility notations\
\n\
\n -load-ml-object f load ML object file f\
\n -load-ml-source f load ML file f\
@@ -45,6 +47,7 @@ let print_usage_channel co command =
\n -require path load Coq library path and import it (Require Import path.)\
\n -compile f.v compile Coq file f.v (implies -batch)\
\n -compile-verbose f.v verbosely compile Coq file f.v (implies -batch)\
+\n -o f.vo use f.vo as the output file name\
\n -quick quickly compile .v files to .vio files (skip proofs)\
\n -schedule-vio2vo j f1..fn run up to j instances of Coq to turn each fi.vio\
\n into fi.vo\
@@ -56,8 +59,8 @@ let print_usage_channel co command =
\n -v print Coq version and exit\
\n -list-tags print highlight color tags known by Coq and exit\
\n\
-\n -quiet unset display of extra information (implies -w none)\
-\n -w (all|none) configure display of warnings\
+\n -quiet unset display of extra information (implies -w \"-all\")\
+\n -w (w1,..,wn) configure display of warnings\
\n -color (yes|no|auto) configure color output\
\n\
\n -q skip loading of rcfile\
@@ -77,6 +80,7 @@ let print_usage_channel co command =
\n the directory $COQ_XML_LIBRARY_ROOT (if set) or to\
\n stdout (if unset)\
\n -time display the time taken by each command\
+\n -profile-ltac display the time taken by each (sub)tactic\
\n -m, --memory display total heap size at program exit\
\n (use environment variable\
\n OCAML_GC_STATS=\"/tmp/gclog.txt\"\
@@ -117,12 +121,7 @@ let print_config () =
if Coq_config.local then Printf.printf "LOCAL=1\n" else Printf.printf "LOCAL=0\n";
Printf.printf "COQLIB=%s/\n" (Envars.coqlib ());
Printf.printf "DOCDIR=%s/\n" (Envars.docdir ());
- Printf.printf "OCAMLDEP=%s\n" Coq_config.ocamldep;
- Printf.printf "OCAMLC=%s\n" Coq_config.ocamlc;
- Printf.printf "OCAMLOPT=%s\n" Coq_config.ocamlopt;
- Printf.printf "OCAMLDOC=%s\n" Coq_config.ocamldoc;
- Printf.printf "CAMLBIN=%s/\n" (Envars.camlbin ());
- Printf.printf "CAMLLIB=%s/\n" (Envars.camllib ());
+ Printf.printf "OCAMLFIND=%s\n" (Envars.ocamlfind ());
Printf.printf "CAMLP4=%s\n" Coq_config.camlp4;
Printf.printf "CAMLP4O=%s\n" Coq_config.camlp4o;
Printf.printf "CAMLP4BIN=%s/\n" (Envars.camlp4bin ());
diff --git a/toplevel/usage.mli b/toplevel/usage.mli
index 3ce9e93e..dccb40e7 100644
--- a/toplevel/usage.mli
+++ b/toplevel/usage.mli
@@ -9,6 +9,7 @@
(** {6 Prints the version number on the standard output and exits (with 0). } *)
val version : int -> 'a
+val machine_readable_version : int -> 'a
(** {6 Prints the usage on the error output, preceeded by a user-provided message. } *)
val print_usage : string -> unit
diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml
index 7c4920df..bfdae85d 100644
--- a/toplevel/vernac.ml
+++ b/toplevel/vernac.ml
@@ -9,27 +9,29 @@
(* Parsing of vernacular. *)
open Pp
-open Errors
+open CErrors
open Util
open Flags
-open System
open Vernacexpr
(* The functions in this module may raise (unexplainable!) exceptions.
Use the module Coqtoplevel, which catches these exceptions
(the exceptions are explained only at the toplevel). *)
-(* The navigation vernac commands will be handled separately *)
+let user_error loc s = CErrors.user_err_loc (loc,"_",str s)
+
+(* Navigation commands are allowed in a coqtop session but not in a .v file *)
let rec is_navigation_vernac = function
| VernacResetInitial
| VernacResetName _
| VernacBacktrack _
| VernacBackTo _
- | VernacBack _ -> true
- | VernacRedirect (_, l) | VernacTime l ->
- List.exists
- (fun (_,c) -> is_navigation_vernac c) l (* Time Back* is harmless *)
+ | VernacBack _
+ | VernacStm _ -> true
+ | VernacRedirect (_, (_,c))
+ | VernacTime (_,c) ->
+ is_navigation_vernac c (* Time Back* is harmless *)
| c -> is_deep_navigation_vernac c
and is_deep_navigation_vernac = function
@@ -42,6 +44,14 @@ let is_reset = function
| VernacResetInitial | VernacResetName _ -> true
| _ -> false
+let checknav_simple loc cmd =
+ if is_navigation_vernac cmd && not (is_reset cmd) then
+ user_error loc "Navigation commands forbidden in files."
+
+let checknav_deep loc ast =
+ if is_deep_navigation_vernac ast then
+ user_error loc "Navigation commands forbidden in nested commands."
+
(* When doing Load on a file, two behaviors are possible:
- either the history stack is grown by only one command,
@@ -66,32 +76,10 @@ let _ =
Goptions.optread = (fun () -> !atomic_load);
Goptions.optwrite = ((:=) atomic_load) }
-(* In case of error, register the file being currently Load'ed and the
- inner file in which the error has been encountered. Any intermediate files
- between the two are discarded. *)
-
-type location_files = { outer : string; inner : string }
-
-let files_of_exn : location_files Exninfo.t = Exninfo.make ()
-
-let get_exn_files e = Exninfo.get e files_of_exn
-
-let add_exn_files e f = Exninfo.add e files_of_exn f
-
-let enrich_with_file f (e, info) =
- let inner = match get_exn_files info with None -> f | Some x -> x.inner in
- (e, add_exn_files info { outer = f; inner })
-
-let raise_with_file f e = iraise (enrich_with_file f e)
-
-let cur_file = ref None
-
let disable_drop = function
- | Drop -> Errors.error "Drop is forbidden."
+ | Drop -> CErrors.error "Drop is forbidden."
| e -> e
-let user_error loc s = Errors.user_err_loc (loc,"_",str s)
-
(* Opening and closing a channel. Open it twice when verbose: the first
channel is used to read the commands, and the second one to print them.
Note: we could use only one thanks to seek_in, but seeking on and on in
@@ -101,7 +89,7 @@ let open_file_twice_if verbosely longfname =
let in_chan = open_utf8_file_in longfname in
let verb_ch =
if verbosely then Some (open_utf8_file_in longfname) else None in
- let po = Pcoq.Gram.parsable (Stream.of_channel in_chan) in
+ let po = Pcoq.Gram.parsable ~file:longfname (Stream.of_channel in_chan) in
(in_chan, longfname, (po, verb_ch))
let close_input in_chan (_,verb) =
@@ -110,7 +98,7 @@ let close_input in_chan (_,verb) =
match verb with
| Some verb_ch -> close_in verb_ch
| _ -> ()
- with e when Errors.noncritical e -> ()
+ with e when CErrors.noncritical e -> ()
let verbose_phrase verbch loc =
let loc = Loc.unloc loc in
@@ -120,8 +108,7 @@ let verbose_phrase verbch loc =
let s = String.create len in
seek_in ch (fst loc);
really_input ch s 0 len;
- ppnl (str s);
- pp_flush()
+ Feedback.msg_notice (str s)
| None -> ()
exception End_of_input
@@ -135,51 +122,44 @@ let parse_sentence = Flags.with_option Flags.we_are_parsing
(* vernac parses the given stream, executes interpfun on the syntax tree it
* parses, and is verbose on "primitives" commands if verbosely is true *)
-let just_parsing = ref false
let chan_beautify = ref stdout
let beautify_suffix = ".beautified"
-let set_formatter_translator() =
- let ch = !chan_beautify in
+let set_formatter_translator ch =
let out s b e = output ch s b e in
Format.set_formatter_output_functions out (fun () -> flush ch);
Format.set_max_boxes max_int
-let pr_new_syntax loc ocom =
+let pr_new_syntax_in_context loc chan_beautify ocom =
let loc = Loc.unloc loc in
- if !beautify_file then set_formatter_translator();
+ if !beautify_file then set_formatter_translator chan_beautify;
let fs = States.freeze ~marshallable:`No in
- let com = match ocom with
- | Some VernacNop -> mt()
- | Some com -> Ppvernac.pr_vernac com
- | None -> mt() in
- if !beautify_file then
- msg (hov 0 (comment (fst loc) ++ com ++ comment (snd loc)))
- else
- msg_info (hov 4 (str"New Syntax:" ++ fnl() ++ (hov 0 com)));
- States.unfreeze fs;
- Format.set_formatter_out_channel stdout
-
-let save_translator_coqdoc () =
- (* translator state *)
- let ch = !chan_beautify in
- let cl = !Pp.comments in
- let cs = Lexer.com_state() in
- (* end translator state *)
- let coqdocstate = Lexer.location_table () in
- ch,cl,cs,coqdocstate
-
-let restore_translator_coqdoc (ch,cl,cs,coqdocstate) =
- if !Flags.beautify_file then close_out !chan_beautify;
- chan_beautify := ch;
- Pp.comments := cl;
- Lexer.restore_com_state cs;
- Lexer.restore_location_table coqdocstate
+ (* The content of this is not supposed to fail, but if ever *)
+ try
+ (* Side-effect: order matters *)
+ let before = comment (CLexer.extract_comments (fst loc)) in
+ let com = match ocom with
+ | Some com -> Ppvernac.pr_vernac com
+ | None -> mt() in
+ let after = comment (CLexer.extract_comments (snd loc)) in
+ if !beautify_file then
+ Pp.msg_with !Pp_control.std_ft (hov 0 (before ++ com ++ after))
+ else
+ Feedback.msg_info (hov 4 (str"New Syntax:" ++ fnl() ++ (hov 0 com)));
+ States.unfreeze fs;
+ Format.set_formatter_out_channel stdout
+ with any ->
+ States.unfreeze fs;
+ Format.set_formatter_out_channel stdout
+
+let pr_new_syntax po loc chan_beautify ocom =
+ (* Reinstall the context of parsing which includes the bindings of comments to locations *)
+ Pcoq.Gram.with_parsable po (pr_new_syntax_in_context chan_beautify loc) ocom
(* For coqtop -time, we display the position in the file,
and a glimpse of the executed command *)
-let display_cmd_header loc com =
+let pp_cmd_header loc com =
let shorten s = try (String.sub s 0 30)^"..." with _ -> s in
let noblank s =
for i = 0 to String.length s - 1 do
@@ -194,75 +174,63 @@ let display_cmd_header loc com =
try Ppvernac.pr_vernac x
with e -> str (Printexc.to_string e) in
let cmd = noblank (shorten (string_of_ppcmds (safe_pr_vernac com)))
- in
- Pp.pp (str "Chars " ++ int start ++ str " - " ++ int stop ++
- str " [" ++ str cmd ++ str "] ");
- Pp.flush_all ()
+ in str "Chars " ++ int start ++ str " - " ++ int stop ++
+ str " [" ++ str cmd ++ str "] "
-let rec vernac_com verbose checknav (loc,com) =
+(* This is a special case where we assume we are in console batch mode
+ and take control of the console.
+ *)
+let print_cmd_header loc com =
+ Pp.pp_with ~pp_tag:Ppstyle.pp_tag !Pp_control.std_ft (pp_cmd_header loc com);
+ Format.pp_print_flush !Pp_control.std_ft ()
+
+let rec interp_vernac po chan_beautify checknav (loc,com) =
let interp = function
| VernacLoad (verbosely, fname) ->
- let fname = Envars.expand_path_macros ~warn:(fun x -> msg_warning (str x)) fname in
+ let fname = Envars.expand_path_macros ~warn:(fun x -> Feedback.msg_warning (str x)) fname in
let fname = CUnix.make_suffix fname ".v" in
let f = Loadpath.locate_file fname in
- let st = save_translator_coqdoc () in
- if !Flags.beautify_file then
- begin
- chan_beautify := open_out (f^beautify_suffix);
- Pp.comments := []
- end;
- begin
- try
- read_vernac_file verbosely f;
- restore_translator_coqdoc st;
- with reraise ->
- let reraise = Errors.push reraise in
- restore_translator_coqdoc st;
- iraise reraise
- end
-
- | v when !just_parsing -> ()
-
- | v -> Stm.interp verbose (loc,v)
+ load_vernac verbosely f
+
+ | v -> Stm.interp (Flags.is_verbose()) (loc,v)
in
try
checknav loc com;
- if do_beautify () then pr_new_syntax loc (Some com);
- if !Flags.time then display_cmd_header loc com;
- let com = if !Flags.time then VernacTime [loc,com] else com in
+ if !beautify then pr_new_syntax po chan_beautify loc (Some com);
+ (* XXX: This is not 100% correct if called from an IDE context *)
+ if !Flags.time then print_cmd_header loc com;
+ let com = if !Flags.time then VernacTime (loc,com) else com in
interp com
with reraise ->
- let (reraise, info) = Errors.push reraise in
- Format.set_formatter_out_channel stdout;
+ let (reraise, info) = CErrors.push reraise in
let loc' = Option.default Loc.ghost (Loc.get_loc info) in
if Loc.is_ghost loc' then iraise (reraise, Loc.add_loc info loc)
else iraise (reraise, info)
-and read_vernac_file verbosely s =
- let checknav loc cmd =
- if is_navigation_vernac cmd && not (is_reset cmd) then
- user_error loc "Navigation commands forbidden in files"
- in
- let (in_chan, fname, input) = open_file_twice_if verbosely s in
- cur_file := Some fname;
+(* Load a vernac file. CErrors are annotated with file and location *)
+and load_vernac verbosely file =
+ let chan_beautify =
+ if !Flags.beautify_file then open_out (file^beautify_suffix) else stdout in
+ let (in_chan, fname, input) = open_file_twice_if verbosely file in
try
(* we go out of the following infinite loop when a End_of_input is
* raised, which means that we raised the end of the file being loaded *)
while true do
- let loc_ast = parse_sentence input in
- vernac_com verbosely checknav loc_ast;
- pp_flush ()
+ let loc_ast = Flags.silently parse_sentence input in
+ CWarnings.set_current_loc (fst loc_ast);
+ Flags.silently (interp_vernac (fst input) chan_beautify checknav_simple) loc_ast;
done
with any -> (* whatever the exception *)
- let (e, info) = Errors.push any in
- Format.set_formatter_out_channel stdout;
+ let (e, info) = CErrors.push any in
close_input in_chan input; (* we must close the file first *)
match e with
| End_of_input ->
- cur_file := None;
- if do_beautify () then
- pr_new_syntax (Loc.make_loc (max_int,max_int)) None
- | _ -> raise_with_file fname (disable_drop e, info)
+ if !beautify then
+ pr_new_syntax (fst input) chan_beautify (Loc.make_loc (max_int,max_int)) None;
+ if !Flags.beautify_file then close_out chan_beautify;
+ | reraise ->
+ if !Flags.beautify_file then close_out chan_beautify;
+ iraise (disable_drop e, info)
(** [eval_expr : ?preserving:bool -> Loc.t * Vernacexpr.vernac_expr -> unit]
It executes one vernacular command. By default the command is
@@ -271,34 +239,50 @@ and read_vernac_file verbosely s =
of a new state label). An example of state-preserving command is one coming
from the query panel of Coqide. *)
-let checknav loc ast =
- if is_deep_navigation_vernac ast then
- user_error loc "Navigation commands forbidden in nested commands"
-
-let eval_expr loc_ast = vernac_com (Flags.is_verbose()) checknav loc_ast
+let process_expr po loc_ast = interp_vernac po stdout checknav_deep loc_ast
(* XML output hooks *)
let (f_xml_start_library, xml_start_library) = Hook.make ~default:ignore ()
let (f_xml_end_library, xml_end_library) = Hook.make ~default:ignore ()
-(* Load a vernac file. Errors are annotated with file and location *)
-let load_vernac verb file =
- chan_beautify :=
- if !Flags.beautify_file then open_out (file^beautify_suffix) else stdout;
- try
- Flags.silently (read_vernac_file verb) file;
- if !Flags.beautify_file then close_out !chan_beautify;
- with any ->
- let (e, info) = Errors.push any in
- if !Flags.beautify_file then close_out !chan_beautify;
- raise_with_file file (disable_drop e, info)
+let warn_file_no_extension =
+ CWarnings.create ~name:"file-no-extension" ~category:"filesystem"
+ (fun (f,ext) ->
+ str "File \"" ++ str f ++
+ strbrk "\" has been implicitly expanded to \"" ++
+ str f ++ str ext ++ str "\"")
-let ensure_v f =
- if Filename.check_suffix f ".v" then f
+let ensure_ext ext f =
+ if Filename.check_suffix f ext then f
else begin
- msg_warning (str "File \"" ++ str f ++ strbrk "\" has been implicitly \
- expanded to \"" ++ str f ++ str ".v\"");
- f ^ ".v"
+ warn_file_no_extension (f,ext);
+ f ^ ext
+ end
+
+let chop_extension f =
+ try Filename.chop_extension f with _ -> f
+
+let ensure_bname src tgt =
+ let src, tgt = Filename.basename src, Filename.basename tgt in
+ let src, tgt = chop_extension src, chop_extension tgt in
+ if src <> tgt then begin
+ Feedback.msg_error (str "Source and target file names must coincide, directories can differ");
+ Feedback.msg_error (str "Source: " ++ str src);
+ Feedback.msg_error (str "Target: " ++ str tgt);
+ flush_all ();
+ exit 1
+ end
+
+let ensure ext src tgt = ensure_bname src tgt; ensure_ext ext tgt
+
+let ensure_v v = ensure ".v" v v
+let ensure_vo v vo = ensure ".vo" v vo
+let ensure_vio v vio = ensure ".vio" v vio
+
+let ensure_exists f =
+ if not (Sys.file_exists f) then begin
+ Feedback.msg_error (hov 0 (str "Can't find file" ++ spc () ++ str f));
+ exit 1
end
(* Compile a vernac file *)
@@ -306,14 +290,21 @@ let compile verbosely f =
let check_pending_proofs () =
let pfs = Pfedit.get_all_proof_names () in
if not (List.is_empty pfs) then
- (msg_error (str "There are pending proofs"); flush_all (); exit 1) in
+ (Feedback.msg_error (str "There are pending proofs"); flush_all (); exit 1) in
match !Flags.compilation_mode with
| BuildVo ->
let long_f_dot_v = ensure_v f in
- let ldir = Flags.verbosely Library.start_library long_f_dot_v in
- Stm.set_compilation_hints long_f_dot_v;
- Aux_file.start_aux_file_for long_f_dot_v;
- Dumpglob.start_dump_glob long_f_dot_v;
+ ensure_exists long_f_dot_v;
+ let long_f_dot_vo =
+ match !Flags.compilation_output_name with
+ | None -> long_f_dot_v ^ "o"
+ | Some f -> ensure_vo long_f_dot_v f in
+ let ldir = Flags.verbosely Library.start_library long_f_dot_vo in
+ Stm.set_compilation_hints long_f_dot_vo;
+ Aux_file.(start_aux_file
+ ~aux_file:(aux_file_name_for long_f_dot_vo)
+ ~v_file:long_f_dot_v);
+ Dumpglob.start_dump_glob ~vfile:long_f_dot_v ~vofile:long_f_dot_vo;
Dumpglob.dump_string ("F" ^ Names.DirPath.to_string ldir ^ "\n");
if !Flags.xml_export then Hook.get f_xml_start_library ();
let wall_clock1 = Unix.gettimeofday () in
@@ -321,7 +312,7 @@ let compile verbosely f =
Stm.join ();
let wall_clock2 = Unix.gettimeofday () in
check_pending_proofs ();
- Library.save_library_to ldir long_f_dot_v (Global.opaque_tables ());
+ Library.save_library_to ldir long_f_dot_vo (Global.opaque_tables ());
Aux_file.record_in_aux_at Loc.ghost "vo_compile_time"
(Printf.sprintf "%.3f" (wall_clock2 -. wall_clock1));
Aux_file.stop_aux_file ();
@@ -329,13 +320,18 @@ let compile verbosely f =
Dumpglob.end_dump_glob ()
| BuildVio ->
let long_f_dot_v = ensure_v f in
- let ldir = Flags.verbosely Library.start_library long_f_dot_v in
+ ensure_exists long_f_dot_v;
+ let long_f_dot_vio =
+ match !Flags.compilation_output_name with
+ | None -> long_f_dot_v ^ "io"
+ | Some f -> ensure_vio long_f_dot_v f in
+ let ldir = Flags.verbosely Library.start_library long_f_dot_vio in
Dumpglob.noglob ();
- Stm.set_compilation_hints long_f_dot_v;
+ Stm.set_compilation_hints long_f_dot_vio;
let _ = load_vernac verbosely long_f_dot_v in
Stm.finish ();
check_pending_proofs ();
- Stm.snapshot_vio ldir long_f_dot_v;
+ Stm.snapshot_vio ldir long_f_dot_vio;
Stm.reset_task_queue ()
| Vio2Vo ->
let open Filename in
@@ -352,8 +348,5 @@ let compile v f =
compile v f;
CoqworkmgrApi.giveback 1
-let () = Hook.set Stm.process_error_hook (fun e ->
- match !cur_file with
- | None -> Cerrors.process_vernac_interp_error e
- | Some f -> enrich_with_file f (Cerrors.process_vernac_interp_error e)
-)
+let () = Hook.set Stm.process_error_hook
+ ExplainErr.process_vernac_interp_error
diff --git a/toplevel/vernac.mli b/toplevel/vernac.mli
index 008d7a31..0d9f5871 100644
--- a/toplevel/vernac.mli
+++ b/toplevel/vernac.mli
@@ -11,17 +11,14 @@
(** Read a vernac command on the specified input (parse only).
Raises [End_of_file] if EOF (or Ctrl-D) is reached. *)
-val parse_sentence : Pcoq.Gram.parsable * in_channel option ->
+val parse_sentence : Pcoq.Gram.coq_parsable * in_channel option ->
Loc.t * Vernacexpr.vernac_expr
-(** Reads and executes vernac commands from a stream.
- The boolean [just_parsing] disables interpretation of commands. *)
+(** Reads and executes vernac commands from a stream. *)
exception End_of_input
-val just_parsing : bool ref
-
-val eval_expr : Loc.t * Vernacexpr.vernac_expr -> unit
+val process_expr : Pcoq.Gram.coq_parsable -> Loc.t * Vernacexpr.vernac_expr -> unit
(** Set XML hooks *)
val xml_start_library : (unit -> unit) Hook.t
@@ -38,9 +35,3 @@ val load_vernac : bool -> string -> unit
val compile : bool -> string -> unit
val is_navigation_vernac : Vernacexpr.vernac_expr -> bool
-
-(** Has an exception been annotated with some file locations ? *)
-
-type location_files = { outer : string; inner : string }
-
-val get_exn_files : Exninfo.info -> location_files option
diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml
index 72dd967b..41ee165f 100644
--- a/toplevel/vernacentries.ml
+++ b/toplevel/vernacentries.ml
@@ -9,7 +9,7 @@
(* Concrete syntax of the mathematical vernacular MV V2.6 *)
open Pp
-open Errors
+open CErrors
open Util
open Flags
open Names
@@ -20,7 +20,6 @@ open Tacmach
open Constrintern
open Prettyp
open Printer
-open Tacinterp
open Command
open Goptions
open Libnames
@@ -32,10 +31,14 @@ open Redexpr
open Lemmas
open Misctypes
open Locality
+open Sigma.Notations
+
+(** TODO: make this function independent of Ltac *)
+let (f_interp_redexp, interp_redexp_hook) = Hook.make ()
let debug = false
-let prerr_endline =
- if debug then prerr_endline else fun _ -> ()
+let prerr_endline x =
+ if debug then prerr_endline (x ()) else ()
(* Misc *)
@@ -45,7 +48,7 @@ let cl_of_qualid = function
| RefClass r -> Class.class_of_global (Smartlocate.smart_global ~head:true r)
let scope_class_of_qualid qid =
- Notation.scope_class_of_reference (Smartlocate.smart_global qid)
+ Notation.scope_class_of_class (cl_of_qualid qid)
(*******************)
(* "Show" commands *)
@@ -54,7 +57,7 @@ let show_proof () =
(* spiwack: this would probably be cooler with a bit of polishing. *)
let p = Proof_global.give_me_the_proof () in
let pprf = Proof.partial_proof p in
- msg_notice (Pp.prlist_with_sep Pp.fnl Printer.pr_constr pprf)
+ Feedback.msg_notice (Pp.prlist_with_sep Pp.fnl Printer.pr_constr pprf)
let show_node () =
(* spiwack: I'm have little clue what this function used to do. I deactivated it,
@@ -62,22 +65,22 @@ let show_node () =
()
let show_thesis () =
- msg_error (anomaly (Pp.str "TODO") )
+ Feedback.msg_error (anomaly (Pp.str "TODO") )
let show_top_evars () =
(* spiwack: new as of Feb. 2010: shows goal evars in addition to non-goal evars. *)
let pfts = get_pftreestate () in
let gls = Proof.V82.subgoals pfts in
let sigma = gls.Evd.sigma in
- msg_notice (pr_evars_int sigma 1 (Evarutil.non_instantiated sigma))
+ Feedback.msg_notice (pr_evars_int sigma 1 (Evarutil.non_instantiated sigma))
let show_universes () =
let pfts = get_pftreestate () in
let gls = Proof.V82.subgoals pfts in
let sigma = gls.Evd.sigma in
let ctx = Evd.universe_context_set (Evd.nf_constraints sigma) in
- msg_notice (Evd.pr_evar_universe_context (Evd.evar_universe_context sigma));
- msg_notice (str"Normalized constraints: " ++ Univ.pr_universe_context_set (Evd.pr_evd_level sigma) ctx)
+ Feedback.msg_notice (Evd.pr_evar_universe_context (Evd.evar_universe_context sigma));
+ Feedback.msg_notice (str"Normalized constraints: " ++ Univ.pr_universe_context_set (Evd.pr_evd_level sigma) ctx)
let show_prooftree () =
(* Spiwack: proof tree is currently not working *)
@@ -88,11 +91,10 @@ let enable_goal_printing = ref true
let print_subgoals () =
if !enable_goal_printing && is_verbose ()
then begin
- msg_notice (pr_open_subgoals ())
+ Feedback.msg_notice (pr_open_subgoals ())
end
let try_print_subgoals () =
- Pp.flush_all();
try print_subgoals () with Proof_global.NoCurrentProof | UserError _ -> ()
@@ -106,10 +108,10 @@ let show_intro all =
let l,_= decompose_prod_assum (strip_outer_cast (pf_concl gl)) in
if all then
let lid = Tactics.find_intro_names l gl in
- msg_notice (hov 0 (prlist_with_sep spc pr_id lid))
+ Feedback.msg_notice (hov 0 (prlist_with_sep spc pr_id lid))
else if not (List.is_empty l) then
let n = List.last l in
- msg_notice (pr_id (List.hd (Tactics.find_intro_names [n] gl)))
+ Feedback.msg_notice (pr_id (List.hd (Tactics.find_intro_names [n] gl)))
end
(** Prepare a "match" template for a given inductive type.
@@ -118,9 +120,7 @@ let show_intro all =
[Not_found] is raised if the given string isn't the qualid of
a known inductive type. *)
-let make_cases s =
- let qualified_name = Libnames.qualid_of_string s in
- let glob_ref = Nametab.locate qualified_name in
+let make_cases_aux glob_ref =
match glob_ref with
| Globnames.IndRef i ->
let {Declarations.mind_nparams = np}
@@ -133,30 +133,35 @@ let make_cases s =
let rec rename avoid = function
| [] -> []
| (n,_)::l ->
- let n' = Namegen.next_name_away_in_cases_pattern ([],mkMeta 0) n avoid in
+ let n' = Namegen.next_name_away_with_default (Id.to_string Namegen.default_dependent_ident) n avoid in
Id.to_string n' :: rename (n'::avoid) l in
let al' = rename [] al in
(Id.to_string consname :: al') :: l)
carr tarr []
| _ -> raise Not_found
+let make_cases s =
+ let qualified_name = Libnames.qualid_of_string s in
+ let glob_ref = Nametab.locate qualified_name in
+ make_cases_aux glob_ref
+
(** Textual display of a generic "match" template *)
let show_match id =
let patterns =
- try make_cases (Id.to_string (snd id))
+ try make_cases_aux (Nametab.global id)
with Not_found -> error "Unknown inductive type."
in
let pr_branch l =
str "| " ++ hov 1 (prlist_with_sep spc str l) ++ str " =>"
in
- msg_notice (v 1 (str "match # with" ++ fnl () ++
+ Feedback.msg_notice (v 1 (str "match # with" ++ fnl () ++
prlist_with_sep fnl pr_branch patterns ++ fnl () ++ str "end" ++ fnl ()))
(* "Print" commands *)
let print_path_entry p =
- let dir = str (DirPath.to_string (Loadpath.logical p)) in
+ let dir = pr_dirpath (Loadpath.logical p) in
let path = str (Loadpath.physical p) in
(dir ++ str " " ++ tbrk (0, 0) ++ path)
@@ -191,23 +196,23 @@ let print_module r =
let globdir = Nametab.locate_dir qid in
match globdir with
DirModule (dirpath,(mp,_)) ->
- msg_notice (Printmod.print_module (Printmod.printable_body dirpath) mp)
+ Feedback.msg_notice (Printmod.print_module (Printmod.printable_body dirpath) mp)
| _ -> raise Not_found
with
- Not_found -> msg_error (str"Unknown Module " ++ pr_qualid qid)
+ Not_found -> Feedback.msg_error (str"Unknown Module " ++ pr_qualid qid)
let print_modtype r =
let (loc,qid) = qualid_of_reference r in
try
let kn = Nametab.locate_modtype qid in
- msg_notice (Printmod.print_modtype kn)
+ Feedback.msg_notice (Printmod.print_modtype kn)
with Not_found ->
(* Is there a module of this name ? If yes we display its type *)
try
let mp = Nametab.locate_module qid in
- msg_notice (Printmod.print_module false mp)
+ Feedback.msg_notice (Printmod.print_module false mp)
with Not_found ->
- msg_error (str"Unknown Module Type or Module " ++ pr_qualid qid)
+ Feedback.msg_error (str"Unknown Module Type or Module " ++ pr_qualid qid)
let print_namespace ns =
let ns = List.rev (Names.DirPath.repr ns) in
@@ -276,7 +281,7 @@ let print_namespace ns =
acc
) constants (str"")
in
- msg_notice ((print_list pr_id ns)++str":"++fnl()++constants_in_namespace)
+ Feedback.msg_notice ((print_list pr_id ns)++str":"++fnl()++constants_in_namespace)
let print_strategy r =
let open Conv_oracle in
@@ -306,7 +311,7 @@ let print_strategy r =
else str "Constant strategies" ++ fnl () ++
hov 0 (prlist_with_sep fnl pr_strategy cst_lvl)
in
- msg_notice (var_msg ++ cst_msg)
+ Feedback.msg_notice (var_msg ++ cst_msg)
| Some r ->
let r = Smartlocate.smart_global r in
let key = match r with
@@ -315,7 +320,7 @@ let print_strategy r =
| IndRef _ | ConstructRef _ -> error "The reference is not unfoldable"
in
let lvl = get_strategy oracle key in
- msg_notice (pr_strategy (r, lvl))
+ Feedback.msg_notice (pr_strategy (r, lvl))
let dump_universes_gen g s =
let output = open_out s in
@@ -333,7 +338,7 @@ let dump_universes_gen g s =
| Univ.Eq ->
Printf.fprintf output " \"%s\" -> \"%s\" [style=dashed];\n" left right
end, begin fun () ->
- if Lazy.lazy_is_val init then Printf.fprintf output "}\n";
+ if Lazy.is_val init then Printf.fprintf output "}\n";
close_out output
end
end else begin
@@ -347,11 +352,11 @@ let dump_universes_gen g s =
end
in
try
- Univ.dump_universes output_constraint g;
+ UGraph.dump_universes output_constraint g;
close ();
- msg_info (str "Universes written to file \"" ++ str s ++ str "\".")
+ Feedback.msg_info (str "Universes written to file \"" ++ str s ++ str "\".")
with reraise ->
- let reraise = Errors.push reraise in
+ let reraise = CErrors.push reraise in
close ();
iraise reraise
@@ -364,11 +369,11 @@ let locate_file f =
let msg_found_library = function
| Library.LibLoaded, fulldir, file ->
- msg_info (hov 0
+ Feedback.msg_info (hov 0
(pr_dirpath fulldir ++ strbrk " has been loaded from file " ++
str file))
| Library.LibInPath, fulldir, file ->
- msg_info (hov 0
+ Feedback.msg_info (hov 0
(pr_dirpath fulldir ++ strbrk " is bound to file " ++ str file))
let err_unmapped_library loc ?from qid =
@@ -409,7 +414,7 @@ let dump_global r =
try
let gr = Smartlocate.smart_global r in
Dumpglob.add_glob (Constrarg.loc_of_or_by_notation loc_of_reference r) gr
- with e when Errors.noncritical e -> ()
+ with e when CErrors.noncritical e -> ()
(**********)
(* Syntax *)
@@ -443,7 +448,27 @@ let vernac_notation locality local =
(***********)
(* Gallina *)
-let start_proof_and_print k l hook = start_proof_com k l hook
+let start_proof_and_print k l hook =
+ let inference_hook =
+ if Flags.is_program_mode () then
+ let hook env sigma ev =
+ let tac = !Obligations.default_tactic in
+ let evi = Evd.find sigma ev in
+ let env = Evd.evar_filtered_env evi in
+ try
+ let concl = Evarutil.nf_evars_universes sigma evi.Evd.evar_concl in
+ if Evarutil.has_undefined_evars sigma concl then raise Exit;
+ let c, _, ctx =
+ Pfedit.build_by_tactic env (Evd.evar_universe_context sigma)
+ concl (Tacticals.New.tclCOMPLETE tac)
+ in Evd.set_universe_context sigma ctx, c
+ with Logic_monad.TacticFailure e when Logic.catchable_exception e ->
+ error "The statement obligations could not be resolved \
+ automatically, write a statement definition first."
+ in Some hook
+ else None
+ in
+ start_proof_com ?inference_hook k l hook
let no_hook = Lemmas.mk_hook (fun _ _ -> ())
@@ -463,14 +488,14 @@ let vernac_definition locality p (local,k) ((loc,id as lid),pl) def =
in
(match def with
| ProveBody (bl,t) -> (* local binders, typ *)
- start_proof_and_print (local,p,DefinitionBody Definition)
- [Some (lid,pl), (bl,t,None)] no_hook
+ start_proof_and_print (local,p,DefinitionBody k)
+ [Some (lid,pl), (bl,t,None)] hook
| DefineBody (bl,red_option,c,typ_opt) ->
let red_option = match red_option with
| None -> None
| Some r ->
let (evc,env)= get_current_context () in
- Some (snd (interp_redexp env evc r)) in
+ Some (snd (Hook.get f_interp_redexp env evc r)) in
do_definition id (local,p,k) pl bl red_option c typ_opt hook)
let vernac_start_proof locality p kind l lettop =
@@ -501,9 +526,9 @@ let vernac_end_proof ?proof = function
let vernac_exact_proof c =
(* spiwack: for simplicity I do not enforce that "Proof proof_term" is
called only at the begining of a proof. *)
- let status = by (Tactics.New.exact_proof c) in
+ let status = by (Tactics.exact_proof c) in
save_proof (Vernacexpr.(Proved(Opaque None,None)));
- if not status then Pp.feedback Feedback.AddedAxiom
+ if not status then Feedback.feedback Feedback.AddedAxiom
let vernac_assumption locality poly (local, kind) l nl =
let local = enforce_locality_exp locality local in
@@ -515,7 +540,7 @@ let vernac_assumption locality poly (local, kind) l nl =
if global then Dumpglob.dump_definition lid false "ax"
else Dumpglob.dump_definition lid true "var") idl) l;
let status = do_assumptions kind nl l in
- if not status then Pp.feedback Feedback.AddedAxiom
+ if not status then Feedback.feedback Feedback.AddedAxiom
let vernac_record k poly finite struc binders sort nameopt cfs =
let const = match nameopt with
@@ -530,6 +555,10 @@ let vernac_record k poly finite struc binders sort nameopt cfs =
| _ -> ()) cfs);
ignore(Record.definition_structure (k,poly,finite,struc,binders,cfs,const,sort))
+(** When [poly] is true the type is declared polymorphic. When [lo] is true,
+ then the type is declared private (as per the [Private] keyword). [finite]
+ indicates whether the type is inductive, co-inductive or
+ neither. *)
let vernac_inductive poly lo finite indl =
if Dumpglob.dump () then
List.iter (fun (((coe,(lid,_)), _, _, _, cstrs), _) ->
@@ -542,29 +571,29 @@ let vernac_inductive poly lo finite indl =
indl;
match indl with
| [ ( _ , _ , _ ,Record, Constructors _ ),_ ] ->
- Errors.error "The Record keyword cannot be used to define a variant type. Use Variant instead."
+ CErrors.error "The Record keyword cannot be used to define a variant type. Use Variant instead."
| [ (_ , _ , _ ,Variant, RecordDecl _),_ ] ->
- Errors.error "The Variant keyword cannot be used to define a record type. Use Record instead."
+ CErrors.error "The Variant keyword cannot be used to define a record type. Use Record instead."
| [ ( id , bl , c , b, RecordDecl (oc,fs) ), [] ] ->
- vernac_record (match b with Class true -> Class false | _ -> b)
+ vernac_record (match b with Class _ -> Class false | _ -> b)
poly finite id bl c oc fs
- | [ ( id , bl , c , Class true, Constructors [l]), _ ] ->
+ | [ ( id , bl , c , Class _, Constructors [l]), [] ] ->
let f =
let (coe, ((loc, id), ce)) = l in
let coe' = if coe then Some true else None in
(((coe', AssumExpr ((loc, Name id), ce)), None), [])
in vernac_record (Class true) poly finite id bl c None [f]
- | [ ( id , bl , c , Class true, _), _ ] ->
- Errors.error "Definitional classes must have a single method"
- | [ ( id , bl , c , Class false, Constructors _), _ ] ->
- Errors.error "Inductive classes not supported"
+ | [ ( _ , _, _, Class _, Constructors _), [] ] ->
+ CErrors.error "Inductive classes not supported"
+ | [ ( id , bl , c , Class _, _), _ :: _ ] ->
+ CErrors.error "where clause not supported for classes"
| [ ( _ , _ , _ , _, RecordDecl _ ) , _ ] ->
- Errors.error "where clause not supported for (co)inductive records"
+ CErrors.error "where clause not supported for (co)inductive records"
| _ -> let unpack = function
| ( (false, id) , bl , c , _ , Constructors l ) , ntn -> ( id , bl , c , l ) , ntn
| ( (true,_),_,_,_,Constructors _),_ ->
- Errors.error "Variant types do not handle the \"> Name\" syntax, which is reserved for records. Use the \":>\" syntax on constructors instead."
- | _ -> Errors.error "Cannot handle mutually (co)inductive records."
+ CErrors.error "Variant types do not handle the \"> Name\" syntax, which is reserved for records. Use the \":>\" syntax on constructors instead."
+ | _ -> CErrors.error "Cannot handle mutually (co)inductive records."
in
let indl = List.map unpack indl in
do_mutual_inductive indl poly lo finite
@@ -633,7 +662,7 @@ let vernac_declare_module export (loc, id) binders_ast mty_ast =
id binders_ast (Enforce mty_ast) []
in
Dumpglob.dump_moddef loc mp "mod";
- if_verbose msg_info (str "Module " ++ pr_id id ++ str " is declared");
+ if_verbose Feedback.msg_info (str "Module " ++ pr_id id ++ str " is declared");
Option.iter (fun export -> vernac_import export [Ident (Loc.ghost,id)]) export
let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l =
@@ -654,7 +683,7 @@ let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l =
export id binders_ast mty_ast_o
in
Dumpglob.dump_moddef loc mp "mod";
- if_verbose msg_info
+ if_verbose Feedback.msg_info
(str "Interactive Module " ++ pr_id id ++ str " started");
List.iter
(fun (export,id) ->
@@ -672,7 +701,7 @@ let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l =
id binders_ast mty_ast_o mexpr_ast_l
in
Dumpglob.dump_moddef loc mp "mod";
- if_verbose msg_info
+ if_verbose Feedback.msg_info
(str "Module " ++ pr_id id ++ str " is defined");
Option.iter (fun export -> vernac_import export [Ident (Loc.ghost,id)])
export
@@ -680,7 +709,7 @@ let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l =
let vernac_end_module export (loc,id as lid) =
let mp = Declaremods.end_module () in
Dumpglob.dump_modref loc mp "mod";
- if_verbose msg_info (str "Module " ++ pr_id id ++ str " is defined");
+ if_verbose Feedback.msg_info (str "Module " ++ pr_id id ++ str " is defined");
Option.iter (fun export -> vernac_import export [Ident lid]) export
let vernac_declare_module_type (loc,id) binders_ast mty_sign mty_ast_l =
@@ -701,7 +730,7 @@ let vernac_declare_module_type (loc,id) binders_ast mty_sign mty_ast_l =
id binders_ast mty_sign
in
Dumpglob.dump_moddef loc mp "modtype";
- if_verbose msg_info
+ if_verbose Feedback.msg_info
(str "Interactive Module Type " ++ pr_id id ++ str " started");
List.iter
(fun (export,id) ->
@@ -720,13 +749,13 @@ let vernac_declare_module_type (loc,id) binders_ast mty_sign mty_ast_l =
id binders_ast mty_sign mty_ast_l
in
Dumpglob.dump_moddef loc mp "modtype";
- if_verbose msg_info
+ if_verbose Feedback.msg_info
(str "Module Type " ++ pr_id id ++ str " is defined")
let vernac_end_modtype (loc,id) =
let mp = Declaremods.end_modtype () in
Dumpglob.dump_modref loc mp "modtype";
- if_verbose msg_info (str "Module Type " ++ pr_id id ++ str " is defined")
+ if_verbose Feedback.msg_info (str "Module Type " ++ pr_id id ++ str " is defined")
let vernac_include l =
Declaremods.declare_include Modintern.interp_module_ast l
@@ -794,7 +823,7 @@ let vernac_coercion locality poly local ref qids qidt =
let source = cl_of_qualid qids in
let ref' = smart_global ref in
Class.try_add_new_coercion_with_target ref' ~local poly ~source ~target;
- if_verbose msg_info (pr_global ref' ++ str " is now a coercion")
+ if_verbose Feedback.msg_info (pr_global ref' ++ str " is now a coercion")
let vernac_identity_coercion locality poly local id qids qidt =
let local = enforce_locality locality local in
@@ -810,11 +839,11 @@ let vernac_instance abst locality poly sup inst props pri =
ignore(Classes.new_instance ~abstract:abst ~global poly sup inst props pri)
let vernac_context poly l =
- if not (Classes.context poly l) then Pp.feedback Feedback.AddedAxiom
+ if not (Classes.context poly l) then Feedback.feedback Feedback.AddedAxiom
-let vernac_declare_instances locality ids pri =
+let vernac_declare_instances locality insts =
let glob = not (make_section_locality locality) in
- List.iter (fun id -> Classes.existing_instance glob id pri) ids
+ List.iter (fun (id, info) -> Classes.existing_instance glob id (Some info)) insts
let vernac_declare_class id =
Record.declare_existing_class (Nametab.global id)
@@ -825,35 +854,6 @@ let vernac_declare_class id =
let command_focus = Proof.new_focus_kind ()
let focus_command_cond = Proof.no_cond command_focus
-
-let print_info_trace = ref None
-
-let _ = let open Goptions in declare_int_option {
- optsync = true;
- optdepr = false;
- optname = "print info trace";
- optkey = ["Info" ; "Level"];
- optread = (fun () -> !print_info_trace);
- optwrite = fun n -> print_info_trace := n;
-}
-
-let vernac_solve n info tcom b =
- if not (refining ()) then
- error "Unknown command of the non proof-editing mode.";
- let status = Proof_global.with_current_proof (fun etac p ->
- let with_end_tac = if b then Some etac else None in
- let global = match n with SelectAll -> true | _ -> false in
- let info = Option.append info !print_info_trace in
- let (p,status) =
- solve n info (Tacinterp.hide_interp global tcom None) ?with_end_tac p
- in
- (* in case a strict subtree was completed,
- go back to the top of the prooftree *)
- let p = Proof.maximal_unfocus command_focus p in
- p,status) in
- if not status then Pp.feedback Feedback.AddedAxiom
-
-
(* A command which should be a tactic. It has been
added by Christine to patch an error in the design of the proof
machine, and enables to instantiate existential variables when
@@ -871,31 +871,36 @@ let vernac_set_end_tac tac =
(* TO DO verifier s'il faut pas mettre exist s | TacId s ici*)
let vernac_set_used_variables e =
+ let open Context.Named.Declaration in
let env = Global.env () in
let tys =
List.map snd (Proof.initial_goals (Proof_global.give_me_the_proof ())) in
let l = Proof_using.process_expr env e tys in
let vars = Environ.named_context env in
List.iter (fun id ->
- if not (List.exists (fun (id',_,_) -> Id.equal id id') vars) then
+ if not (List.exists (Id.equal id % get_id) vars) then
errorlabstrm "vernac_set_used_variables"
(str "Unknown variable: " ++ pr_id id))
l;
let _, to_clear = set_used_variables l in
- vernac_solve
- SelectAll None Tacexpr.(TacAtom (Loc.ghost,TacClear(false,to_clear))) false
-
+ let to_clear = List.map snd to_clear in
+ Proof_global.with_current_proof begin fun _ p ->
+ if List.is_empty to_clear then (p, ())
+ else
+ let tac = Tactics.clear to_clear in
+ fst (solve SelectAll None tac p), ()
+ end
(*****************************)
(* Auxiliary file management *)
let expand filename =
- Envars.expand_path_macros ~warn:(fun x -> msg_warning (str x)) filename
+ Envars.expand_path_macros ~warn:(fun x -> Feedback.msg_warning (str x)) filename
let vernac_add_loadpath implicit pdir ldiropt =
let pdir = expand pdir in
let alias = Option.default Nameops.default_root_prefix ldiropt in
- Mltop.add_rec_path ~unix_path:pdir ~coq_root:alias ~implicit
+ Mltop.add_rec_path Mltop.AddTopML ~unix_path:pdir ~coq_root:alias ~implicit
let vernac_remove_loadpath path =
Loadpath.remove_load_path (expand path)
@@ -910,13 +915,17 @@ let vernac_declare_ml_module locality l =
Mltop.declare_ml_modules local (List.map expand l)
let vernac_chdir = function
- | None -> msg_notice (str (Sys.getcwd()))
+ | None -> Feedback.msg_notice (str (Sys.getcwd()))
| Some path ->
begin
try Sys.chdir (expand path)
- with Sys_error err -> msg_warning (str "Cd failed: " ++ str err)
+ with Sys_error err ->
+ (* Cd is typically used to control the output directory of
+ extraction. A failed Cd could lead to overwriting .ml files
+ so we make it an error. *)
+ CErrors.error ("Cd failed: " ^ err)
end;
- if_verbose msg_info (str (Sys.getcwd()))
+ if_verbose Feedback.msg_info (str (Sys.getcwd()))
(********************)
@@ -935,85 +944,6 @@ let vernac_restore_state file =
(************)
(* Commands *)
-type tacdef_kind =
- | NewTac of Id.t
- | UpdateTac of Nametab.ltac_constant
-
-let is_defined_tac kn =
- try ignore (Tacenv.interp_ltac kn); true with Not_found -> false
-
-let make_absolute_name ident repl =
- let loc = loc_of_reference ident in
- if repl then
- let kn =
- try Nametab.locate_tactic (snd (qualid_of_reference ident))
- with Not_found ->
- Errors.user_err_loc (loc, "",
- str "There is no Ltac named " ++ pr_reference ident ++ str ".")
- in
- UpdateTac kn
- else
- let id = Constrexpr_ops.coerce_reference_to_id ident in
- let kn = Lib.make_kn id in
- let () = if is_defined_tac kn then
- Errors.user_err_loc (loc, "",
- str "There is already an Ltac named " ++ pr_reference ident ++ str".")
- in
- let is_primitive =
- try
- match Pcoq.parse_string Pcoq.Tactic.tactic (Id.to_string id) with
- | Tacexpr.TacArg _ -> false
- | _ -> true (* most probably TacAtom, i.e. a primitive tactic ident *)
- with e when Errors.noncritical e -> true (* prim tactics with args, e.g. "apply" *)
- in
- let () = if is_primitive then
- msg_warning (str "The Ltac name " ++ pr_reference ident ++
- str " may be unusable because of a conflict with a notation.")
- in
- NewTac id
-
-let register_ltac local isrec tacl =
- let map (ident, repl, body) =
- let name = make_absolute_name ident repl in
- (name, body)
- in
- let rfun = List.map map tacl in
- let recvars =
- let fold accu (op, _) = match op with
- | UpdateTac _ -> accu
- | NewTac id -> (Lib.make_path id, Lib.make_kn id) :: accu
- in
- if isrec then List.fold_left fold [] rfun
- else []
- in
- let ist = Tacintern.make_empty_glob_sign () in
- let map (name, body) =
- let body = Flags.with_option Tacintern.strict_check (Tacintern.intern_tactic_or_tacarg ist) body in
- (name, body)
- in
- let defs () =
- (** Register locally the tactic to handle recursivity. This function affects
- the whole environment, so that we transactify it afterwards. *)
- let iter_rec (sp, kn) = Nametab.push_tactic (Nametab.Until 1) sp kn in
- let () = List.iter iter_rec recvars in
- List.map map rfun
- in
- let defs = Future.transactify defs () in
- let iter (def, tac) = match def with
- | NewTac id ->
- Tacenv.register_ltac false local id tac;
- Flags.if_verbose msg_info (Nameops.pr_id id ++ str " is defined")
- | UpdateTac kn ->
- Tacenv.redefine_ltac local kn tac;
- let name = Nametab.shortest_qualid_of_tactic kn in
- Flags.if_verbose msg_info (Libnames.pr_qualid name ++ str " is redefined")
- in
- List.iter iter defs
-
-let vernac_declare_tactic_definition locality (x,def) =
- let local = make_module_locality locality in
- register_ltac local x def
-
let vernac_create_hintdb locality id b =
let local = make_module_locality locality in
Hints.create_hint_db local id full_transparent_state b
@@ -1040,141 +970,265 @@ let vernac_declare_implicits locality r l =
Impargs.declare_manual_implicits local (smart_global r) ~enriching:false
(List.map (List.map (fun (ex,b,f) -> ex, (b,true,f))) imps)
-let vernac_declare_arguments locality r l nargs flags =
- let extra_scope_flag = List.mem `ExtraScopes flags in
- let names = List.map (List.map (fun (id, _,_,_,_) -> id)) l in
- let names, rest = List.hd names, List.tl names in
- let scopes = List.map (List.map (fun (_,_, s, _,_) -> s)) l in
- if List.exists (fun na -> not (List.equal Name.equal na names)) rest then
- error "All arguments lists must declare the same names.";
- if not (List.distinct_f Name.compare (List.filter ((!=) Anonymous) names))
- then error "Arguments names must be distinct.";
- let sr = smart_global r in
+let warn_arguments_assert =
+ CWarnings.create ~name:"arguments-assert" ~category:"vernacular"
+ (fun sr ->
+ strbrk "This command is just asserting the names of arguments of " ++
+ pr_global sr ++ strbrk". If this is what you want add " ++
+ strbrk "': assert' to silence the warning. If you want " ++
+ strbrk "to clear implicit arguments add ': clear implicits'. " ++
+ strbrk "If you want to clear notation scopes add ': clear scopes'")
+
+(* [nargs_for_red] is the number of arguments required to trigger reduction,
+ [args] is the main list of arguments statuses,
+ [more_implicits] is a list of extra lists of implicit statuses *)
+let vernac_arguments locality reference args more_implicits nargs_for_red flags =
+ let assert_flag = List.mem `Assert flags in
+ let rename_flag = List.mem `Rename flags in
+ let clear_scopes_flag = List.mem `ClearScopes flags in
+ let extra_scopes_flag = List.mem `ExtraScopes flags in
+ let clear_implicits_flag = List.mem `ClearImplicits flags in
+ let default_implicits_flag = List.mem `DefaultImplicits flags in
+ let never_unfold_flag = List.mem `ReductionNeverUnfold flags in
+
+ let err_incompat x y =
+ error ("Options \""^x^"\" and \""^y^"\" are incompatible.") in
+
+ if assert_flag && rename_flag then
+ err_incompat "assert" "rename";
+ if Option.has_some nargs_for_red && never_unfold_flag then
+ err_incompat "simpl never" "/";
+ if never_unfold_flag && List.mem `ReductionDontExposeCase flags then
+ err_incompat "simpl never" "simpl nomatch";
+ if clear_scopes_flag && extra_scopes_flag then
+ err_incompat "clear scopes" "extra scopes";
+ if clear_implicits_flag && default_implicits_flag then
+ err_incompat "clear implicits" "default implicits";
+
+ let sr = smart_global reference in
let inf_names =
let ty = Global.type_of_global_unsafe sr in
- Impargs.compute_implicits_names (Global.env ()) ty in
- let rec check li ld ls = match li, ld, ls with
- | [], [], [] -> ()
- | [], Anonymous::ld, (Some _)::ls when extra_scope_flag -> check li ld ls
- | [], _::_, (Some _)::ls when extra_scope_flag ->
- error "Extra notation scopes can be set on anonymous arguments only"
- | [], x::_, _ -> errorlabstrm "vernac_declare_arguments"
- (str "Extra argument " ++ pr_name x ++ str ".")
- | l, [], _ -> errorlabstrm "vernac_declare_arguments"
- (str "The following arguments are not declared: " ++
- prlist_with_sep pr_comma pr_name l ++ str ".")
- | _::li, _::ld, _::ls -> check li ld ls
- | _ -> assert false in
- let () = match l with
- | [[]] -> ()
- | _ ->
- List.iter2 (fun l -> check inf_names l) (names :: rest) scopes
+ Impargs.compute_implicits_names (Global.env ()) ty
in
- (* we take extra scopes apart, and we check they are consistent *)
- let l, scopes =
- let scopes, rest = List.hd scopes, List.tl scopes in
- if List.exists (List.exists ((!=) None)) rest then
- error "Notation scopes can be given only once";
- if not extra_scope_flag then l, scopes else
- let l, _ = List.split (List.map (List.chop (List.length inf_names)) l) in
- l, scopes in
- (* we interpret _ as the inferred names *)
- let l = match l with
- | [[]] -> l
- | _ ->
- let name_anons = function
- | (Anonymous, a,b,c,d), Name id -> Name id, a,b,c,d
- | x, _ -> x in
- List.map (fun ns -> List.map name_anons (List.combine ns inf_names)) l in
- let names_decl = List.map (List.map (fun (id, _,_,_,_) -> id)) l in
- let renamed_arg = ref None in
- let set_renamed a b =
- if Option.is_empty !renamed_arg && not (Id.equal a b) then renamed_arg := Some(b,a) in
- let some_renaming_specified =
- try
- let names = Arguments_renaming.arguments_names sr in
- not (List.equal (List.equal Name.equal) names names_decl)
- with Not_found -> false in
- let some_renaming_specified, implicits =
- match l with
- | [[]] -> false, [[]]
+ let prev_names =
+ try Arguments_renaming.arguments_names sr with Not_found -> inf_names
+ in
+ let num_args = List.length inf_names in
+ assert (Int.equal num_args (List.length prev_names));
+
+ let names_of args = List.map (fun a -> a.name) args in
+
+ (* Checks *)
+
+ let err_extra_args names =
+ errorlabstrm "vernac_declare_arguments"
+ (strbrk "Extra arguments: " ++
+ prlist_with_sep pr_comma pr_name names ++ str ".")
+ in
+ let err_missing_args names =
+ errorlabstrm "vernac_declare_arguments"
+ (strbrk "The following arguments are not declared: " ++
+ prlist_with_sep pr_comma pr_name names ++ str ".")
+ in
+
+ let rec check_extra_args extra_args =
+ match extra_args with
+ | [] -> ()
+ | { notation_scope = None } :: _ -> err_extra_args (names_of extra_args)
+ | { name = Anonymous; notation_scope = Some _ } :: args ->
+ check_extra_args args
| _ ->
- List.fold_map (fun sr il ->
- let sr', impl = List.fold_map (fun b -> function
- | (Anonymous, _,_, true, max), Name id -> assert false
- | (Name x, _,_, true, _), Anonymous ->
- errorlabstrm "vernac_declare_arguments"
- (str "Argument " ++ pr_id x ++ str " cannot be declared implicit.")
- | (Name iid, _,_, true, max), Name id ->
- set_renamed iid id;
- b || not (Id.equal iid id), Some (ExplByName id, max, false)
- | (Name iid, _,_, _, _), Name id ->
- set_renamed iid id;
- b || not (Id.equal iid id), None
- | _ -> b, None)
- sr (List.combine il inf_names) in
- sr || sr', List.map_filter (fun x -> x) impl)
- some_renaming_specified l in
- if some_renaming_specified then
- if not (List.mem `Rename flags) then
- errorlabstrm "vernac_declare_arguments"
- (str "To rename arguments the \"rename\" flag must be specified." ++
- match !renamed_arg with
- | None -> mt ()
- | Some (o,n) ->
- str "\nArgument " ++ pr_id o ++ str " renamed to " ++ pr_id n ++ str ".")
- else
- Arguments_renaming.rename_arguments
- (make_section_locality locality) sr names_decl;
- (* All other infos are in the first item of l *)
- let l = List.hd l in
- let some_implicits_specified = match implicits with
- | [[]] -> false | _ -> true in
- let scopes = List.map (function
- | None -> None
- | Some (o, k) ->
- try ignore (Notation.find_scope k); Some k
- with UserError _ ->
- Some (Notation.find_delimiters_scope o k)) scopes
+ error "Extra notation scopes can be set on anonymous and explicit arguments only."
+ in
+
+ let args, scopes =
+ let scopes = List.map (fun { notation_scope = s } -> s) args in
+ if List.length args > num_args then
+ let args, extra_args = List.chop num_args args in
+ if extra_scopes_flag then
+ (check_extra_args extra_args; (args, scopes))
+ else err_extra_args (names_of extra_args)
+ else args, scopes
+ in
+
+ if Option.cata (fun n -> n > num_args) false nargs_for_red then
+ error "The \"/\" modifier should be put before any extra scope.";
+
+ let scopes_specified = List.exists Option.has_some scopes in
+
+ if scopes_specified && clear_scopes_flag then
+ error "The \"clear scopes\" flag is incompatible with scope annotations.";
+
+ let names = List.map (fun { name } -> name) args in
+ let names = names :: List.map (List.map fst) more_implicits in
+
+ let rename_flag_required = ref false in
+ let example_renaming = ref None in
+ let save_example_renaming renaming =
+ rename_flag_required := !rename_flag_required
+ || not (Name.equal (fst renaming) Anonymous);
+ if Option.is_empty !example_renaming then
+ example_renaming := Some renaming
in
- let some_scopes_specified = List.exists ((!=) None) scopes in
+
+ let rec names_union names1 names2 =
+ match names1, names2 with
+ | [], [] -> []
+ | _ :: _, [] -> names1
+ | [], _ :: _ -> names2
+ | (Name _ as name) :: names1, Anonymous :: names2
+ | Anonymous :: names1, (Name _ as name) :: names2 ->
+ name :: names_union names1 names2
+ | name1 :: names1, name2 :: names2 ->
+ if Name.equal name1 name2 then
+ name1 :: names_union names1 names2
+ else error "Argument lists should agree on the names they provide."
+ in
+
+ let names = List.fold_left names_union [] names in
+
+ let rec rename prev_names names =
+ match prev_names, names with
+ | [], [] -> []
+ | [], _ :: _ -> err_extra_args names
+ | _ :: _, [] when assert_flag ->
+ (* Error messages are expressed in terms of original names, not
+ renamed ones. *)
+ err_missing_args (List.lastn (List.length prev_names) inf_names)
+ | _ :: _, [] -> prev_names
+ | prev :: prev_names, Anonymous :: names ->
+ prev :: rename prev_names names
+ | prev :: prev_names, (Name id as name) :: names ->
+ if not (Name.equal prev name) then save_example_renaming (prev,name);
+ name :: rename prev_names names
+ in
+
+ let names = rename prev_names names in
+ let renaming_specified = Option.has_some !example_renaming in
+
+ if !rename_flag_required && not rename_flag then
+ errorlabstrm "vernac_declare_arguments"
+ (strbrk "To rename arguments the \"rename\" flag must be specified."
+ ++ spc () ++
+ match !example_renaming with
+ | None -> mt ()
+ | Some (o,n) ->
+ str "Argument " ++ pr_name o ++
+ str " renamed to " ++ pr_name n ++ str ".");
+
+ let duplicate_names =
+ List.duplicates Name.equal (List.filter ((!=) Anonymous) names)
+ in
+ if not (List.is_empty duplicate_names) then begin
+ let duplicates = prlist_with_sep pr_comma pr_name duplicate_names in
+ errorlabstrm "_" (strbrk "Some argument names are duplicated: " ++ duplicates)
+ end;
+
+ (* Parts of this code are overly complicated because the implicit arguments
+ API is completely crazy: positions (ExplByPos) are elaborated to
+ names. This is broken by design, since not all arguments have names. So
+ even though we eventually want to map only positions to implicit statuses,
+ we have to check whether the corresponding arguments have names, not to
+ trigger an error in the impargs code. Even better, the names we have to
+ check are not the current ones (after previous renamings), but the original
+ ones (inferred from the type). *)
+
+ let implicits =
+ List.map (fun { name; implicit_status = i } -> (name,i)) args
+ in
+ let implicits = implicits :: more_implicits in
+
+ let open Vernacexpr in
+ let rec build_implicits inf_names implicits =
+ match inf_names, implicits with
+ | _, [] -> []
+ | _ :: inf_names, (_, NotImplicit) :: implicits ->
+ build_implicits inf_names implicits
+
+ (* With the current impargs API, it is impossible to make an originally
+ anonymous argument implicit *)
+ | Anonymous :: _, (name, _) :: _ ->
+ errorlabstrm "vernac_declare_arguments"
+ (strbrk"Argument "++ pr_name name ++
+ strbrk " cannot be declared implicit.")
+
+ | Name id :: inf_names, (name, impl) :: implicits ->
+ let max = impl = MaximallyImplicit in
+ (ExplByName id,max,false) :: build_implicits inf_names implicits
+
+ | _ -> assert false (* already checked in [names_union] *)
+ in
+
+ let implicits = List.map (build_implicits inf_names) implicits in
+ let implicits_specified = match implicits with [[]] -> false | _ -> true in
+
+ if implicits_specified && clear_implicits_flag then
+ error "The \"clear implicits\" flag is incompatible with implicit annotations";
+
+ if implicits_specified && default_implicits_flag then
+ error "The \"default implicits\" flag is incompatible with implicit annotations";
+
let rargs =
Util.List.map_filter (function (n, true) -> Some n | _ -> None)
- (Util.List.map_i (fun i (_, b, _,_,_) -> i, b) 0 l) in
- if some_scopes_specified || List.mem `ClearScopes flags then
- vernac_arguments_scope locality r scopes;
- if not some_implicits_specified && List.mem `DefaultImplicits flags then
- vernac_declare_implicits locality r []
- else if some_implicits_specified || List.mem `ClearImplicits flags then
- vernac_declare_implicits locality r implicits;
- if nargs >= 0 && nargs < List.fold_left max 0 rargs then
- error "The \"/\" option must be placed after the last \"!\".";
- let no_flags = List.is_empty flags in
+ (Util.List.map_i (fun i { recarg_like = b } -> i, b) 0 args)
+ in
+
let rec narrow = function
| #Reductionops.ReductionBehaviour.flag as x :: tl -> x :: narrow tl
- | [] -> [] | _ :: tl -> narrow tl in
- let flags = narrow flags in
- let some_simpl_flags_specified =
- not (List.is_empty rargs) || nargs >= 0 || not (List.is_empty flags) in
- if some_simpl_flags_specified then begin
+ | [] -> [] | _ :: tl -> narrow tl
+ in
+ let red_flags = narrow flags in
+ let red_modifiers_specified =
+ not (List.is_empty rargs) || Option.has_some nargs_for_red
+ || not (List.is_empty red_flags)
+ in
+
+ if not (List.is_empty rargs) && never_unfold_flag then
+ err_incompat "simpl never" "!";
+
+
+ (* Actions *)
+
+ if renaming_specified then begin
+ let local = make_section_locality locality in
+ Arguments_renaming.rename_arguments local sr names
+ end;
+
+ if scopes_specified || clear_scopes_flag then begin
+ let scopes = List.map (Option.map (fun (o,k) ->
+ try ignore (Notation.find_scope k); k
+ with UserError _ ->
+ Notation.find_delimiters_scope o k)) scopes
+ in
+ vernac_arguments_scope locality reference scopes
+ end;
+
+ if implicits_specified || clear_implicits_flag then
+ vernac_declare_implicits locality reference implicits;
+
+ if default_implicits_flag then
+ vernac_declare_implicits locality reference [];
+
+ if red_modifiers_specified then begin
match sr with
| ConstRef _ as c ->
Reductionops.ReductionBehaviour.set
- (make_section_locality locality) c (rargs, nargs, flags)
- | _ -> errorlabstrm "" (strbrk "Modifiers of the behavior of the simpl tactic are relevant for constants only.")
+ (make_section_locality locality) c
+ (rargs, Option.default ~-1 nargs_for_red, red_flags)
+ | _ -> errorlabstrm ""
+ (strbrk "Modifiers of the behavior of the simpl tactic "++
+ strbrk "are relevant for constants only.")
end;
- if not (some_renaming_specified ||
- some_implicits_specified ||
- some_scopes_specified ||
- some_simpl_flags_specified) &&
- no_flags then
- msg_warning (strbrk "This command is just asserting the number and names of arguments of " ++ pr_global sr ++ strbrk". If this is what you want add ': assert' to silence the warning. If you want to clear implicit arguments add ': clear implicits'. If you want to clear notation scopes add ': clear scopes'")
+ if not (renaming_specified ||
+ implicits_specified ||
+ scopes_specified ||
+ red_modifiers_specified) && (List.is_empty flags) then
+ warn_arguments_assert sr
let default_env () = {
Notation_term.ninterp_var_type = Id.Map.empty;
ninterp_rec_vars = Id.Map.empty;
- ninterp_only_parse = false;
}
let vernac_reserve bl =
@@ -1183,7 +1237,7 @@ let vernac_reserve bl =
let sigma = Evd.from_env env in
let t,ctx = Constrintern.interp_type env sigma c in
let t = Detyping.detype false [] env (Evd.from_ctx ctx) t in
- let t = Notation_ops.notation_constr_of_glob_constr (default_env ()) t in
+ let t,_ = Notation_ops.notation_constr_of_glob_constr (default_env ()) t in
Reserve.declare_reserved_type idl t)
in List.iter sb_decl bl
@@ -1370,8 +1424,8 @@ let _ =
optdepr = false;
optname = "kernel term sharing";
optkey = ["Kernel"; "Term"; "Sharing"];
- optread = (fun () -> !Closure.share);
- optwrite = (fun b -> Closure.share := b) }
+ optread = (fun () -> !CClosure.share);
+ optwrite = (fun b -> CClosure.share := b) }
(* No more undo limit in the new proof engine.
The command still exists for compatibility (e.g. with ProofGeneral) *)
@@ -1430,18 +1484,6 @@ let _ =
optread = Flags.get_dump_bytecode;
optwrite = Flags.set_dump_bytecode }
-let vernac_debug b =
- set_debug (if b then Tactic_debug.DebugOn 0 else Tactic_debug.DebugOff)
-
-let _ =
- declare_bool_option
- { optsync = false;
- optdepr = false;
- optname = "Ltac debug";
- optkey = ["Ltac";"Debug"];
- optread = (fun () -> get_debug () != Tactic_debug.DebugOff);
- optwrite = vernac_debug }
-
let _ =
declare_bool_option
{ optsync = true;
@@ -1451,6 +1493,15 @@ let _ =
optread = (fun () -> !Constrintern.parsing_explicit);
optwrite = (fun b -> Constrintern.parsing_explicit := b) }
+let _ =
+ declare_string_option ~preprocess:CWarnings.normalize_flags_string
+ { optsync = true;
+ optdepr = false;
+ optname = "warnings display";
+ optkey = ["Warnings"];
+ optread = CWarnings.get_flags;
+ optwrite = CWarnings.set_flags }
+
let vernac_set_strategy locality l =
let local = make_locality locality in
let glob_ref r =
@@ -1480,6 +1531,9 @@ let vernac_set_option locality key = function
| IntValue n -> set_int_option_value_gen locality key n
| BoolValue b -> set_bool_option_value_gen locality key b
+let vernac_set_append_option locality key s =
+ set_string_option_append_value_gen locality key s
+
let vernac_unset_option locality key =
unset_option_value_gen locality key
@@ -1519,7 +1573,7 @@ let get_current_context_of_args = function
let vernac_check_may_eval redexp glopt rc =
let (sigma, env) = get_current_context_of_args glopt in
let sigma', c = interp_open_constr env sigma rc in
- let sigma' = Evarconv.consider_remaining_unif_problems env sigma' in
+ let sigma' = Evarconv.solve_unif_constraints_with_heuristics env sigma' in
Evarconv.check_problems_are_solved env sigma';
let sigma',nf = Evarutil.nf_evars_and_universes sigma' in
let pl, uctx = Evd.universe_context sigma' in
@@ -1535,18 +1589,22 @@ let vernac_check_may_eval redexp glopt rc =
| None ->
let l = Evar.Set.union (Evd.evars_of_term j.Environ.uj_val) (Evd.evars_of_term j.Environ.uj_type) in
let j = { j with Environ.uj_type = Reductionops.nf_betaiota sigma' j.Environ.uj_type } in
- msg_notice (print_judgment env sigma' j ++
+ Feedback.msg_notice (print_judgment env sigma' j ++
pr_ne_evar_set (fnl () ++ str "where" ++ fnl ()) (mt ()) sigma' l ++
Printer.pr_universe_ctx sigma uctx)
| Some r ->
- Tacintern.dump_glob_red_expr r;
- let (sigma',r_interp) = interp_redexp env sigma' r in
- let redfun env evm c = snd (fst (reduction_of_red_expr env r_interp) env evm c) in
- msg_notice (print_eval redfun env sigma' rc j)
+ let (sigma',r_interp) = Hook.get f_interp_redexp env sigma' r in
+ let redfun env evm c =
+ let (redfun, _) = reduction_of_red_expr env r_interp in
+ let evm = Sigma.Unsafe.of_evar_map evm in
+ let Sigma (c, _, _) = redfun.Reductionops.e_redfun env evm c in
+ c
+ in
+ Feedback.msg_notice (print_eval redfun env sigma' rc j)
let vernac_declare_reduction locality s r =
let local = make_locality locality in
- declare_red_expr local s (snd (interp_redexp (Global.env()) Evd.empty r))
+ declare_red_expr local s (snd (Hook.get f_interp_redexp (Global.env()) Evd.empty r))
(* The same but avoiding the current goal context if any *)
let vernac_global_check c =
@@ -1554,11 +1612,11 @@ let vernac_global_check c =
let sigma = Evd.from_env env in
let c,ctx = interp_constr env sigma c in
let senv = Global.safe_env() in
- let cstrs = snd (Evd.evar_universe_context_set Univ.UContext.empty ctx) in
+ let cstrs = snd (UState.context_set ctx) in
let senv = Safe_typing.add_constraints cstrs senv in
let j = Safe_typing.typing senv c in
let env = Safe_typing.env_of_safe_env senv in
- msg_notice (print_safe_judgment env sigma j)
+ Feedback.msg_notice (print_safe_judgment env sigma j)
let get_nth_goal n =
@@ -1572,6 +1630,7 @@ exception NoHyp
We only print the type and a small statement to this comes from the
goal. Precondition: there must be at least one current goal. *)
let print_about_hyp_globs ref_or_by_not glnumopt =
+ let open Context.Named.Declaration in
try
let gl,id =
match glnumopt,ref_or_by_not with
@@ -1584,17 +1643,17 @@ let print_about_hyp_globs ref_or_by_not glnumopt =
(str "No such goal: " ++ int n ++ str "."))
| _ , _ -> raise NoHyp in
let hyps = pf_hyps gl in
- let (id,bdyopt,typ) = Context.lookup_named id hyps in
- let natureofid = match bdyopt with
- | None -> "Hypothesis"
- | Some bdy ->"Constant (let in)" in
- v 0 (str (Id.to_string id) ++ str":" ++ pr_constr typ ++ fnl() ++ fnl()
+ let decl = Context.Named.lookup id hyps in
+ let natureofid = match decl with
+ | LocalAssum _ -> "Hypothesis"
+ | LocalDef (_,bdy,_) ->"Constant (let in)" in
+ v 0 (pr_id id ++ str":" ++ pr_constr (get_type decl) ++ fnl() ++ fnl()
++ str natureofid ++ str " of the goal context.")
with (* fallback to globals *)
| NoHyp | Not_found -> print_about ref_or_by_not
-let vernac_print = function
+let vernac_print = let open Feedback in function
| PrintTables -> msg_notice (print_tables ())
| PrintFullContext-> msg_notice (print_full_context_typ ())
| PrintSectionContext qid -> msg_notice (print_sec_context_typ qid)
@@ -1613,26 +1672,24 @@ let vernac_print = function
| PrintClasses -> msg_notice (Prettyp.print_classes())
| PrintTypeClasses -> msg_notice (Prettyp.print_typeclasses())
| PrintInstances c -> msg_notice (Prettyp.print_instances (smart_global c))
- | PrintLtac qid -> msg_notice (Tacintern.print_ltac (snd (qualid_of_reference qid)))
| PrintCoercions -> msg_notice (Prettyp.print_coercions())
| PrintCoercionPaths (cls,clt) ->
msg_notice (Prettyp.print_path_between (cl_of_qualid cls) (cl_of_qualid clt))
| PrintCanonicalConversions -> msg_notice (Prettyp.print_canonical_projections ())
| PrintUniverses (b, dst) ->
let univ = Global.universes () in
- let univ = if b then Univ.sort_universes univ else univ in
+ let univ = if b then UGraph.sort_universes univ else univ in
let pr_remaining =
if Global.is_joined_environment () then mt ()
else str"There may remain asynchronous universe constraints"
in
begin match dst with
- | None -> msg_notice (Univ.pr_universes Universes.pr_with_global_universes univ ++ pr_remaining)
+ | None -> msg_notice (UGraph.pr_universes Universes.pr_with_global_universes univ ++ pr_remaining)
| Some s -> dump_universes_gen univ s
end
| PrintHint r -> msg_notice (Hints.pr_hint_ref (smart_global r))
| PrintHintGoal -> msg_notice (Hints.pr_applicable_hint ())
| PrintHintDbName s -> msg_notice (Hints.pr_hint_db_by_name s)
- | PrintRewriteHintDbName s -> msg_notice (Autorewrite.print_rewrite_hintdb s)
| PrintHintDb -> msg_notice (Hints.pr_searchtable ())
| PrintScopes ->
msg_notice (Notation.pr_scopes (Constrextern.without_symbols pr_lglob_constr))
@@ -1684,6 +1741,28 @@ let interp_search_about_item env =
errorlabstrm "interp_search_about_item"
(str "Unable to interp \"" ++ str s ++ str "\" either as a reference or as an identifier component")
+(* 05f22a5d6d5b8e3e80f1a37321708ce401834430 introduced the
+ `search_output_name_only` option to avoid excessive printing when
+ searching.
+
+ The motivation was to make search usable for IDE completion,
+ however, it is still too slow due to the non-indexed nature of the
+ underlying search mechanism.
+
+ In the future we should deprecate the option and provide a fast,
+ indexed name-searching interface.
+*)
+let search_output_name_only = ref false
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "output-name-only search";
+ optkey = ["Search";"Output";"Name";"Only"];
+ optread = (fun () -> !search_output_name_only);
+ optwrite = (:=) search_output_name_only }
+
let vernac_search s gopt r =
let r = interp_search_restriction r in
let env,gopt =
@@ -1695,18 +1774,28 @@ let vernac_search s gopt r =
| Some g -> snd (Pfedit.get_goal_context g) , Some g
in
let get_pattern c = snd (intern_constr_pattern env c) in
+ let pr_search ref env c =
+ let pr = pr_global ref in
+ let pp = if !search_output_name_only
+ then pr
+ else begin
+ let pc = pr_lconstr_env env Evd.empty c in
+ hov 2 (pr ++ str":" ++ spc () ++ pc)
+ end
+ in Feedback.msg_notice pp
+ in
match s with
| SearchPattern c ->
- msg_notice (Search.search_pattern gopt (get_pattern c) r)
+ Search.search_pattern gopt (get_pattern c) r pr_search
| SearchRewrite c ->
- msg_notice (Search.search_rewrite gopt (get_pattern c) r)
+ Search.search_rewrite gopt (get_pattern c) r pr_search
| SearchHead c ->
- msg_notice (Search.search_by_head gopt (get_pattern c) r)
+ Search.search_by_head gopt (get_pattern c) r pr_search
| SearchAbout sl ->
- msg_notice (Search.search_about gopt (List.map (on_snd (interp_search_about_item env)) sl) r)
+ Search.search_about gopt (List.map (on_snd (interp_search_about_item env)) sl) r pr_search
-let vernac_locate = function
- | LocateAny (AN qid) -> msg_notice (print_located_qualid qid)
+let vernac_locate = let open Feedback in function
+ | LocateAny (AN qid) -> msg_notice (print_located_qualid qid)
| LocateTerm (AN qid) -> msg_notice (print_located_term qid)
| LocateAny (ByNotation (_, ntn, sc)) (** TODO : handle Ltac notations *)
| LocateTerm (ByNotation (_, ntn, sc)) ->
@@ -1736,7 +1825,7 @@ let vernac_focus gln =
match gln with
| None -> Proof.focus focus_command_cond () 1 p
| Some 0 ->
- Errors.error "Invalid goal number: 0. Goal numbering starts with 1."
+ CErrors.error "Invalid goal number: 0. Goal numbering starts with 1."
| Some n ->
Proof.focus focus_command_cond () n p)
@@ -1749,7 +1838,7 @@ let vernac_unfocus () =
let vernac_unfocused () =
let p = Proof_global.give_me_the_proof () in
if Proof.unfocused p then
- msg_notice (str"The proof is indeed fully unfocused.")
+ Feedback.msg_notice (str"The proof is indeed fully unfocused.")
else
error "The proof is not fully unfocused."
@@ -1777,7 +1866,7 @@ let vernac_bullet (bullet:Proof_global.Bullet.t) =
Proof_global.simple_with_current_proof (fun _ p ->
Proof_global.Bullet.put p bullet)
-let vernac_show = function
+let vernac_show = let open Feedback in function
| ShowGoal goalref ->
let info = match goalref with
| OpenSubgoals -> pr_open_subgoals ()
@@ -1815,33 +1904,36 @@ let vernac_check_guard () =
with UserError(_,s) ->
(str ("Condition violated: ") ++s)
in
- msg_notice message
+ Feedback.msg_notice message
exception End_of_input
let vernac_load interp fname =
+ let interp x =
+ let proof_mode = Proof_global.get_default_proof_mode_name () in
+ Proof_global.activate_proof_mode proof_mode;
+ interp x in
let parse_sentence = Flags.with_option Flags.we_are_parsing
(fun po ->
match Pcoq.Gram.entry_parse Pcoq.main_entry po with
| Some x -> x
| None -> raise End_of_input) in
let fname =
- Envars.expand_path_macros ~warn:(fun x -> msg_warning (str x)) fname in
+ Envars.expand_path_macros ~warn:(fun x -> Feedback.msg_warning (str x)) fname in
let fname = CUnix.make_suffix fname ".v" in
let input =
let longfname = Loadpath.locate_file fname in
let in_chan = open_utf8_file_in longfname in
- Pcoq.Gram.parsable (Stream.of_channel in_chan) in
+ Pcoq.Gram.parsable ~file:longfname (Stream.of_channel in_chan) in
try while true do interp (snd (parse_sentence input)) done
with End_of_input -> ()
-
(* "locality" is the prefix "Local" attribute, while the "local" component
* is the outdated/deprecated "Local" attribute of some vernacular commands
* still parsed as the obsolete_locality grammar entry for retrocompatibility.
* loc is the Loc.t of the vernacular command being interpreted. *)
let interp ?proof ~loc locality poly c =
- prerr_endline ("interpreting: " ^ Pp.string_of_ppcmds (Ppvernac.pr_vernac c));
+ prerr_endline (fun () -> "interpreting: " ^ Pp.string_of_ppcmds (Ppvernac.pr_vernac c));
match c with
(* Done later in this file *)
| VernacLoad _ -> assert false
@@ -1854,8 +1946,6 @@ let interp ?proof ~loc locality poly c =
| VernacError e -> raise e
(* Syntax *)
- | VernacTacticNotation (n,r,e) ->
- Metasyntax.add_tactic_notation (make_module_locality locality,n,r,e)
| VernacSyntaxExtension (local,sl) ->
vernac_syntax_extension locality local sl
| VernacDelimiters (sc,lr) -> vernac_delimiters sc lr
@@ -1906,14 +1996,13 @@ let interp ?proof ~loc locality poly c =
vernac_identity_coercion locality poly local id s t
(* Type classes *)
- | VernacInstance (abst, sup, inst, props, pri) ->
- vernac_instance abst locality poly sup inst props pri
+ | VernacInstance (abst, sup, inst, props, info) ->
+ vernac_instance abst locality poly sup inst props info
| VernacContext sup -> vernac_context poly sup
- | VernacDeclareInstances (ids, pri) -> vernac_declare_instances locality ids pri
+ | VernacDeclareInstances insts -> vernac_declare_instances locality insts
| VernacDeclareClass id -> vernac_declare_class id
(* Solving *)
- | VernacSolve (n,info,tac,b) -> vernac_solve n info tac b
| VernacSolveExistential (n,c) -> vernac_solve_existential n c
(* Auxiliary file and library management *)
@@ -1934,8 +2023,6 @@ let interp ?proof ~loc locality poly c =
| VernacBackTo _ -> anomaly (str "VernacBackTo not handled by Stm")
(* Commands *)
- | VernacDeclareTacticDefinition def ->
- vernac_declare_tactic_definition locality def
| VernacCreateHintDb (dbname,b) -> vernac_create_hintdb locality dbname b
| VernacRemoveHints (dbnames,ids) -> vernac_remove_hints locality dbnames ids
| VernacHints (local,dbnames,hints) ->
@@ -1944,13 +2031,14 @@ let interp ?proof ~loc locality poly c =
vernac_syntactic_definition locality id c local b
| VernacDeclareImplicits (qid,l) ->
vernac_declare_implicits locality qid l
- | VernacArguments (qid, l, narg, flags) ->
- vernac_declare_arguments locality qid l narg flags
+ | VernacArguments (qid, args, more_implicits, nargs, flags) ->
+ vernac_arguments locality qid args more_implicits nargs flags
| VernacReserve bl -> vernac_reserve bl
| VernacGeneralizable gen -> vernac_generalizable locality gen
| VernacSetOpacity qidl -> vernac_set_opacity locality qidl
| VernacSetStrategy l -> vernac_set_strategy locality l
| VernacSetOption (key,v) -> vernac_set_option locality key v
+ | VernacSetAppendOption (key,v) -> vernac_set_append_option locality key v
| VernacUnsetOption key -> vernac_unset_option locality key
| VernacRemoveOption (key,v) -> vernac_remove_option key v
| VernacAddOption (key,v) -> vernac_add_option key v
@@ -1963,16 +2051,15 @@ let interp ?proof ~loc locality poly c =
| VernacSearch (s,g,r) -> vernac_search s g r
| VernacLocate l -> vernac_locate l
| VernacRegister (id, r) -> vernac_register id r
- | VernacComments l -> if_verbose msg_info (str "Comments ok\n")
- | VernacNop -> ()
+ | VernacComments l -> if_verbose Feedback.msg_info (str "Comments ok\n")
(* The STM should handle that, but LOAD bypasses the STM... *)
- | VernacAbort id -> msg_warning (str "VernacAbort not handled by Stm")
- | VernacAbortAll -> msg_warning (str "VernacAbortAll not handled by Stm")
- | VernacRestart -> msg_warning (str "VernacRestart not handled by Stm")
- | VernacUndo _ -> msg_warning (str "VernacUndo not handled by Stm")
- | VernacUndoTo _ -> msg_warning (str "VernacUndoTo not handled by Stm")
- | VernacBacktrack _ -> msg_warning (str "VernacBacktrack not handled by Stm")
+ | VernacAbort id -> CErrors.errorlabstrm "" (str "Abort cannot be used through the Load command")
+ | VernacAbortAll -> CErrors.errorlabstrm "" (str "AbortAll cannot be used through the Load command")
+ | VernacRestart -> CErrors.errorlabstrm "" (str "Restart cannot be used through the Load command")
+ | VernacUndo _ -> CErrors.errorlabstrm "" (str "Undo cannot be used through the Load command")
+ | VernacUndoTo _ -> CErrors.errorlabstrm "" (str "UndoTo cannot be used through the Load command")
+ | VernacBacktrack _ -> CErrors.errorlabstrm "" (str "Backtrack cannot be used through the Load command")
(* Proof management *)
| VernacGoal t -> vernac_start_proof locality poly Theorem [None,([],t,None)] false
@@ -2012,25 +2099,23 @@ let check_vernac_supports_locality c l =
match l, c with
| None, _ -> ()
| Some _, (
- VernacTacticNotation _
- | VernacOpenCloseScope _
+ VernacOpenCloseScope _
| VernacSyntaxExtension _ | VernacInfix _ | VernacNotation _
| VernacDefinition _ | VernacFixpoint _ | VernacCoFixpoint _
| VernacAssumption _ | VernacStartTheoremProof _
| VernacCoercion _ | VernacIdentityCoercion _
| VernacInstance _ | VernacDeclareInstances _
| VernacDeclareMLModule _
- | VernacDeclareTacticDefinition _
| VernacCreateHintDb _ | VernacRemoveHints _ | VernacHints _
| VernacSyntacticDefinition _
| VernacArgumentsScope _ | VernacDeclareImplicits _ | VernacArguments _
| VernacGeneralizable _
| VernacSetOpacity _ | VernacSetStrategy _
- | VernacSetOption _ | VernacUnsetOption _
+ | VernacSetOption _ | VernacSetAppendOption _ | VernacUnsetOption _
| VernacDeclareReduction _
| VernacExtend _
| VernacInductive _) -> ()
- | Some _, _ -> Errors.error "This command does not support Locality"
+ | Some _, _ -> CErrors.error "This command does not support Locality"
(* Vernaculars that take a polymorphism flag *)
let check_vernac_supports_polymorphism c p =
@@ -2044,13 +2129,13 @@ let check_vernac_supports_polymorphism c p =
| VernacInstance _ | VernacDeclareInstances _
| VernacHints _ | VernacContext _
| VernacExtend _ | VernacUniverse _ | VernacConstraint _) -> ()
- | Some _, _ -> Errors.error "This command does not support Polymorphism"
+ | Some _, _ -> CErrors.error "This command does not support Polymorphism"
let enforce_polymorphism = function
| None -> Flags.is_universe_polymorphism ()
| Some b -> Flags.make_polymorphic_flag b; b
-(** A global default timeout, controled by option "Set Default Timeout n".
+(** A global default timeout, controlled by option "Set Default Timeout n".
Use "Unset Default Timeout" to deactivate it (or set it to 0). *)
let default_timeout = ref None
@@ -2098,17 +2183,17 @@ let with_fail b f =
with
| HasNotFailed as e -> raise e
| e ->
- let e = Errors.push e in
- raise (HasFailed (Errors.iprint
- (Cerrors.process_vernac_interp_error ~allow_uncaught:false ~with_header:false e))))
+ let e = CErrors.push e in
+ raise (HasFailed (CErrors.iprint
+ (ExplainErr.process_vernac_interp_error ~allow_uncaught:false ~with_header:false e))))
()
- with e when Errors.noncritical e ->
- let (e, _) = Errors.push e in
+ with e when CErrors.noncritical e ->
+ let (e, _) = CErrors.push e in
match e with
| HasNotFailed ->
errorlabstrm "Fail" (str "The command has not failed!")
| HasFailed msg ->
- if is_verbose () || !test_mode || !ide_slave then msg_info
+ if is_verbose () || !test_mode || !ide_slave then Feedback.msg_info
(str "The command has indeed failed with message:" ++ fnl () ++ msg)
| _ -> assert false
end
@@ -2117,13 +2202,13 @@ let interp ?(verbosely=true) ?proof (loc,c) =
let orig_program_mode = Flags.is_program_mode () in
let rec aux ?locality ?polymorphism isprogcmd = function
| VernacProgram c when not isprogcmd -> aux ?locality ?polymorphism true c
- | VernacProgram _ -> Errors.error "Program mode specified twice"
+ | VernacProgram _ -> CErrors.error "Program mode specified twice"
| VernacLocal (b, c) when Option.is_empty locality ->
aux ~locality:b ?polymorphism isprogcmd c
| VernacPolymorphic (b, c) when polymorphism = None ->
aux ?locality ~polymorphism:b isprogcmd c
- | VernacPolymorphic (b, c) -> Errors.error "Polymorphism specified twice"
- | VernacLocal _ -> Errors.error "Locality specified twice"
+ | VernacPolymorphic (b, c) -> CErrors.error "Polymorphism specified twice"
+ | VernacLocal _ -> CErrors.error "Locality specified twice"
| VernacStm (Command c) -> aux ?locality ?polymorphism isprogcmd c
| VernacStm (PGLast c) -> aux ?locality ?polymorphism isprogcmd c
| VernacStm _ -> assert false (* Done by Stm *)
@@ -2132,11 +2217,11 @@ let interp ?(verbosely=true) ?proof (loc,c) =
| VernacTimeout (n,v) ->
current_timeout := Some n;
aux ?locality ?polymorphism isprogcmd v
- | VernacRedirect (s, v) ->
- Pp.with_output_to_file s (aux_list ?locality ?polymorphism isprogcmd) v;
- | VernacTime v ->
+ | VernacRedirect (s, (_,v)) ->
+ Feedback.with_output_to_file s (aux false) v
+ | VernacTime (_,v) ->
System.with_time !Flags.time
- (aux_list ?locality ?polymorphism isprogcmd) v;
+ (aux ?locality ?polymorphism isprogcmd) v;
| VernacLoad (_,fname) -> vernac_load (aux false) fname
| c ->
check_vernac_supports_locality c locality;
@@ -2156,16 +2241,14 @@ let interp ?(verbosely=true) ?proof (loc,c) =
| reraise when
(match reraise with
| Timeout -> true
- | e -> Errors.noncritical e)
+ | e -> CErrors.noncritical e)
->
- let e = Errors.push reraise in
+ let e = CErrors.push reraise in
let e = locate_if_not_already loc e in
let () = restore_timeout () in
Flags.program_mode := orig_program_mode;
ignore (Flags.use_polymorphic_flag ());
iraise e
- and aux_list ?locality ?polymorphism isprogcmd l =
- List.iter (aux false) (List.map snd l)
in
if verbosely then Flags.verbosely (aux false) c
else aux false c
diff --git a/toplevel/vernacentries.mli b/toplevel/vernacentries.mli
index 451ccdb4..4e7fa4a0 100644
--- a/toplevel/vernacentries.mli
+++ b/toplevel/vernacentries.mli
@@ -59,3 +59,8 @@ val vernac_end_proof :
?proof:Proof_global.closed_proof -> Vernacexpr.proof_end -> unit
val with_fail : bool -> (unit -> unit) -> unit
+
+val command_focus : unit Proof.focus_kind
+
+val interp_redexp_hook : (Environ.env -> Evd.evar_map -> Tacexpr.raw_red_expr ->
+ Evd.evar_map * Redexpr.red_expr) Hook.t
diff --git a/toplevel/vernacinterp.ml b/toplevel/vernacinterp.ml
index 7fbd2b11..d81e3d6b 100644
--- a/toplevel/vernacinterp.ml
+++ b/toplevel/vernacinterp.ml
@@ -8,7 +8,7 @@
open Util
open Pp
-open Errors
+open CErrors
type deprecation = bool
type vernac_command = Genarg.raw_generic_argument list -> unit -> unit
@@ -42,6 +42,11 @@ let vinterp_map s =
let vinterp_init () = Hashtbl.clear vernac_tab
+let warn_deprecated_command =
+ let open CWarnings in
+ create ~name:"deprecated-command" ~category:"deprecated"
+ (fun pr -> str "Deprecated vernacular command: " ++ pr)
+
(* Interpretation of a vernac command *)
let call ?locality (opn,converted_args) =
@@ -55,7 +60,7 @@ let call ?locality (opn,converted_args) =
| Egramml.GramNonTerminal _ -> str "_"
in
let pr = pr_sequence pr_gram rules in
- msg_warning (str "Deprecated vernacular command: " ++ pr)
+ warn_deprecated_command pr;
in
loc:= "Checking arguments";
let hunk = callback converted_args in
@@ -66,7 +71,7 @@ let call ?locality (opn,converted_args) =
with
| Drop -> raise Drop
| reraise ->
- let reraise = Errors.push reraise in
+ let reraise = CErrors.push reraise in
if !Flags.debug then
- msg_debug (str"Vernac Interpreter " ++ str !loc);
+ Feedback.msg_debug (str"Vernac Interpreter " ++ str !loc);
iraise reraise