summaryrefslogtreecommitdiff
path: root/toplevel
diff options
context:
space:
mode:
authorGravatar Enrico Tassi <gareuselesinge@debian.org>2015-01-25 14:42:51 +0100
committerGravatar Enrico Tassi <gareuselesinge@debian.org>2015-01-25 14:42:51 +0100
commit7cfc4e5146be5666419451bdd516f1f3f264d24a (patch)
treee4197645da03dc3c7cc84e434cc31d0a0cca7056 /toplevel
parent420f78b2caeaaddc6fe484565b2d0e49c66888e5 (diff)
Imported Upstream version 8.5~beta1+dfsg
Diffstat (limited to 'toplevel')
-rw-r--r--toplevel/auto_ind_decl.ml674
-rw-r--r--toplevel/auto_ind_decl.mli15
-rw-r--r--toplevel/autoinstance.ml320
-rw-r--r--toplevel/autoinstance.mli34
-rw-r--r--toplevel/backtrack.ml243
-rw-r--r--toplevel/backtrack.mli99
-rw-r--r--toplevel/cerrors.ml120
-rw-r--r--toplevel/cerrors.mli11
-rw-r--r--toplevel/class.ml159
-rw-r--r--toplevel/class.mli29
-rw-r--r--toplevel/classes.ml418
-rw-r--r--toplevel/classes.mli48
-rw-r--r--toplevel/command.ml1072
-rw-r--r--toplevel/command.mli91
-rw-r--r--toplevel/coqinit.ml135
-rw-r--r--toplevel/coqinit.mli8
-rw-r--r--toplevel/coqloop.ml (renamed from toplevel/toplevel.ml)271
-rw-r--r--toplevel/coqloop.mli (renamed from toplevel/toplevel.mli)7
-rw-r--r--toplevel/coqtop.ml687
-rw-r--r--toplevel/coqtop.mli10
-rw-r--r--toplevel/discharge.ml72
-rw-r--r--toplevel/discharge.mli8
-rw-r--r--toplevel/g_obligations.ml4135
-rw-r--r--toplevel/himsg.ml1078
-rw-r--r--toplevel/himsg.mli10
-rw-r--r--toplevel/ide_intf.ml713
-rw-r--r--toplevel/ide_intf.mli62
-rw-r--r--toplevel/ide_slave.ml466
-rw-r--r--toplevel/ide_slave.mli17
-rw-r--r--toplevel/ind_tables.ml119
-rw-r--r--toplevel/ind_tables.mli24
-rw-r--r--toplevel/indschemes.ml126
-rw-r--r--toplevel/indschemes.mli15
-rw-r--r--toplevel/interface.mli231
-rw-r--r--toplevel/lemmas.ml353
-rw-r--r--toplevel/lemmas.mli63
-rw-r--r--toplevel/libtypes.ml111
-rw-r--r--toplevel/libtypes.mli26
-rw-r--r--toplevel/locality.ml99
-rw-r--r--toplevel/locality.mli51
-rw-r--r--toplevel/metasyntax.ml832
-rw-r--r--toplevel/metasyntax.mli25
-rw-r--r--toplevel/mltop.ml439
-rw-r--r--toplevel/mltop.ml4337
-rw-r--r--toplevel/mltop.mli35
-rw-r--r--toplevel/obligations.ml1075
-rw-r--r--toplevel/obligations.mli116
-rw-r--r--toplevel/record.ml478
-rw-r--r--toplevel/record.mli26
-rw-r--r--toplevel/search.ml465
-rw-r--r--toplevel/search.mli87
-rw-r--r--toplevel/toplevel.mllib12
-rw-r--r--toplevel/usage.ml57
-rw-r--r--toplevel/usage.mli5
-rw-r--r--toplevel/vernac.ml379
-rw-r--r--toplevel/vernac.mli26
-rw-r--r--toplevel/vernacentries.ml1724
-rw-r--r--toplevel/vernacentries.mli39
-rw-r--r--toplevel/vernacexpr.ml508
-rw-r--r--toplevel/vernacinterp.ml34
-rw-r--r--toplevel/vernacinterp.mli12
-rw-r--r--toplevel/whelp.ml473
-rw-r--r--toplevel/whelp.mli4
63 files changed, 7737 insertions, 7281 deletions
diff --git a/toplevel/auto_ind_decl.ml b/toplevel/auto_ind_decl.ml
index 2e775f30..26b54a73 100644
--- a/toplevel/auto_ind_decl.ml
+++ b/toplevel/auto_ind_decl.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -10,24 +10,22 @@
decidable equality, created by Vincent Siles, Oct 2007 *)
open Tacmach
+open Errors
open Util
-open Flags
-open Decl_kinds
open Pp
-open Entries
+open Term
+open Vars
open Termops
open Declarations
-open Declare
-open Term
open Names
-open Libnames
-open Goptions
-open Mod_subst
-open Indrec
+open Globnames
open Inductiveops
open Tactics
-open Tacticals
open Ind_tables
+open Misctypes
+open Proofview.Notations
+
+let out_punivs = Univ.out_punivs
(**********************************************************************)
(* Generic synthesis of boolean equality *)
@@ -45,9 +43,9 @@ and aux = function
if n > (List.length l) then failwith "quick_chop args"
else kick_last (aux (n,l) )
-let rec deconstruct_type t =
+let deconstruct_type t =
let l,r = decompose_prod t in
- (List.map (fun (_,b) -> b) (List.rev l))@[r]
+ (List.rev_map snd l)@[r]
exception EqNotFound of inductive * inductive
exception EqUnknown of string
@@ -57,7 +55,9 @@ exception InductiveWithSort
exception ParameterWithoutEquality of constant
exception NonSingletonProp of inductive
-let dl = dummy_loc
+let dl = Loc.ghost
+
+let constr_of_global g = lazy (Universes.constr_of_global g)
(* Some pre declaration of constant we are going to use *)
let bb = constr_of_global Coqlib.glob_bool
@@ -77,27 +77,21 @@ let sumbool = Coqlib.build_coq_sumbool
let andb = fun _ -> (Coqlib.build_bool_type()).Coqlib.andb
-let induct_on c =
- new_induct false
- [Tacexpr.ElimOnConstr (Evd.empty,(c,Glob_term.NoBindings))]
- None (None,None) None
+let induct_on c = induction false None c None None
+
+let destruct_on c = destruct false None c None None
let destruct_on_using c id =
- new_destruct false
- [Tacexpr.ElimOnConstr (Evd.empty,(c,Glob_term.NoBindings))]
- None
- (None,Some (dl,Genarg.IntroOrAndPattern [
- [dl,Genarg.IntroAnonymous];
- [dl,Genarg.IntroIdentifier id]]))
+ destruct false None c
+ (Some (dl,[[dl,IntroNaming IntroAnonymous];
+ [dl,IntroNaming (IntroIdentifier id)]]))
None
-let destruct_on c =
- new_destruct false
- [Tacexpr.ElimOnConstr (Evd.empty,(c,Glob_term.NoBindings))]
- None (None,None) None
+let destruct_on_as c l =
+ destruct false None c (Some (dl,l)) None
(* reconstruct the inductive with the correct deBruijn indexes *)
-let mkFullInd ind n =
+let mkFullInd (ind,u) n =
let mib = Global.lookup_mind (fst ind) in
let nparams = mib.mind_nparams in
let nparrec = mib.mind_nparams_rec in
@@ -105,12 +99,12 @@ let mkFullInd ind n =
let lnonparrec,lnamesparrec =
context_chop (nparams-nparrec) mib.mind_params_ctxt in
if nparrec > 0
- then mkApp (mkInd ind,
+ then mkApp (mkIndU (ind,u),
Array.of_list(extended_rel_list (nparrec+n) lnamesparrec))
- else mkInd ind
+ else mkIndU (ind,u)
let check_bool_is_defined () =
- try let _ = Global.type_of_global Coqlib.glob_bool in ()
+ try let _ = Global.type_of_global_unsafe Coqlib.glob_bool in ()
with e when Errors.noncritical e -> raise (UndefinedCst "bool")
let beq_scheme_kind_aux = ref (fun _ -> failwith "Undefined")
@@ -131,22 +125,22 @@ let build_beq_scheme kn =
context_chop (nparams-nparrec) mib.mind_params_ctxt in
(* predef coq's boolean type *)
(* rec name *)
- let rec_name i =(string_of_id (Array.get mib.mind_packets i).mind_typename)^
+ let rec_name i =(Id.to_string (Array.get mib.mind_packets i).mind_typename)^
"_eqrec"
in
(* construct the "fun A B ... N, eqA eqB eqC ... N => fixpoint" part *)
let create_input c =
let myArrow u v = mkArrow u (lift 1 v)
and eqName = function
- | Name s -> id_of_string ("eq_"^(string_of_id s))
- | Anonymous -> id_of_string "eq_A"
+ | 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 lift_cnt = ref 0 in
let eqs_typ = List.map (fun aa ->
let a = lift !lift_cnt aa in
incr lift_cnt;
- myArrow a (myArrow a bb)
+ myArrow a (myArrow a (Lazy.force bb))
) ext_rel_list in
let eq_input = List.fold_left2
@@ -159,15 +153,16 @@ let build_beq_scheme kn =
List.fold_left (fun a (n,_,t) ->(* 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")
+ (match n with Name s -> s | Anonymous -> Id.of_string "A")
t a) eq_input lnamesparrec
in
let make_one_eq cur =
- let ind = kn,cur in
+ let u = Univ.Instance.empty in
+ let ind = (kn,cur),u (* FIXME *) in
(* current inductive we are working on *)
- let cur_packet = mib.mind_packets.(snd ind) in
+ let cur_packet = mib.mind_packets.(snd (fst ind)) in
(* Inductive toto : [rettyp] := *)
- let rettyp = Inductive.type_of_inductive env (mib,cur_packet) in
+ let rettyp = Inductive.type_of_inductive env ((mib,cur_packet),u) in
(* split rettyp in a list without the non rec params and the last ->
e.g. Inductive vec (A:Set) : nat -> Set := ... will do [nat] *)
let rettyp_l = quick_chop nparrec (deconstruct_type rettyp) in
@@ -182,31 +177,39 @@ let build_beq_scheme kn =
let rec aux c =
let (c,a) = Reductionops.whd_betaiota_stack Evd.empty c in
match kind_of_term c with
- | Rel x -> mkRel (x-nlist+ndx)
- | Var x -> mkVar (id_of_string ("eq_"^(string_of_id x)))
+ | Rel x -> mkRel (x-nlist+ndx), Declareops.no_seff
+ | Var x -> mkVar (id_of_string ("eq_"^(string_of_id x))), Declareops.no_seff
| Cast (x,_,_) -> aux (applist (x,a))
| App _ -> assert false
- | Ind (kn',i as ind') -> if eq_mind kn kn' then mkRel(eqA-nlist-i+nb_ind-1)
- else ( try
- let a = Array.of_list a in
- let eq = mkConst (find_scheme (!beq_scheme_kind_aux()) (kn',i))
- and eqa = Array.map aux a
- in
- let args = Array.append
- (Array.map (fun x->lift lifti x) a) eqa
- in if args = [||] then eq
- else mkApp (eq,Array.append
- (Array.map (fun x->lift lifti x) a) eqa)
- with Not_found -> raise(EqNotFound (ind',ind))
- )
+ | Ind ((kn',i as ind'),u) (*FIXME: universes *) ->
+ if eq_mind kn kn' then mkRel(eqA-nlist-i+nb_ind-1), Declareops.no_seff
+ else begin
+ try
+ let eq, eff =
+ let c, eff = find_scheme (!beq_scheme_kind_aux()) (kn',i) in
+ mkConst c, eff in
+ let eqa, eff =
+ let eqa, effs = List.split (List.map aux a) in
+ Array.of_list eqa,
+ Declareops.union_side_effects
+ (Declareops.flatten_side_effects (List.rev effs))
+ eff in
+ let args =
+ Array.append
+ (Array.of_list (List.map (fun x -> lift lifti x) a)) eqa in
+ if Int.equal (Array.length args) 0 then eq, eff
+ else mkApp (eq, args), eff
+ with Not_found -> raise(EqNotFound (ind', fst ind))
+ end
| Sort _ -> raise InductiveWithSort
| Prod _ -> raise InductiveWithProduct
| Lambda _-> raise (EqUnknown "Lambda")
| LetIn _ -> raise (EqUnknown "LetIn")
| Const kn ->
- (match Environ.constant_opt_value env kn with
- | None -> raise (ParameterWithoutEquality kn)
+ (match Environ.constant_opt_value_in env kn with
+ | None -> raise (ParameterWithoutEquality (fst kn))
| Some c -> aux (applist (c,a)))
+ | Proj _ -> raise (EqUnknown "Proj")
| Construct _ -> raise (EqUnknown "Construct")
| Case _ -> raise (EqUnknown "Case")
| CoFix _ -> raise (EqUnknown "CoFix")
@@ -221,35 +224,37 @@ let build_beq_scheme kn =
List.fold_left (fun a b -> mkLambda(Anonymous,b,a))
(mkLambda (Anonymous,
mkFullInd ind (n+3+(List.length rettyp_l)+nb_ind-1),
- bb))
+ (Lazy.force bb)))
(List.rev rettyp_l) in
(* make_one_eq *)
(* do the [| C1 ... => match Y with ... end
...
Cn => match Y with ... end |] part *)
- let ci = make_case_info env ind MatchStyle in
+ 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
let constrsi = constrs (3+nparrec) in
let n = Array.length constrsi in
- let ar = Array.create n ff in
+ let ar = Array.make n (Lazy.force ff) in
+ let eff = ref Declareops.no_seff in
for i=0 to n-1 do
let nb_cstr_args = List.length constrsi.(i).cs_args in
- let ar2 = Array.create n ff in
+ let ar2 = Array.make n (Lazy.force ff) in
let constrsj = constrs (3+nparrec+nb_cstr_args) in
for j=0 to n-1 do
- if (i=j) then
+ if Int.equal i j then
ar2.(j) <- let cc = (match nb_cstr_args with
- | 0 -> tt
- | _ -> let eqs = Array.make nb_cstr_args tt in
+ | 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 eqA = compute_A_equality rel_list
+ let eqA, eff' = compute_A_equality rel_list
nparrec
(nparrec+3+2*nb_cstr_args)
(nb_cstr_args+ndx+1)
cc
in
+ eff := Declareops.union_side_effects eff' !eff;
Array.set eqs ndx
(mkApp (eqA,
[|mkRel (ndx+1+nb_cstr_args);mkRel (ndx+1)|]
@@ -265,33 +270,40 @@ let build_beq_scheme kn =
(constrsj.(j).cs_args)
)
else ar2.(j) <- (List.fold_left (fun a (p,q,r) ->
- mkLambda (p,r,a)) ff (constrsj.(j).cs_args) )
+ mkLambda (p,r,a)) (Lazy.force ff) (constrsj.(j).cs_args) )
done;
ar.(i) <- (List.fold_left (fun a (p,q,r) -> mkLambda (p,r,a))
(mkCase (ci,do_predicate rel_list nb_cstr_args,
- mkVar (id_of_string "Y") ,ar2))
+ mkVar (Id.of_string "Y") ,ar2))
(constrsi.(i).cs_args))
done;
- mkNamedLambda (id_of_string "X") (mkFullInd ind (nb_ind-1+1)) (
- mkNamedLambda (id_of_string "Y") (mkFullInd ind (nb_ind-1+2)) (
- mkCase (ci, do_predicate rel_list 0,mkVar (id_of_string "X"),ar)))
+ mkNamedLambda (Id.of_string "X") (mkFullInd ind (nb_ind-1+1)) (
+ mkNamedLambda (Id.of_string "Y") (mkFullInd ind (nb_ind-1+2)) (
+ mkCase (ci, do_predicate rel_list 0,mkVar (Id.of_string "X"),ar))),
+ !eff
in (* build_beq_scheme *)
let names = Array.make nb_ind Anonymous and
types = Array.make nb_ind mkSet and
cores = Array.make nb_ind mkSet in
+ let eff = ref Declareops.no_seff in
+ let u = Univ.Instance.empty in
for i=0 to (nb_ind-1) do
- names.(i) <- Name (id_of_string (rec_name i));
- types.(i) <- mkArrow (mkFullInd (kn,i) 0)
- (mkArrow (mkFullInd (kn,i) 1) bb);
- cores.(i) <- make_one_eq i
+ names.(i) <- Name (Id.of_string (rec_name i));
+ types.(i) <- mkArrow (mkFullInd ((kn,i),u) 0)
+ (mkArrow (mkFullInd ((kn,i),u) 1) (Lazy.force bb));
+ let c, eff' = make_one_eq i in
+ cores.(i) <- c;
+ eff := Declareops.union_side_effects eff' !eff
done;
- Array.init nb_ind (fun i ->
+ (Array.init nb_ind (fun i ->
let kelim = Inductive.elim_sorts (mib,mib.mind_packets.(i)) in
- if not (List.mem InSet kelim) then
+ if not (Sorts.List.mem InSet kelim) then
raise (NonSingletonProp (kn,i));
let fix = mkFix (((Array.make nb_ind 0),i),(names,types,cores)) in
- create_input fix)
+ create_input fix),
+ Evd.empty_evar_universe_context (* FIXME *)),
+ !eff
let beq_scheme_kind = declare_mutual_scheme_object "_beq" build_beq_scheme
@@ -304,9 +316,8 @@ let destruct_ind c =
try let u,v = destApp c in
let indc = destInd u in
indc,v
- with e when Errors.noncritical e ->
- let indc = destInd c in
- indc,[||]
+ with DestKO -> let indc = destInd c in
+ indc,[||]
(*
In the following, avoid is the list of names to avoid.
@@ -317,7 +328,7 @@ let destruct_ind c =
so from Ai we can find the the correct eq_Ai bl_ai or lb_ai
*)
(* used in the leib -> bool side*)
-let do_replace_lb lb_scheme_key aavoid narg gls p q =
+let do_replace_lb lb_scheme_key aavoid narg p q =
let avoid = Array.of_list aavoid in
let do_arg v offset =
try
@@ -325,47 +336,56 @@ let do_replace_lb lb_scheme_key aavoid narg gls p q =
let s = destVar v in
let n = Array.length avoid in
let rec find i =
- if avoid.(n-i) = s then avoid.(n-i-x)
+ if Id.equal avoid.(n-i) s then avoid.(n-i-x)
else (if i<n then find (i+1)
- else error ("Var "^(string_of_id s)^" seems unknown.")
+ else error ("Var "^(Id.to_string s)^" seems unknown.")
)
in mkVar (find 1)
with e when Errors.noncritical e ->
(* if this happen then the args have to be already declared as a
- Parameter*)
+ Parameter*)
(
- let mp,dir,lbl = repr_con (destConst v) in
+ let mp,dir,lbl = repr_con (fst (destConst v)) in
mkConst (make_con mp dir (mk_label (
- if offset=1 then ("eq_"^(string_of_label lbl))
- else ((string_of_label lbl)^"_lb")
+ if Int.equal offset 1 then ("eq_"^(Label.to_string lbl))
+ else ((Label.to_string lbl)^"_lb")
)))
)
in
- let type_of_pq = pf_type_of gls p in
+ Proofview.Goal.nf_enter begin fun gl ->
+ let type_of_pq = Tacmach.New.of_old (fun gl -> pf_type_of gl p) gl in
let u,v = destruct_ind type_of_pq
in let lb_type_of_p =
- try mkConst (find_scheme lb_scheme_key u)
+ try
+ let c, eff = find_scheme lb_scheme_key (out_punivs u) (*FIXME*) in
+ Proofview.tclUNIT (mkConst c, eff)
with Not_found ->
(* spiwack: the format of this error message should probably
be improved. *)
- let err_msg = string_of_ppcmds
- (str "Leibniz->boolean:" ++
- str "You have to declare the" ++
- str "decidability over " ++
- Printer.pr_constr type_of_pq ++
- str " first.")
+ let err_msg =
+ (str "Leibniz->boolean:" ++
+ str "You have to declare the" ++
+ str "decidability over " ++
+ Printer.pr_constr type_of_pq ++
+ str " first.")
in
- error err_msg
- in let lb_args = Array.append (Array.append
+ Proofview.tclZERO (Errors.UserError("",err_msg))
+ in
+ lb_type_of_p >>= fun (lb_type_of_p,eff) ->
+ let lb_args = Array.append (Array.append
(Array.map (fun x -> x) v)
(Array.map (fun x -> do_arg x 1) v))
(Array.map (fun x -> do_arg x 2) v)
- in let app = if lb_args = [||]
+ in let app = if Array.equal eq_constr lb_args [||]
then lb_type_of_p else mkApp (lb_type_of_p,lb_args)
- in [Equality.replace p q ; apply app ; Auto.default_auto]
+ in
+ Tacticals.New.tclTHENLIST [
+ Proofview.tclEFFECTS eff;
+ Equality.replace p q ; apply app ; Auto.default_auto]
+ end
(* used in the bool -> leib side *)
-let do_replace_bl bl_scheme_key ind gls aavoid narg lft rgt =
+let do_replace_bl bl_scheme_key (ind,u as indu) aavoid narg lft rgt =
let avoid = Array.of_list aavoid in
let do_arg v offset =
try
@@ -373,36 +393,40 @@ let do_replace_bl bl_scheme_key ind gls aavoid narg lft rgt =
let s = destVar v in
let n = Array.length avoid in
let rec find i =
- if avoid.(n-i) = s then avoid.(n-i-x)
+ if Id.equal avoid.(n-i) s then avoid.(n-i-x)
else (if i<n then find (i+1)
- else error ("Var "^(string_of_id s)^" seems unknown.")
+ else error ("Var "^(Id.to_string s)^" seems unknown.")
)
in mkVar (find 1)
with e when Errors.noncritical e ->
(* if this happen then the args have to be already declared as a
Parameter*)
(
- let mp,dir,lbl = repr_con (destConst v) in
+ let mp,dir,lbl = repr_con (fst (destConst v)) in
mkConst (make_con mp dir (mk_label (
- if offset=1 then ("eq_"^(string_of_label lbl))
- else ((string_of_label lbl)^"_bl")
+ if Int.equal offset 1 then ("eq_"^(Label.to_string lbl))
+ else ((Label.to_string lbl)^"_bl")
)))
)
in
let rec aux l1 l2 =
match (l1,l2) with
- | (t1::q1,t2::q2) -> let tt1 = pf_type_of gls t1 in
- if t1=t2 then aux q1 q2
+ | (t1::q1,t2::q2) ->
+ Proofview.Goal.enter begin fun gl ->
+ let tt1 = Tacmach.New.pf_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 -> ind,[||]
- in if u = ind
- then (Equality.replace t1 t2)::(Auto.default_auto)::(aux q1 q2)
+ with e when Errors.noncritical e -> indu,[||]
+ in if eq_ind (fst u) ind
+ then Tacticals.New.tclTHENLIST [Equality.replace t1 t2; Auto.default_auto ; aux q1 q2 ]
else (
- let bl_t1 =
- try mkConst (find_scheme bl_scheme_key u)
+ let bl_t1, eff =
+ try
+ let c, eff = find_scheme bl_scheme_key (out_punivs u) (*FIXME*) in
+ mkConst c, eff
with Not_found ->
(* spiwack: the format of this error message should probably
be improved. *)
@@ -420,33 +444,41 @@ let do_replace_bl bl_scheme_key ind gls aavoid narg lft rgt =
(Array.map (fun x -> do_arg x 1) v))
(Array.map (fun x -> do_arg x 2) v )
in
- let app = if bl_args = [||]
+ let app = if Array.equal eq_constr bl_args [||]
then bl_t1 else mkApp (bl_t1,bl_args)
in
- (Equality.replace_by t1 t2
- (tclTHEN (apply app) (Auto.default_auto)))::(aux q1 q2)
+ Tacticals.New.tclTHENLIST [
+ Proofview.tclEFFECTS eff;
+ Equality.replace_by t1 t2
+ (Tacticals.New.tclTHEN (apply app) (Auto.default_auto)) ;
+ aux q1 q2 ]
)
)
- | ([],[]) -> []
- | _ -> error "Both side of the equality must have the same arity."
- in
- let (ind1,ca1) =
- try destApp lft with e when Errors.noncritical e -> error "replace failed."
- and (ind2,ca2) =
- try destApp rgt with e when Errors.noncritical e -> error "replace failed."
+ end
+ | ([],[]) -> Proofview.tclUNIT ()
+ | _ -> Proofview.tclZERO (UserError ("" , str"Both side of the equality must have the same arity."))
in
- let (sp1,i1) =
- try destInd ind1 with e when Errors.noncritical e ->
- try fst (destConstruct ind1) with e when Errors.noncritical e ->
- error "The expected type is an inductive one."
- and (sp2,i2) =
- try destInd ind2 with e when Errors.noncritical e ->
- try fst (destConstruct ind2) with e when Errors.noncritical e ->
- error "The expected type is an inductive one."
- in
- if (sp1 <> sp2) || (i1 <> i2)
- then (error "Eq should be on the same type")
- else (aux (Array.to_list ca1) (Array.to_list ca2))
+ begin try Proofview.tclUNIT (destApp lft)
+ with DestKO -> Proofview.tclZERO (UserError ("" , str"replace failed."))
+ end >>= fun (ind1,ca1) ->
+ begin try Proofview.tclUNIT (destApp rgt)
+ with DestKO -> Proofview.tclZERO (UserError ("" , str"replace failed."))
+ end >>= fun (ind2,ca2) ->
+ begin try Proofview.tclUNIT (out_punivs (destInd ind1))
+ with DestKO ->
+ begin try Proofview.tclUNIT (fst (fst (destConstruct ind1)))
+ with DestKO -> Proofview.tclZERO (UserError ("" , str"The expected type is an inductive one."))
+ end
+ end >>= fun (sp1,i1) ->
+ begin try Proofview.tclUNIT (out_punivs (destInd ind2))
+ with DestKO ->
+ begin try Proofview.tclUNIT (fst (fst (destConstruct ind2)))
+ with DestKO -> Proofview.tclZERO (UserError ("" , str"The expected type is an inductive one."))
+ end
+ end >>= fun (sp2,i2) ->
+ if not (eq_mind sp1 sp2) || not (Int.equal i1 i2)
+ then Proofview.tclZERO (UserError ("" , str"Eq should be on the same type"))
+ else aux (Array.to_list ca1) (Array.to_list ca2)
(*
create, from a list of ids [i1,i2,...,in] the list
@@ -454,11 +486,11 @@ let do_replace_bl bl_scheme_key ind gls aavoid narg lft rgt =
*)
let list_id l = List.fold_left ( fun a (n,_,t) -> let s' =
match n with
- Name s -> string_of_id s
+ Name s -> Id.to_string s
| Anonymous -> "A" in
- (id_of_string s',id_of_string ("eq_"^s'),
- id_of_string (s'^"_bl"),
- id_of_string (s'^"_lb"))
+ (Id.of_string s',Id.of_string ("eq_"^s'),
+ Id.of_string (s'^"_bl"),
+ Id.of_string (s'^"_lb"))
::a
) [] l
(*
@@ -468,52 +500,54 @@ let eqI ind l =
let list_id = list_id l in
let eA = Array.of_list((List.map (fun (s,_,_,_) -> mkVar s) list_id)@
(List.map (fun (_,seq,_,_)-> mkVar seq) list_id ))
- and e = try mkConst (find_scheme beq_scheme_kind ind) with
- Not_found -> error
+ and e, eff =
+ try let c, eff = find_scheme beq_scheme_kind ind in mkConst c, eff
+ with Not_found -> error
("The boolean equality on "^(string_of_mind (fst ind))^" is needed.");
- in (if eA = [||] then e else mkApp(e,eA))
+ in (if Array.equal eq_constr eA [||] then e else mkApp(e,eA)), eff
(**********************************************************************)
(* Boolean->Leibniz *)
let compute_bl_goal ind lnamesparrec nparrec =
- let eqI = eqI ind lnamesparrec in
+ let eqI, eff = eqI ind lnamesparrec in
let list_id = list_id lnamesparrec in
let create_input c =
- let x = id_of_string "x" and
- y = id_of_string "y" in
+ let x = Id.of_string "x" and
+ y = Id.of_string "y" in
let bl_typ = List.map (fun (s,seq,_,_) ->
mkNamedProd x (mkVar s) (
mkNamedProd y (mkVar s) (
mkArrow
- ( mkApp(eq,[|bb;mkApp(mkVar seq,[|mkVar x;mkVar y|]);tt|]))
- ( mkApp(eq,[|mkVar s;mkVar x;mkVar y|]))
+ ( mkApp(Lazy.force eq,[|(Lazy.force bb);mkApp(mkVar seq,[|mkVar x;mkVar y|]);(Lazy.force tt)|]))
+ ( mkApp(Lazy.force eq,[|mkVar s;mkVar x;mkVar y|]))
))
) list_id in
let bl_input = List.fold_left2 ( fun a (s,_,sbl,_) b ->
mkNamedProd sbl b a
) c (List.rev list_id) (List.rev bl_typ) in
let eqs_typ = List.map (fun (s,_,_,_) ->
- mkProd(Anonymous,mkVar s,mkProd(Anonymous,mkVar s,bb))
+ mkProd(Anonymous,mkVar s,mkProd(Anonymous,mkVar s,(Lazy.force bb)))
) list_id in
let eq_input = List.fold_left2 ( fun a (s,seq,_,_) b ->
mkNamedProd seq b a
) bl_input (List.rev list_id) (List.rev eqs_typ) in
List.fold_left (fun a (n,_,t) -> mkNamedProd
- (match n with Name s -> s | Anonymous -> id_of_string "A")
+ (match n with Name s -> s | Anonymous -> Id.of_string "A")
t a) eq_input lnamesparrec
in
- let n = id_of_string "x" and
- m = id_of_string "y" in
+ let n = Id.of_string "x" and
+ m = Id.of_string "y" in
+ let u = Univ.Instance.empty in
create_input (
- mkNamedProd n (mkFullInd ind nparrec) (
- mkNamedProd m (mkFullInd ind (nparrec+1)) (
+ mkNamedProd n (mkFullInd (ind,u) nparrec) (
+ mkNamedProd m (mkFullInd (ind,u) (nparrec+1)) (
mkArrow
- (mkApp(eq,[|bb;mkApp(eqI,[|mkVar n;mkVar m|]);tt|]))
- (mkApp(eq,[|mkFullInd ind (nparrec+3);mkVar n;mkVar m|]))
- )))
+ (mkApp(Lazy.force eq,[|(Lazy.force bb);mkApp(eqI,[|mkVar n;mkVar m|]);(Lazy.force tt)|]))
+ (mkApp(Lazy.force eq,[|mkFullInd (ind,u) (nparrec+3);mkVar n;mkVar m|]))
+ ))), eff
-let compute_bl_tact bl_scheme_key ind lnamesparrec nparrec gsig =
+let compute_bl_tact bl_scheme_key ind lnamesparrec nparrec =
let list_id = list_id lnamesparrec in
let avoid = ref [] in
let first_intros =
@@ -521,83 +555,88 @@ let compute_bl_tact bl_scheme_key ind lnamesparrec nparrec gsig =
( List.map (fun (_,seq,_,_ ) -> seq) list_id ) @
( List.map (fun (_,_,sbl,_ ) -> sbl) list_id )
in
- let fresh_first_intros = List.map ( fun s ->
- let fresh = fresh_id (!avoid) s gsig in
- avoid := fresh::(!avoid); fresh ) first_intros in
- let freshn = fresh_id (!avoid) (id_of_string "x") gsig in
- let freshm = avoid := freshn::(!avoid);
- fresh_id (!avoid) (id_of_string "y") gsig in
- let freshz = avoid := freshm::(!avoid);
- fresh_id (!avoid) (id_of_string "Z") gsig in
+ let fresh_id s gl =
+ Tacmach.New.of_old begin fun gsig ->
+ let fresh = fresh_id (!avoid) s gsig in
+ avoid := fresh::(!avoid); fresh
+ end gl
+ in
+ Proofview.Goal.nf_enter begin fun gl ->
+ let fresh_first_intros = List.map (fun id -> fresh_id id gl) first_intros in
+ let freshn = fresh_id (Id.of_string "x") gl in
+ let freshm = fresh_id (Id.of_string "y") gl in
+ let freshz = fresh_id (Id.of_string "Z") gl in
(* try with *)
- avoid := freshz::(!avoid);
- tclTHENSEQ [ intros_using fresh_first_intros;
+ Tacticals.New.tclTHENLIST [ intros_using fresh_first_intros;
intro_using freshn ;
induct_on (mkVar freshn);
intro_using freshm;
destruct_on (mkVar freshm);
intro_using freshz;
intros;
- tclTRY (
- tclORELSE reflexivity (Equality.discr_tac false None)
+ Tacticals.New.tclTRY (
+ Tacticals.New.tclORELSE reflexivity (Equality.discr_tac false None)
);
- simpl_in_hyp (freshz,InHyp);
+ Proofview.V82.tactic (simpl_in_hyp (freshz,Locus.InHyp));
(*
repeat ( apply andb_prop in z;let z1:= fresh "Z" in destruct z as [z1 z]).
*)
- tclREPEAT (
- tclTHENSEQ [
- simple_apply_in freshz (andb_prop());
- fun gl ->
- let fresht = fresh_id (!avoid) (id_of_string "Z") gsig
- in
- avoid := fresht::(!avoid);
- (new_destruct false [Tacexpr.ElimOnConstr
- (Evd.empty,((mkVar freshz,Glob_term.NoBindings)))]
- None
- (None, Some (dl,Genarg.IntroOrAndPattern [[
- dl,Genarg.IntroIdentifier fresht;
- dl,Genarg.IntroIdentifier freshz]])) None) gl
+ Tacticals.New.tclREPEAT (
+ Tacticals.New.tclTHENLIST [
+ Simple.apply_in freshz (andb_prop());
+ Proofview.Goal.nf_enter begin fun gl ->
+ let fresht = fresh_id (Id.of_string "Z") gl in
+ (destruct_on_as (mkVar freshz)
+ [[dl,IntroNaming (IntroIdentifier fresht);
+ dl,IntroNaming (IntroIdentifier freshz)]])
+ end
]);
(*
Ci a1 ... an = Ci b1 ... bn
replace bi with ai; auto || replace bi with ai by apply typeofbi_prod ; auto
*)
- fun gls-> let gl = pf_concl gls in
- match (kind_of_term gl) with
- | App (c,ca) -> (
- match (kind_of_term c) with
- | Ind indeq ->
- if IndRef indeq = Coqlib.glob_eq
- then (
- tclTHENSEQ ((do_replace_bl bl_scheme_key ind gls
- (!avoid)
- nparrec (ca.(2))
- (ca.(1)))@[Auto.default_auto]) gls
- )
- else
- (error "Failure while solving Boolean->Leibniz.")
- | _ -> error "Failure while solving Boolean->Leibniz."
- )
- | _ -> error "Failure while solving Boolean->Leibniz."
-
- ] gsig
+ Proofview.Goal.nf_enter begin fun gls ->
+ let gl = Proofview.Goal.concl gls in
+ match (kind_of_term gl) with
+ | App (c,ca) -> (
+ match (kind_of_term c) with
+ | Ind (indeq, u) ->
+ if eq_gr (IndRef indeq) Coqlib.glob_eq
+ then
+ Tacticals.New.tclTHEN
+ (do_replace_bl bl_scheme_key ind
+ (!avoid)
+ nparrec (ca.(2))
+ (ca.(1)))
+ Auto.default_auto
+ else
+ Proofview.tclZERO (UserError ("",str"Failure while solving Boolean->Leibniz."))
+ | _ -> Proofview.tclZERO (UserError ("", str"Failure while solving Boolean->Leibniz."))
+ )
+ | _ -> Proofview.tclZERO (UserError ("", str"Failure while solving Boolean->Leibniz."))
+ end
+
+ ]
+ end
let bl_scheme_kind_aux = ref (fun _ -> failwith "Undefined")
let make_bl_scheme mind =
let mib = Global.lookup_mind mind in
- if Array.length mib.mind_packets <> 1 then
+ if not (Int.equal (Array.length mib.mind_packets) 1) then
errorlabstrm ""
(str "Automatic building of boolean->Leibniz lemmas not supported");
let ind = (mind,0) in
let nparams = mib.mind_nparams in
let nparrec = mib.mind_nparams_rec in
- let lnonparrec,lnamesparrec =
+ let lnonparrec,lnamesparrec = (* TODO subst *)
context_chop (nparams-nparrec) mib.mind_params_ctxt in
- [|Pfedit.build_by_tactic (Global.env())
- (compute_bl_goal ind lnamesparrec nparrec)
- (compute_bl_tact (!bl_scheme_kind_aux()) ind lnamesparrec nparrec)|]
+ let bl_goal, eff = compute_bl_goal ind lnamesparrec nparrec in
+ let ctx = Evd.empty_evar_universe_context (*FIXME univs *) in
+ let (ans, _, ctx) = Pfedit.build_by_tactic (Global.env()) ctx bl_goal
+ (compute_bl_tact (!bl_scheme_kind_aux()) (ind, Univ.Instance.empty) lnamesparrec nparrec)
+ in
+ ([|ans|], ctx), eff
let bl_scheme_kind = declare_mutual_scheme_object "_dec_bl" make_bl_scheme
@@ -608,10 +647,11 @@ let _ = bl_scheme_kind_aux := fun () -> bl_scheme_kind
let compute_lb_goal ind lnamesparrec nparrec =
let list_id = list_id lnamesparrec in
- let eqI = eqI ind lnamesparrec in
+ let eq = Lazy.force eq and tt = Lazy.force tt and bb = Lazy.force bb in
+ let eqI, eff = eqI ind lnamesparrec in
let create_input c =
- let x = id_of_string "x" and
- y = id_of_string "y" in
+ let x = Id.of_string "x" and
+ y = Id.of_string "y" in
let lb_typ = List.map (fun (s,seq,_,_) ->
mkNamedProd x (mkVar s) (
mkNamedProd y (mkVar s) (
@@ -630,20 +670,21 @@ let compute_lb_goal ind lnamesparrec nparrec =
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")
+ (match n with Name s -> s | Anonymous -> Id.of_string "A")
t a) eq_input lnamesparrec
in
- let n = id_of_string "x" and
- m = id_of_string "y" in
+ let n = Id.of_string "x" and
+ m = Id.of_string "y" in
+ let u = Univ.Instance.empty in
create_input (
- mkNamedProd n (mkFullInd ind nparrec) (
- mkNamedProd m (mkFullInd ind (nparrec+1)) (
+ mkNamedProd n (mkFullInd (ind,u) nparrec) (
+ mkNamedProd m (mkFullInd (ind,u) (nparrec+1)) (
mkArrow
- (mkApp(eq,[|mkFullInd ind (nparrec+2);mkVar n;mkVar m|]))
+ (mkApp(eq,[|mkFullInd (ind,u) (nparrec+2);mkVar n;mkVar m|]))
(mkApp(eq,[|bb;mkApp(eqI,[|mkVar n;mkVar m|]);tt|]))
- )))
+ ))), eff
-let compute_lb_tact lb_scheme_key ind lnamesparrec nparrec gsig =
+let compute_lb_tact lb_scheme_key ind lnamesparrec nparrec =
let list_id = list_id lnamesparrec in
let avoid = ref [] in
let first_intros =
@@ -651,55 +692,60 @@ let compute_lb_tact lb_scheme_key ind lnamesparrec nparrec gsig =
( List.map (fun (_,seq,_,_) -> seq) list_id ) @
( List.map (fun (_,_,_,slb) -> slb) list_id )
in
- let fresh_first_intros = List.map ( fun s ->
- let fresh = fresh_id (!avoid) s gsig in
- avoid := fresh::(!avoid); fresh ) first_intros in
- let freshn = fresh_id (!avoid) (id_of_string "x") gsig in
- let freshm = avoid := freshn::(!avoid);
- fresh_id (!avoid) (id_of_string "y") gsig in
- let freshz = avoid := freshm::(!avoid);
- fresh_id (!avoid) (id_of_string "Z") gsig in
+ let fresh_id s gl =
+ Tacmach.New.of_old begin fun gsig ->
+ let fresh = fresh_id (!avoid) s gsig in
+ avoid := fresh::(!avoid); fresh
+ end gl
+ in
+ Proofview.Goal.nf_enter begin fun gl ->
+ let fresh_first_intros = List.map (fun id -> fresh_id id gl) first_intros in
+ let freshn = fresh_id (Id.of_string "x") gl in
+ let freshm = fresh_id (Id.of_string "y") gl in
+ let freshz = fresh_id (Id.of_string "Z") gl in
(* try with *)
- avoid := freshz::(!avoid);
- tclTHENSEQ [ intros_using fresh_first_intros;
+ Tacticals.New.tclTHENLIST [ intros_using fresh_first_intros;
intro_using freshn ;
induct_on (mkVar freshn);
intro_using freshm;
destruct_on (mkVar freshm);
intro_using freshz;
intros;
- tclTRY (
- tclORELSE reflexivity (Equality.discr_tac false None)
+ Tacticals.New.tclTRY (
+ Tacticals.New.tclORELSE reflexivity (Equality.discr_tac false None)
);
- Equality.inj [] false (mkVar freshz,Glob_term.NoBindings);
- intros; simpl_in_concl;
+ Equality.inj None false None (mkVar freshz,NoBindings);
+ intros; (Proofview.V82.tactic simpl_in_concl);
Auto.default_auto;
- tclREPEAT (
- tclTHENSEQ [apply (andb_true_intro());
+ Tacticals.New.tclREPEAT (
+ Tacticals.New.tclTHENLIST [apply (andb_true_intro());
simplest_split ;Auto.default_auto ]
);
- fun gls -> let gl = pf_concl gls in
- (* assume the goal to be eq (eq_type ...) = true *)
+ Proofview.Goal.nf_enter begin fun gls ->
+ let gl = Proofview.Goal.concl gls in
+ (* assume the goal to be eq (eq_type ...) = true *)
match (kind_of_term gl) with
- | App(c,ca) -> (match (kind_of_term ca.(1)) with
- | App(c',ca') ->
- let n = Array.length ca' in
- tclTHENSEQ (do_replace_lb lb_scheme_key
- (!avoid)
- nparrec gls
- ca'.(n-2) ca'.(n-1)) gls
- | _ -> error
- "Failure while solving Leibniz->Boolean."
- )
- | _ -> error
- "Failure while solving Leibniz->Boolean."
- ] gsig
+ | App(c,ca) -> (match (kind_of_term ca.(1)) with
+ | App(c',ca') ->
+ let n = Array.length ca' in
+ do_replace_lb lb_scheme_key
+ (!avoid)
+ nparrec
+ ca'.(n-2) ca'.(n-1)
+ | _ ->
+ Proofview.tclZERO (UserError ("",str"Failure while solving Leibniz->Boolean."))
+ )
+ | _ ->
+ Proofview.tclZERO (UserError ("",str"Failure while solving Leibniz->Boolean."))
+ end
+ ]
+ end
let lb_scheme_kind_aux = ref (fun () -> failwith "Undefined")
let make_lb_scheme mind =
let mib = Global.lookup_mind mind in
- if Array.length mib.mind_packets <> 1 then
+ if not (Int.equal (Array.length mib.mind_packets) 1) then
errorlabstrm ""
(str "Automatic building of Leibniz->boolean lemmas not supported");
let ind = (mind,0) in
@@ -707,9 +753,12 @@ let make_lb_scheme mind =
let nparrec = mib.mind_nparams_rec in
let lnonparrec,lnamesparrec =
context_chop (nparams-nparrec) mib.mind_params_ctxt in
- [|Pfedit.build_by_tactic (Global.env())
- (compute_lb_goal ind lnamesparrec nparrec)
- (compute_lb_tact (!lb_scheme_kind_aux()) ind lnamesparrec nparrec)|]
+ let lb_goal, eff = compute_lb_goal ind lnamesparrec nparrec in
+ let ctx = Evd.empty_evar_universe_context in
+ let (ans, _, ctx) = Pfedit.build_by_tactic (Global.env()) ctx lb_goal
+ (compute_lb_tact (!lb_scheme_kind_aux()) ind lnamesparrec nparrec)
+ in
+ ([|ans|], ctx (* FIXME *)), eff
let lb_scheme_kind = declare_mutual_scheme_object "_dec_lb" make_lb_scheme
@@ -725,10 +774,11 @@ let check_not_is_defined () =
(* {n=m}+{n<>m} part *)
let compute_dec_goal ind lnamesparrec nparrec =
check_not_is_defined ();
+ let eq = Lazy.force eq and tt = Lazy.force tt and bb = Lazy.force bb in
let list_id = list_id lnamesparrec in
let create_input c =
- let x = id_of_string "x" and
- y = id_of_string "y" in
+ let x = Id.of_string "x" and
+ y = Id.of_string "y" in
let lb_typ = List.map (fun (s,seq,_,_) ->
mkNamedProd x (mkVar s) (
mkNamedProd y (mkVar s) (
@@ -760,11 +810,11 @@ let compute_dec_goal ind lnamesparrec nparrec =
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")
+ (match n with Name s -> s | Anonymous -> Id.of_string "A")
t a) eq_input lnamesparrec
in
- let n = id_of_string "x" and
- m = id_of_string "y" in
+ let n = Id.of_string "x" and
+ m = Id.of_string "y" in
let eqnm = mkApp(eq,[|mkFullInd ind (2*nparrec+2);mkVar n;mkVar m|]) in
create_input (
mkNamedProd n (mkFullInd ind (2*nparrec)) (
@@ -774,9 +824,11 @@ let compute_dec_goal ind lnamesparrec nparrec =
)
)
-let compute_dec_tact ind lnamesparrec nparrec gsig =
+let compute_dec_tact ind lnamesparrec nparrec =
+ let eq = Lazy.force eq and tt = Lazy.force tt
+ and ff = Lazy.force ff and bb = Lazy.force bb in
let list_id = list_id lnamesparrec in
- let eqI = eqI ind lnamesparrec in
+ let eqI, eff = eqI ind lnamesparrec in
let avoid = ref [] in
let eqtrue x = mkApp(eq,[|bb;x;tt|]) in
let eqfalse x = mkApp(eq,[|bb;x;ff|]) in
@@ -786,85 +838,101 @@ let compute_dec_tact ind lnamesparrec nparrec gsig =
( List.map (fun (_,_,sbl,_) -> sbl) list_id ) @
( List.map (fun (_,_,_,slb) -> slb) list_id )
in
- let fresh_first_intros = List.map ( fun s ->
- let fresh = fresh_id (!avoid) s gsig in
- avoid := fresh::(!avoid); fresh ) first_intros in
- let freshn = fresh_id (!avoid) (id_of_string "x") gsig in
- let freshm = avoid := freshn::(!avoid);
- fresh_id (!avoid) (id_of_string "y") gsig in
- let freshH = avoid := freshm::(!avoid);
- fresh_id (!avoid) (id_of_string "H") gsig in
+ let fresh_id s gl =
+ Tacmach.New.of_old begin fun gsig ->
+ let fresh = fresh_id (!avoid) s gsig in
+ avoid := fresh::(!avoid); fresh
+ end gl
+ in
+ Proofview.Goal.nf_enter begin fun gl ->
+ let fresh_first_intros = List.map (fun id -> fresh_id id gl) first_intros in
+ let freshn = fresh_id (Id.of_string "x") gl in
+ let freshm = fresh_id (Id.of_string "y") gl in
+ let freshH = fresh_id (Id.of_string "H") gl in
let eqbnm = mkApp(eqI,[|mkVar freshn;mkVar freshm|]) in
- avoid := freshH::(!avoid);
let arfresh = Array.of_list fresh_first_intros in
let xargs = Array.sub arfresh 0 (2*nparrec) in
- let blI = try mkConst (find_scheme bl_scheme_kind ind) with
- Not_found -> error (
- "Error during the decidability part, boolean to leibniz"^
- " equality is required.")
- in
- let lbI = try mkConst (find_scheme lb_scheme_kind ind) with
- Not_found -> error (
- "Error during the decidability part, leibniz to boolean"^
- " equality is required.")
- in
- tclTHENSEQ [
+ begin try
+ let c, eff = find_scheme bl_scheme_kind ind in
+ Proofview.tclUNIT (mkConst c,eff) with
+ Not_found ->
+ Proofview.tclZERO (UserError ("",str"Error during the decidability part, boolean to leibniz"++
+ str" equality is required."))
+ end >>= fun (blI,eff') ->
+ begin try
+ let c, eff = find_scheme lb_scheme_kind ind in
+ Proofview.tclUNIT (mkConst c,eff) with
+ Not_found ->
+ Proofview.tclZERO (UserError ("",str"Error during the decidability part, leibniz to boolean"++
+ str" equality is required."))
+ end >>= fun (lbI,eff'') ->
+ let eff = (Declareops.union_side_effects eff'' (Declareops.union_side_effects eff' eff)) in
+ Tacticals.New.tclTHENLIST [
+ Proofview.tclEFFECTS eff;
intros_using fresh_first_intros;
intros_using [freshn;freshm];
(*we do this so we don't have to prove the same goal twice *)
assert_by (Name freshH) (
mkApp(sumbool(),[|eqtrue eqbnm; eqfalse eqbnm|])
)
- (tclTHEN (destruct_on eqbnm) Auto.default_auto);
- (fun gsig ->
- let freshH2 = fresh_id (!avoid) (id_of_string "H") gsig in
- avoid := freshH2::(!avoid);
- tclTHENS (destruct_on_using (mkVar freshH) freshH2) [
+ (Tacticals.New.tclTHEN (destruct_on eqbnm) Auto.default_auto);
+
+ Proofview.Goal.nf_enter begin fun gl ->
+ let freshH2 = fresh_id (Id.of_string "H") gl in
+ Tacticals.New.tclTHENS (destruct_on_using (mkVar freshH) freshH2) [
(* left *)
- tclTHENSEQ [
+ Tacticals.New.tclTHENLIST [
simplest_left;
apply (mkApp(blI,Array.map(fun x->mkVar x) xargs));
Auto.default_auto
- ];
+ ]
+ ;
+
(*right *)
- (fun gsig ->
- let freshH3 = fresh_id (!avoid) (id_of_string "H") gsig in
- avoid := freshH3::(!avoid);
- tclTHENSEQ [
+ Proofview.Goal.nf_enter begin fun gl ->
+ let freshH3 = fresh_id (Id.of_string "H") gl in
+ Tacticals.New.tclTHENLIST [
simplest_right ;
- unfold_constr (Lazy.force Coqlib.coq_not_ref);
+ Proofview.V82.tactic (unfold_constr (Lazy.force Coqlib.coq_not_ref));
intro;
- Equality.subst_all;
+ Equality.subst_all ();
assert_by (Name freshH3)
(mkApp(eq,[|bb;mkApp(eqI,[|mkVar freshm;mkVar freshm|]);tt|]))
- (tclTHENSEQ [
+ (Tacticals.New.tclTHENLIST [
apply (mkApp(lbI,Array.map (fun x->mkVar x) xargs));
Auto.default_auto
]);
Equality.general_rewrite_bindings_in true
- all_occurrences true false
+ Locus.AllOccurrences true false
(List.hd !avoid)
((mkVar (List.hd (List.tl !avoid))),
- Glob_term.NoBindings
+ NoBindings
)
true;
Equality.discr_tac false None
- ] gsig)
- ] gsig)
- ] gsig
+ ]
+ end
+ ]
+ end
+ ]
+ end
let make_eq_decidability mind =
let mib = Global.lookup_mind mind in
- if Array.length mib.mind_packets <> 1 then
- anomaly "Decidability lemma for mutual inductive types not supported";
+ if not (Int.equal (Array.length mib.mind_packets) 1) then
+ anomaly (Pp.str "Decidability lemma for mutual inductive types not supported");
let ind = (mind,0) in
let nparams = mib.mind_nparams in
let nparrec = mib.mind_nparams_rec in
+ let u = Univ.Instance.empty in
+ let ctx = Evd.empty_evar_universe_context (* FIXME *)in
let lnonparrec,lnamesparrec =
context_chop (nparams-nparrec) mib.mind_params_ctxt in
- [|Pfedit.build_by_tactic (Global.env())
- (compute_dec_goal ind lnamesparrec nparrec)
- (compute_dec_tact ind lnamesparrec nparrec)|]
+ let (ans, _, ctx) = Pfedit.build_by_tactic (Global.env()) ctx
+ (compute_dec_goal (ind,u) lnamesparrec nparrec)
+ (compute_dec_tact ind lnamesparrec nparrec)
+ in
+ ([|ans|], ctx), Declareops.no_seff
let eq_dec_scheme_kind =
declare_mutual_scheme_object "_eq_dec" make_eq_decidability
diff --git a/toplevel/auto_ind_decl.mli b/toplevel/auto_ind_decl.mli
index 1cb8b767..5dbd5379 100644
--- a/toplevel/auto_ind_decl.mli
+++ b/toplevel/auto_ind_decl.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,10 +8,6 @@
open Term
open Names
-open Libnames
-open Mod_subst
-open Sign
-open Proof_type
open Ind_tables
(** This file is about the automatic generation of schemes about
@@ -30,17 +26,16 @@ exception ParameterWithoutEquality of constant
exception NonSingletonProp of inductive
val beq_scheme_kind : mutual scheme_kind
-val build_beq_scheme : mutual_inductive -> constr array
+val build_beq_scheme : mutual_scheme_object_function
(** {6 Build equivalence between boolean equality and Leibniz equality } *)
val lb_scheme_kind : mutual scheme_kind
-val make_lb_scheme : mutual_inductive -> constr array
-
+val make_lb_scheme : mutual_scheme_object_function
val bl_scheme_kind : mutual scheme_kind
-val make_bl_scheme : mutual_inductive -> constr array
+val make_bl_scheme : mutual_scheme_object_function
(** {6 Build decidability of equality } *)
val eq_dec_scheme_kind : mutual scheme_kind
-val make_eq_decidability : mutual_inductive -> constr array
+val make_eq_decidability : mutual_scheme_object_function
diff --git a/toplevel/autoinstance.ml b/toplevel/autoinstance.ml
deleted file mode 100644
index ff09f73d..00000000
--- a/toplevel/autoinstance.ml
+++ /dev/null
@@ -1,320 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(*i*)
-open Pp
-open Printer
-open Names
-open Term
-open Evd
-open Sign
-open Libnames
-(*i*)
-
-(*s
- * Automatic detection of (some) record instances
- *)
-
-(* Datatype for wannabe-instances: a signature is a typeclass along
- with the collection of evars corresponding to the parameters/fields
- of the class. Each evar can be uninstantiated (we're still looking
- for them) or defined (the instance for the field is fixed) *)
-type signature = global_reference * evar list * evar_map
-
-type instance_decl_function = global_reference -> rel_context -> constr list -> unit
-
-(*
- * Search algorithm
- *)
-
-let rec subst_evar evar def n c =
- match kind_of_term c with
- | Evar (e,_) when e=evar -> lift n def
- | _ -> map_constr_with_binders (fun n->n+1) (subst_evar evar def) n c
-
-let subst_evar_in_evm evar def evm =
- Evd.fold
- (fun ev evi acc ->
- let evar_body = match evi.evar_body with
- | Evd.Evar_empty -> Evd.Evar_empty
- | Evd.Evar_defined c -> Evd.Evar_defined (subst_evar evar def 0 c) in
- let evar_concl = subst_evar evar def 0 evi.evar_concl in
- Evd.add acc ev {evi with evar_body=evar_body; evar_concl=evar_concl}
- ) evm empty
-
-(* Tries to define ev by c in evd. Fails if ev := c1 and c1 /= c ev :
- * T1, c : T2 and T1 /= T2. Defines recursively all evars instantiated
- * by this definition. *)
-
-let rec safe_define evm ev c =
- if not (closedn (-1) c) then raise Termops.CannotFilter else
-(* msgnl(str"safe_define "++pr_evar_map evm++spc()++str" |- ?"++Util.pr_int ev++str" := "++pr_constr c);*)
- let evi = (Evd.find evm ev) in
- let define_subst evm sigma =
- Util.Intmap.fold
- ( fun ev (e,c) evm ->
- match kind_of_term c with Evar (i,_) when i=ev -> evm | _ ->
- safe_define evm ev (lift (-List.length e) c)
- ) sigma evm in
- match evi.evar_body with
- | Evd.Evar_defined def ->
- define_subst evm (Termops.filtering [] Reduction.CUMUL def c)
- | Evd.Evar_empty ->
- let t = Libtypes.reduce (Typing.type_of (Global.env()) evm c) in
- let u = Libtypes.reduce (evar_concl evi) in
- let evm = subst_evar_in_evm ev c evm in
- define_subst (Evd.define ev c evm) (Termops.filtering [] Reduction.CUMUL t u)
-
-let add_gen_ctx (cl,gen,evm) ctx : signature * constr list =
- let rec really_new_evar () =
- let ev = Evarutil.new_untyped_evar() in
- if Evd.is_evar evm ev then really_new_evar() else ev in
- let add_gen_evar (cl,gen,evm) ev ty : signature =
- let evm = Evd.add evm ev (Evd.make_evar Environ.empty_named_context_val ty) in
- (cl,ev::gen,evm) in
- let rec mksubst b = function
- | [] -> []
- | a::tl -> b::(mksubst (a::b) tl) in
- let evl = List.map (fun _ -> really_new_evar()) ctx in
- let evcl = List.map (fun i -> mkEvar (i,[||])) evl in
- let substl = List.rev (mksubst [] (evcl)) in
- let ctx = List.map2 (fun s t -> substnl s 0 t) substl ctx in
- let sign = List.fold_left2 add_gen_evar (cl,gen,evm) (List.rev evl) ctx in
- sign,evcl
-
-(* TODO : for full proof-irrelevance in the search, provide a real
- compare function for constr instead of Pervasive's one! *)
-module SubstSet : Set.S with type elt = Termops.subst
- = Set.Make (struct type t = Termops.subst
- let compare = Util.Intmap.compare (Pervasives.compare)
- end)
-
-(* searches instatiations in the library for just one evar [ev] of a
- signature. [k] is called on each resulting signature *)
-let complete_evar (cl,gen,evm:signature) (ev,evi) (k:signature -> unit) =
- let ev_typ = Libtypes.reduce (evar_concl evi) in
- let sort_is_prop = is_Prop (Typing.type_of (Global.env()) evm (evar_concl evi)) in
-(* msgnl(str"cherche "++pr_constr ev_typ++str" pour "++Util.pr_int ev);*)
- let substs = ref SubstSet.empty in
- try List.iter
- ( fun (gr,(pat,_),s) ->
- let (_,genl,_) = Termops.decompose_prod_letin pat in
- let genl = List.map (fun (_,_,t) -> t) genl in
- let ((cl,gen,evm),argl) = add_gen_ctx (cl,gen,evm) genl in
- let def = applistc (Libnames.constr_of_global gr) argl in
-(* msgnl(str"essayons ?"++Util.pr_int ev++spc()++str":="++spc()
- ++pr_constr def++spc()++str":"++spc()++pr_constr (Global.type_of_global gr)*)
- (*++spc()++str"dans"++spc()++pr_evar_map evm++spc());*)
- try
- let evm = safe_define evm ev def in
- k (cl,gen,evm);
- if sort_is_prop && SubstSet.mem s !substs then raise Exit;
- substs := SubstSet.add s !substs
- with Termops.CannotFilter -> ()
- ) (Libtypes.search_concl ev_typ)
- with Exit -> ()
-
-let evm_fold_rev f evm acc =
- let l = Evd.fold (fun ev evi acc -> (ev,evi)::acc) evm [] in
- List.fold_left (fun acc (ev,evi) -> f ev evi acc) acc l
-
-exception Continue of Evd.evar * Evd.evar_info
-
-(* searches matches for all the uninstantiated evars of evd in the
- context. For each totally instantiated evar_map found, apply
- k. *)
-let rec complete_signature (k:signature -> unit) (cl,gen,evm:signature) =
- try
- evm_fold_rev
- ( fun ev evi _ ->
- if not (is_defined evm ev) && not (List.mem ev gen) then
- raise (Continue (ev,evi))
- ) evm (); k (cl,gen,evm)
- with Continue (ev,evi) -> complete_evar (cl,gen,evm) (ev,evi) (complete_signature k)
-
-(* define all permutations of the evars to evd and call k on the
- resulting evd *)
-let complete_with_evars_permut (cl,gen,evm:signature) evl c (k:signature -> unit) : unit =
- let rec aux evm = List.iter
- ( fun (ctx,ev) ->
- let tyl = List.map (fun (_,_,t) -> t) ctx in
- let ((cl,gen,evm),argl) = add_gen_ctx (cl,gen,evm) tyl in
- let def = applistc c argl in
-(* msgnl(str"trouvé def ?"++Util.pr_int ev++str" := "++pr_constr def++str " dans "++pr_evar_map evm);*)
- try
- if not (Evd.is_defined evm ev) then
- let evm = safe_define evm ev def in
- aux evm; k (cl,gen,evm)
- with Termops.CannotFilter -> ()
- ) evl in
- aux evm
-
-let new_inst_no =
- let cnt = ref 0 in
- fun () -> incr cnt; string_of_int !cnt
-
-let make_instance_ident gr =
- Nameops.add_suffix (Nametab.basename_of_global gr) ("_autoinstance_"^new_inst_no())
-
-let new_instance_message ident typ def =
- Flags.if_verbose
- msgnl (str"new instance"++spc()
- ++Nameops.pr_id ident++spc()++str":"++spc()
- ++pr_constr typ++spc()++str":="++spc()
- ++pr_constr def)
-
-open Entries
-
-let rec deep_refresh_universes c =
- match kind_of_term c with
- | Sort (Type _) -> Termops.new_Type()
- | _ -> map_constr deep_refresh_universes c
-
-let declare_record_instance gr ctx params =
- let ident = make_instance_ident gr in
- let def = it_mkLambda_or_LetIn (applistc (constr_of_global gr) params) ctx in
- let def = deep_refresh_universes def in
- let ce = { const_entry_body= def;
- const_entry_secctx = None;
- const_entry_type=None;
- const_entry_opaque=false } in
- let cst = Declare.declare_constant ident
- (DefinitionEntry ce,Decl_kinds.IsDefinition Decl_kinds.StructureComponent) in
- new_instance_message ident (Typeops.type_of_constant (Global.env()) cst) def
-
-let declare_class_instance gr ctx params =
- let ident = make_instance_ident gr in
- let cl = Typeclasses.class_info gr in
- let (def,typ) = Typeclasses.instance_constructor cl params in
- let (def,typ) = it_mkLambda_or_LetIn (Option.get def) ctx, it_mkProd_or_LetIn typ ctx in
- let def = deep_refresh_universes def in
- let typ = deep_refresh_universes typ in
- let ce = Entries.DefinitionEntry
- { const_entry_type = Some typ;
- const_entry_secctx = None;
- const_entry_body= def;
- const_entry_opaque=false } in
- try
- let cst = Declare.declare_constant ident
- (ce,Decl_kinds.IsDefinition Decl_kinds.Instance) in
- Typeclasses.add_instance (Typeclasses.new_instance cl (Some 100) true (ConstRef cst));
- new_instance_message ident typ def
- with e when Errors.noncritical e ->
- msgnl (str"Error defining instance := "++pr_constr def++str" : "++pr_constr typ++str" "++Errors.print e)
-
-let rec iter_under_prod (f:rel_context->constr->unit) (ctx:rel_context) t = f ctx t;
- match kind_of_term t with
- | Prod (n,t,c) -> iter_under_prod f ((n,None,t)::ctx) c
- | _ -> ()
-
-(* main search function: search for total instances containing gr, and
- apply k to each of them *)
-let complete_signature_with_def gr deftyp (k:instance_decl_function -> signature -> unit) : unit =
- let gr_c = Libnames.constr_of_global gr in
- let (smap:(Libnames.global_reference * Evd.evar_map,
- ('a * 'b * Term.constr) list * Evd.evar)
- Gmapl.t ref) = ref Gmapl.empty in
- iter_under_prod
- ( fun ctx typ ->
- List.iter
- (fun ((cl,ev,evm),_,_) ->
-(* msgnl(pr_global gr++str" : "++pr_constr typ++str" matche ?"++Util.pr_int ev++str " dans "++pr_evar_map evm);*)
- smap := Gmapl.add (cl,evm) (ctx,ev) !smap)
- (Recordops.methods_matching typ)
- ) [] deftyp;
- Gmapl.iter
- ( fun (cl,evm) evl ->
- let f = if Typeclasses.is_class cl then
- declare_class_instance else declare_record_instance in
- complete_with_evars_permut (cl,[],evm) evl gr_c
- (fun sign -> complete_signature (k f) sign)
- ) !smap
-
-(*
- * Interface with other parts: hooks & declaration
- *)
-
-
-let evar_definition evi = match evar_body evi with
- Evar_empty -> assert false | Evar_defined c -> c
-
-let gen_sort_topo l evm =
- let iter_evar f ev =
- let rec aux c = match kind_of_term c with
- Evar (e,_) -> f e
- | _ -> iter_constr aux c in
- aux (Evd.evar_concl (Evd.find evm ev));
- if Evd.is_defined evm ev then aux (evar_definition (Evd.find evm ev)) in
- let r = ref [] in
- let rec dfs ev = iter_evar dfs ev;
- if not(List.mem ev !r) then r := ev::!r in
- List.iter dfs l; List.rev !r
-
-(* register real typeclass instance given a totally defined evd *)
-let declare_instance (k:global_reference -> rel_context -> constr list -> unit)
- (cl,gen,evm:signature) =
- let evm = Evarutil.nf_evar_map evm in
- let gen = gen_sort_topo gen evm in
- let (evm,gen) = List.fold_right
- (fun ev (evm,gen) ->
- if Evd.is_defined evm ev
- then Evd.remove evm ev,gen
- else evm,(ev::gen))
- gen (evm,[]) in
-(* msgnl(str"instance complète : ["++Util.prlist_with_sep (fun _ -> str";") Util.pr_int gen++str"] : "++spc()++pr_evar_map evm);*)
- let ngen = List.length gen in
- let (_,ctx,evm) = List.fold_left
- ( fun (i,ctx,evm) ev ->
- let ctx = (Anonymous,None,lift (-i) (Evd.evar_concl(Evd.find evm ev)))::ctx in
- let evm = subst_evar_in_evm ev (mkRel i) (Evd.remove evm ev) in
- (i-1,ctx,evm)
- ) (ngen,[],evm) gen in
- let fields = List.rev (Evd.fold ( fun ev evi l -> evar_definition evi::l ) evm []) in
- k cl ctx fields
-
-let autoinstance_opt = ref true
-
-let search_declaration gr =
- if !autoinstance_opt &&
- not (Lib.is_modtype()) then
- let deftyp = Global.type_of_global gr in
- complete_signature_with_def gr deftyp declare_instance
-
-let search_record k cons sign =
- if !autoinstance_opt && not (Lib.is_modtype()) then
- complete_signature (declare_instance k) (cons,[],sign)
-
-(*
-let dh_key = Profile.declare_profile "declaration_hook"
-let ch_key = Profile.declare_profile "class_decl_hook"
-let declaration_hook = Profile.profile1 dh_key declaration_hook
-let class_decl_hook = Profile.profile1 ch_key class_decl_hook
-*)
-
-(*
- * Options and bookeeping
- *)
-
-let begin_autoinstance () =
- if not !autoinstance_opt then (
- autoinstance_opt := true;
- )
-
-let end_autoinstance () =
- if !autoinstance_opt then (
- autoinstance_opt := false;
- )
-
-let _ =
- Goptions.declare_bool_option
- { Goptions.optsync=true;
- Goptions.optdepr=false;
- Goptions.optkey=["Autoinstance"];
- Goptions.optname="automatic typeclass instance recognition";
- Goptions.optread=(fun () -> !autoinstance_opt);
- Goptions.optwrite=(fun b -> if b then begin_autoinstance() else end_autoinstance()) }
diff --git a/toplevel/autoinstance.mli b/toplevel/autoinstance.mli
deleted file mode 100644
index 6e26a784..00000000
--- a/toplevel/autoinstance.mli
+++ /dev/null
@@ -1,34 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-open Term
-open Libnames
-open Typeclasses
-open Names
-open Evd
-open Sign
-
-(** {6 Automatic detection of (some) record instances } *)
-
-(** What to do if we find an instance. Passed are : the reference
- * representing the record/class (definition or constructor) *)
-type instance_decl_function = global_reference -> rel_context -> constr list -> unit
-
-(** [search_declaration gr] Search in the library if the (new)
- * declaration gr can form an instance of a registered record/class *)
-val search_declaration : global_reference -> unit
-
-(** [search_record declf gr evm] Search the library for instances of
- the (new) record/class declaration [gr], and register them using
- [declf]. [evm] is the signature of the record (to avoid recomputing
- it) *)
-val search_record : instance_decl_function -> global_reference -> evar_map -> unit
-
-(** Instance declaration for both scenarios *)
-val declare_record_instance : instance_decl_function
-val declare_class_instance : instance_decl_function
diff --git a/toplevel/backtrack.ml b/toplevel/backtrack.ml
deleted file mode 100644
index eb100379..00000000
--- a/toplevel/backtrack.ml
+++ /dev/null
@@ -1,243 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-open Names
-open Vernacexpr
-
-(** Command history stack
-
- We maintain a stack of the past states of the system. Each
- successfully interpreted command adds an [info] element
- to this stack, storing what were the (label / current proof / ...)
- just _after_ the interpretation of this command.
-
- - A label is just an integer, starting from Lib.first_command_label
- initially, and incremented at each new successful command.
- - If some proofs are opened, we have their number in [nproofs],
- the name of the current proof in [prfname], the current depth in
- [prfdepth].
- - Otherwise, [nproofs = 0], [prfname = None], [prfdepth = 0]
- - The text of the command is stored (for Show Script currently).
- - A command can be tagged later as non-"reachable" when the current proof
- at the time of this command has been ended by Qed/Abort/Restart,
- meaning we can't backtrack there.
-*)
-
-type info = {
- label : int;
- nproofs : int;
- prfname : identifier option;
- prfdepth : int;
- ngoals : int;
- cmd : vernac_expr;
- mutable reachable : bool;
-}
-
-let history : info Stack.t = Stack.create ()
-
-(** Is this stack active (i.e. nonempty) ?
- The stack is currently inactive when compiling files (coqc). *)
-
-let is_active () = not (Stack.is_empty history)
-
-(** For debug purpose, a dump of the history *)
-
-let dump_history () =
- let l = ref [] in
- Stack.iter (fun i -> l:=i::!l) history;
- !l
-
-(** Basic manipulation of the command history stack *)
-
-exception Invalid
-
-let pop () = ignore (Stack.pop history)
-
-let npop n =
- (* Since our history stack always contains an initial entry,
- it's invalid to try to completely empty it *)
- if n < 0 || n >= Stack.length history then raise Invalid
- else for i = 1 to n do pop () done
-
-let top () =
- try Stack.top history with Stack.Empty -> raise Invalid
-
-(** Search the history stack for a suitable location. We perform first
- a non-destructive search: in case of search failure, the stack is
- unchanged. *)
-
-exception Found of info
-
-let search test =
- try
- Stack.iter (fun i -> if test i then raise (Found i)) history;
- raise Invalid
- with Found i ->
- while i != Stack.top history do pop () done
-
-(** An auxiliary function to retrieve the number of remaining subgoals *)
-
-let get_ngoals () =
- try
- let prf = Proof_global.give_me_the_proof () in
- List.length (Evd.sig_it (Proof.V82.background_subgoals prf))
- with Proof_global.NoCurrentProof -> 0
-
-(** Register the end of a command and store the current state *)
-
-let mark_command ast =
- Lib.add_frozen_state();
- Lib.mark_end_of_command();
- Stack.push
- { label = Lib.current_command_label ();
- nproofs = List.length (Pfedit.get_all_proof_names ());
- prfname =
- (try Some (Pfedit.get_current_proof_name ())
- with Proof_global.NoCurrentProof -> None);
- prfdepth = max 0 (Pfedit.current_proof_depth ());
- reachable = true;
- ngoals = get_ngoals ();
- cmd = ast }
- history
-
-(** Backtrack by aborting [naborts] proofs, then setting proof-depth back to
- [pnum] and finally going to state number [snum]. *)
-
-let raw_backtrack snum pnum naborts =
- for i = 1 to naborts do Pfedit.delete_current_proof () done;
- Pfedit.undo_todepth pnum;
- Lib.reset_label snum
-
-(** Re-sync the state of the system (label, proofs) with the top
- of the history stack. We may end on some earlier state to avoid
- re-opening proofs. This function will return the final label
- and the number of extra backtracking steps performed. *)
-
-let sync nb_opened_proofs =
- (* Backtrack by enough additional steps to avoid re-opening proofs.
- Typically, when a Qed has been crossed, we backtrack to the proof start.
- NB: We cannot reach the empty stack, since the first entry in the
- stack has no opened proofs and is tagged as reachable.
- *)
- let extra = ref 0 in
- while not (top()).reachable do incr extra; pop () done;
- let target = top ()
- in
- (* Now the opened proofs at target is a subset of the opened proofs before
- the backtrack, we simply abort the extra proofs (if any).
- NB: It is critical here that proofs are nested in a regular way
- (i.e. no more Resume or Suspend commands as earlier). This way, we can
- simply count the extra proofs to abort instead of taking care of their
- names.
- *)
- let naborts = nb_opened_proofs - target.nproofs in
- (* We are now ready to do a low-level backtrack *)
- raw_backtrack target.label target.prfdepth naborts;
- (target.label, !extra)
-
-(** Backtracking by a certain number of (non-state-preserving) commands.
- This is used by Coqide.
- It may actually undo more commands than asked : for instance instead
- of jumping back in the middle of a finished proof, we jump back before
- this proof. The number of extra backtracked command is returned at
- the end. *)
-
-let back count =
- if count = 0 then 0
- else
- let nb_opened_proofs = List.length (Pfedit.get_all_proof_names ()) in
- npop count;
- snd (sync nb_opened_proofs)
-
-(** Backtracking to a certain state number, and reset proofs accordingly.
- We may end on some earlier state if needed to avoid re-opening proofs.
- Return the final state number. *)
-
-let backto snum =
- if snum = Lib.current_command_label () then snum
- else
- let nb_opened_proofs = List.length (Pfedit.get_all_proof_names ()) in
- search (fun i -> i.label = snum);
- fst (sync nb_opened_proofs)
-
-(** Old [Backtrack] code with corresponding update of the history stack.
- [Backtrack] is now deprecated (in favor of [BackTo]) but is kept for
- compatibility with ProofGeneral. It's completely up to ProofGeneral
- to decide where to go and how to adapt proofs. Note that the choices
- of ProofGeneral are currently not always perfect (for instance when
- backtracking an Undo). *)
-
-let backtrack snum pnum naborts =
- raw_backtrack snum pnum naborts;
- search (fun i -> i.label = snum)
-
-(** [reset_initial] resets the system and clears the command history
- stack, only pushing back the initial entry. It should be equivalent
- to [backto Lib.first_command_label], but sligthly more efficient. *)
-
-let reset_initial () =
- let init_label = Lib.first_command_label in
- if Lib.current_command_label () = init_label then ()
- else begin
- Pfedit.delete_all_proofs ();
- Lib.reset_label init_label;
- Stack.clear history;
- Stack.push
- { label = init_label;
- nproofs = 0;
- prfname = None;
- prfdepth = 0;
- ngoals = 0;
- reachable = true;
- cmd = VernacNop }
- history
- end
-
-(** Reset to the last known state just before defining [id] *)
-
-let reset_name id =
- let lbl =
- try Lib.label_before_name id with Not_found -> raise Invalid
- in
- ignore (backto lbl)
-
-(** When a proof is ended (via either Qed/Admitted/Restart/Abort),
- old proof steps should be marked differently to avoid jumping back
- to them:
- - either this proof isn't there anymore in the proof engine
- - either it's there but it's a more recent attempt after a Restart,
- so we shouldn't mix the two.
- We also mark as unreachable the proof steps cancelled via a Undo. *)
-
-let mark_unreachable ?(after=0) prf_lst =
- let fix i = match i.prfname with
- | None -> raise Not_found (* stop hacking the history outside of proofs *)
- | Some p ->
- if List.mem p prf_lst && i.prfdepth > after
- then i.reachable <- false
- in
- try Stack.iter fix history with Not_found -> ()
-
-(** Parse the history stack for printing the script of a proof *)
-
-let get_script prf =
- let script = ref [] in
- let select i = match i.prfname with
- | None -> raise Not_found
- | Some p when p=prf && i.reachable -> script := i :: !script
- | _ -> ()
- in
- (try Stack.iter select history with Not_found -> ());
- (* Get rid of intermediate commands which don't grow the proof depth *)
- let rec filter n = function
- | [] -> []
- | {prfdepth=d; cmd=c; ngoals=ng}::l when n < d -> (c,ng) :: filter d l
- | {prfdepth=d}::l -> filter d l
- in
- (* initial proof depth (after entering the lemma statement) is 1 *)
- filter 1 !script
diff --git a/toplevel/backtrack.mli b/toplevel/backtrack.mli
deleted file mode 100644
index 315575dc..00000000
--- a/toplevel/backtrack.mli
+++ /dev/null
@@ -1,99 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(** Command history stack
-
- We maintain a stack of the past states of the system after each
- (non-state-preserving) interpreted commands
-*)
-
-(** [mark_command ast] marks the end of a command:
- - it stores a frozen state and a end of command in the Lib stack,
- - it stores the current state information in the command history
- stack *)
-
-val mark_command : Vernacexpr.vernac_expr -> unit
-
-(** Is this history stack active (i.e. nonempty) ?
- The stack is currently inactive when compiling files (coqc). *)
-
-val is_active : unit -> bool
-
-(** The [Invalid] exception is raised when one of the following function
- tries to empty the history stack, or reach an unknown states, etc.
- The stack is preserved in these cases. *)
-
-exception Invalid
-
-(** Nota Bene: it is critical for the following functions that proofs
- are nested in a regular way (i.e. no more Resume or Suspend commands
- as earlier). *)
-
-(** Backtracking by a certain number of (non-state-preserving) commands.
- This is used by Coqide.
- It may actually undo more commands than asked : for instance instead
- of jumping back in the middle of a finished proof, we jump back before
- this proof. The number of extra backtracked command is returned at
- the end. *)
-
-val back : int -> int
-
-(** Backtracking to a certain state number, and reset proofs accordingly.
- We may end on some earlier state if needed to avoid re-opening proofs.
- Return the state number on which we finally end. *)
-
-val backto : int -> int
-
-(** Old [Backtrack] code with corresponding update of the history stack.
- [Backtrack] is now deprecated (in favor of [BackTo]) but is kept for
- compatibility with ProofGeneral. It's completely up to ProofGeneral
- to decide where to go and how to adapt proofs. Note that the choices
- of ProofGeneral are currently not always perfect (for instance when
- backtracking an Undo). *)
-
-val backtrack : int -> int -> int -> unit
-
-(** [reset_initial] resets the system and clears the command history
- stack, only pushing back the initial entry. It should be equivalent
- to [backto Lib.first_command_label], but sligthly more efficient. *)
-
-val reset_initial : unit -> unit
-
-(** Reset to the last known state just before defining [id] *)
-
-val reset_name : Names.identifier Util.located -> unit
-
-(** When a proof is ended (via either Qed/Admitted/Restart/Abort),
- old proof steps should be marked differently to avoid jumping back
- to them:
- - either this proof isn't there anymore in the proof engine
- - either a proof with the same name is there, but it's a more recent
- attempt after a Restart/Abort, we shouldn't mix the two.
- We also mark as unreachable the proof steps cancelled via a Undo.
-*)
-
-val mark_unreachable : ?after:int -> Names.identifier list -> unit
-
-(** Parse the history stack for printing the script of a proof *)
-
-val get_script : Names.identifier -> (Vernacexpr.vernac_expr * int) list
-
-
-(** For debug purpose, a dump of the history *)
-
-type info = {
- label : int;
- nproofs : int;
- prfname : Names.identifier option;
- prfdepth : int;
- ngoals : int;
- cmd : Vernacexpr.vernac_expr;
- mutable reachable : bool;
-}
-
-val dump_history : unit -> info list
diff --git a/toplevel/cerrors.ml b/toplevel/cerrors.ml
index 76f020dd..22ea09c5 100644
--- a/toplevel/cerrors.ml
+++ b/toplevel/cerrors.ml
@@ -1,28 +1,30 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Compat
-open Pp
open Util
+open Pp
+open Errors
open Indtypes
open Type_errors
open Pretype_errors
open Indrec
let print_loc loc =
- if loc = dummy_loc then
+ if Loc.is_ghost loc then
(str"<unknown>")
else
- let loc = unloc loc in
+ let loc = Loc.unloc loc in
(int (fst loc) ++ str"-" ++ int (snd loc))
let guill s = "\""^s^"\""
+(** Invariant : exceptions embedded in EvaluatedError satisfy
+ Errors.noncritical *)
exception EvaluatedError of std_ppcmds * exn option
@@ -33,20 +35,16 @@ exception EvaluatedError of std_ppcmds * exn option
let explain_exn_default = function
(* Basic interaction exceptions *)
| Stream.Error txt -> hov 0 (str ("Syntax error: " ^ txt ^ "."))
- | Token.Error txt -> hov 0 (str ("Syntax error: " ^ txt ^ "."))
+ | Compat.Token.Error txt -> hov 0 (str ("Syntax error: " ^ txt ^ "."))
| Lexer.Error.E err -> hov 0 (str (Lexer.Error.to_string err))
| Sys_error msg -> hov 0 (str ("System error: " ^ guill msg))
| Out_of_memory -> hov 0 (str "Out of memory.")
| Stack_overflow -> hov 0 (str "Stack overflow.")
| Timeout -> hov 0 (str "Timeout!")
| Sys.Break -> hov 0 (fnl () ++ str "User interrupt.")
- (* Meta-exceptions *)
- | Loc.Exc_located (loc,exc) ->
- hov 0 ((if loc = dummy_loc then (mt ())
- else (str"At location " ++ print_loc loc ++ str":" ++ fnl ()))
- ++ Errors.print_no_anomaly exc)
+ (* Exceptions with pre-evaluated error messages *)
| EvaluatedError (msg,None) -> msg
- | EvaluatedError (msg,Some reraise) -> msg ++ Errors.print_no_anomaly reraise
+ | EvaluatedError (msg,Some reraise) -> msg ++ Errors.print reraise
(* Otherwise, not handled here *)
| _ -> raise Errors.Unhandled
@@ -55,67 +53,77 @@ let _ = Errors.register_handler explain_exn_default
(** Pre-explain a vernac interpretation error *)
-let wrap_vernac_error strm =
- EvaluatedError (hov 0 (str "Error:" ++ spc () ++ strm), None)
-
-let rec process_vernac_interp_error = function
- | Univ.UniverseInconsistency (o,u,v) ->
- let msg =
- if !Constrextern.print_universes then
- spc() ++ str "(cannot enforce" ++ spc() ++ Univ.pr_uni u ++ spc() ++
- str (match o with Univ.Lt -> "<" | Univ.Le -> "<=" | Univ.Eq -> "=")
- ++ spc() ++ Univ.pr_uni v ++ str")"
- else
- mt() in
- wrap_vernac_error (str "Universe inconsistency" ++ msg ++ str ".")
+let wrap_vernac_error (exn, info) strm =
+ let header = Pp.tag (Pp.Tag.inj Ppstyle.error_tag Ppstyle.tag) (str "Error:") in
+ let e = EvaluatedError (hov 0 (header ++ spc () ++ strm), None) in
+ (e, info)
+
+let process_vernac_interp_error exn = match fst exn with
+ | Univ.UniverseInconsistency i ->
+ let msg =
+ if !Constrextern.print_universes then
+ str "." ++ spc() ++
+ Univ.explain_universe_inconsistency Universes.pr_with_global_universes i
+ else
+ mt() in
+ wrap_vernac_error exn (str "Universe inconsistency" ++ msg ++ str ".")
| TypeError(ctx,te) ->
- wrap_vernac_error (Himsg.explain_type_error ctx Evd.empty te)
+ wrap_vernac_error exn (Himsg.explain_type_error ctx Evd.empty te)
| PretypeError(ctx,sigma,te) ->
- wrap_vernac_error (Himsg.explain_pretype_error ctx sigma te)
+ wrap_vernac_error exn (Himsg.explain_pretype_error ctx sigma te)
| Typeclasses_errors.TypeClassError(env, te) ->
- wrap_vernac_error (Himsg.explain_typeclass_error env te)
+ wrap_vernac_error exn (Himsg.explain_typeclass_error env te)
| InductiveError e ->
- wrap_vernac_error (Himsg.explain_inductive_error e)
+ wrap_vernac_error exn (Himsg.explain_inductive_error e)
| Modops.ModuleTypingError e ->
- wrap_vernac_error (Himsg.explain_module_error e)
+ wrap_vernac_error exn (Himsg.explain_module_error e)
| Modintern.ModuleInternalizationError e ->
- wrap_vernac_error (Himsg.explain_module_internalization_error e)
+ wrap_vernac_error exn (Himsg.explain_module_internalization_error e)
| RecursionSchemeError e ->
- wrap_vernac_error (Himsg.explain_recursion_scheme_error e)
- | Cases.PatternMatchingError (env,e) ->
- wrap_vernac_error (Himsg.explain_pattern_matching_error env e)
+ wrap_vernac_error exn (Himsg.explain_recursion_scheme_error e)
+ | Cases.PatternMatchingError (env,sigma,e) ->
+ wrap_vernac_error exn (Himsg.explain_pattern_matching_error env sigma e)
| Tacred.ReductionTacticError e ->
- wrap_vernac_error (Himsg.explain_reduction_tactic_error e)
+ wrap_vernac_error exn (Himsg.explain_reduction_tactic_error e)
| Logic.RefinerError e ->
- wrap_vernac_error (Himsg.explain_refiner_error e)
+ wrap_vernac_error exn (Himsg.explain_refiner_error e)
| Nametab.GlobalizationError q ->
- wrap_vernac_error
+ wrap_vernac_error exn
(str "The reference" ++ spc () ++ Libnames.pr_qualid q ++
spc () ++ str "was not found" ++
spc () ++ str "in the current" ++ spc () ++ str "environment.")
- | Nametab.GlobalizationConstantError q ->
- wrap_vernac_error
- (str "No constant of this name:" ++ spc () ++
- Libnames.pr_qualid q ++ str ".")
| Refiner.FailError (i,s) ->
- wrap_vernac_error
+ let s = Lazy.force s in
+ wrap_vernac_error exn
(str "Tactic failure" ++
- (if Lazy.force s <> mt() then str ":" ++ Lazy.force s else mt ()) ++
- if i=0 then str "." else str " (level " ++ int i ++ str").")
+ (if Pp.is_empty s then s else str ": " ++ s) ++
+ if Int.equal i 0 then str "." else str " (level " ++ int i ++ str").")
| AlreadyDeclared msg ->
- wrap_vernac_error (msg ++ str ".")
- | Proof_type.LtacLocated (_,(Refiner.FailError (i,s) as exc)) when Lazy.force s <> mt () ->
- process_vernac_interp_error exc
- | Proof_type.LtacLocated (s,exc) ->
- EvaluatedError (hov 0 (Himsg.explain_ltac_call_trace s ++ fnl()),
- Some (process_vernac_interp_error exc))
- | Loc.Exc_located (loc,exc) ->
- Loc.Exc_located (loc,process_vernac_interp_error exc)
- | exc ->
- exc
+ wrap_vernac_error exn (msg ++ str ".")
+ | _ ->
+ exn
+
+let rec strip_wrapping_exceptions = function
+ | Logic_monad.TacticFailure e ->
+ strip_wrapping_exceptions e
+ | exc -> exc
+
+let process_vernac_interp_error (exc, info) =
+ let exc = strip_wrapping_exceptions exc in
+ let e = process_vernac_interp_error (exc, 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
+ | 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 (process_vernac_interp_error e))
+ (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 (process_vernac_interp_error e))
+ (fun e -> Errors.print_no_report (fst (process_vernac_interp_error (e, Exninfo.null))))
diff --git a/toplevel/cerrors.mli b/toplevel/cerrors.mli
index c9cada8c..1768af11 100644
--- a/toplevel/cerrors.mli
+++ b/toplevel/cerrors.mli
@@ -1,24 +1,21 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Pp
-open Util
-
(** Error report. *)
-val print_loc : loc -> std_ppcmds
+val print_loc : Loc.t -> Pp.std_ppcmds
(** Pre-explain a vernac interpretation error *)
-val process_vernac_interp_error : exn -> exn
+val process_vernac_interp_error : Util.iexn -> Util.iexn
(** General explain function. Should not be used directly now,
see instead function [Errors.print] and variants *)
-val explain_exn_default : exn -> std_ppcmds
+val explain_exn_default : exn -> Pp.std_ppcmds
diff --git a/toplevel/class.ml b/toplevel/class.ml
index 214fbf5b..6a485d52 100644
--- a/toplevel/class.ml
+++ b/toplevel/class.ml
@@ -1,31 +1,29 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Errors
open Util
open Pp
open Names
-open Nameops
open Term
+open Vars
open Termops
-open Inductive
-open Declarations
open Entries
open Environ
-open Inductive
-open Lib
open Classops
open Declare
-open Libnames
+open Globnames
open Nametab
open Decl_kinds
-open Safe_typing
-let strength_min l = if List.mem Local l then Local else Global
+let strength_min l = if List.mem `LOCAL l then `LOCAL else `GLOBAL
+
+let loc_of_bool b = if b then `LOCAL else `GLOBAL
(* Errors *)
@@ -38,7 +36,6 @@ type coercion_error_kind =
| NoTarget
| WrongTarget of cl_typ * cl_typ
| NotAClass of global_reference
- | NotEnoughClassArgs of cl_typ
exception CoercionError of coercion_error_kind
@@ -65,18 +62,17 @@ let explain_coercion_error g = function
| NotAClass ref ->
(str "Type of " ++ Printer.pr_global ref ++
str " does not end with a sort")
- | NotEnoughClassArgs cl ->
- (str"Wrong number of parameters for " ++ pr_class cl)
(* Verifications pour l'ajout d'une classe *)
let check_reference_arity ref =
- if not (Reductionops.is_arity (Global.env()) Evd.empty (Global.type_of_global ref)) then
+ if not (Reductionops.is_arity (Global.env()) Evd.empty (Global.type_of_global_unsafe ref)) then
raise (CoercionError (NotAClass ref))
let check_arity = function
| CL_FUN | CL_SORT -> ()
| CL_CONST cst -> check_reference_arity (ConstRef cst)
+ | CL_PROJ cst -> check_reference_arity (ConstRef cst)
| CL_SECVAR id -> check_reference_arity (VarRef id)
| CL_IND kn -> check_reference_arity (IndRef kn)
@@ -84,7 +80,7 @@ let check_arity = function
(* check that the computed target is the provided one *)
let check_target clt = function
- | Some cl when cl <> clt -> raise (CoercionError (WrongTarget(clt,cl)))
+ | Some cl when not (cl_typ_eq cl clt) -> raise (CoercionError (WrongTarget(clt,cl)))
| _ -> ()
(* condition d'heritage uniforme *)
@@ -94,13 +90,15 @@ let uniform_cond nargs lt =
| (0,[]) -> true
| (n,t::l) ->
let t = strip_outer_cast t in
- isRel t && destRel t = n && aux ((n-1),l)
+ isRel t && Int.equal (destRel t) n && aux ((n-1),l)
| _ -> false
in
aux (nargs,lt)
let class_of_global = function
- | ConstRef sp -> CL_CONST sp
+ | ConstRef sp ->
+ if Environ.is_projection sp (Global.env ())
+ then CL_PROJ sp else CL_CONST sp
| IndRef sp -> CL_IND sp
| VarRef id -> CL_SECVAR id
| ConstructRef _ as c ->
@@ -123,19 +121,19 @@ l'indice de la classe source dans la liste lp
let get_source lp source =
match source with
| None ->
- let (cl1,lv1) =
+ let (cl1,u1,lv1) =
match lp with
| [] -> raise Not_found
| t1::_ -> find_class_type Evd.empty t1
in
- (cl1,lv1,1)
+ (cl1,u1,lv1,1)
| Some cl ->
let rec aux = function
| [] -> raise Not_found
| t1::lt ->
try
- let cl1,lv1 = find_class_type Evd.empty t1 in
- if cl = cl1 then cl1,lv1,(List.length lt+1)
+ let cl1,u1,lv1 = find_class_type Evd.empty t1 in
+ if cl_typ_eq cl cl1 then cl1,u1,lv1,(List.length lt+1)
else raise Not_found
with Not_found -> aux lt
in aux (List.rev lp)
@@ -144,7 +142,11 @@ let get_target t ind =
if (ind > 1) then
CL_FUN
else
- fst (find_class_type Evd.empty t)
+ match pi1 (find_class_type Evd.empty t) with
+ | CL_CONST p when Environ.is_projection p (Global.env ()) ->
+ CL_PROJ p
+ | x -> x
+
let prods_of t =
let rec aux acc d = match kind_of_term d with
@@ -155,9 +157,13 @@ let prods_of t =
aux [] t
let strength_of_cl = function
- | CL_CONST kn -> Global
- | CL_SECVAR id -> Local
- | _ -> Global
+ | CL_CONST kn -> `GLOBAL
+ | CL_SECVAR id -> `LOCAL
+ | _ -> `GLOBAL
+
+let strength_of_global = function
+ | VarRef _ -> `LOCAL
+ | _ -> `GLOBAL
let get_strength stre ref cls clt =
let stres = strength_of_cl cls in
@@ -168,28 +174,28 @@ let get_strength stre ref cls clt =
let ident_key_of_class = function
| CL_FUN -> "Funclass"
| CL_SORT -> "Sortclass"
- | CL_CONST sp -> string_of_label (con_label sp)
- | CL_IND (sp,_) -> string_of_label (mind_label sp)
- | CL_SECVAR id -> string_of_id id
+ | CL_CONST sp | CL_PROJ sp -> Label.to_string (con_label sp)
+ | CL_IND (sp,_) -> Label.to_string (mind_label sp)
+ | CL_SECVAR id -> Id.to_string id
-(* coercion identité *)
+(* Identity coercion *)
let error_not_transparent source =
errorlabstrm "build_id_coercion"
(pr_class source ++ str " must be a transparent constant.")
-let build_id_coercion idf_opt source =
+let build_id_coercion idf_opt source poly =
let env = Global.env () in
- let vs = match source with
- | CL_CONST sp -> mkConst sp
+ let vs, ctx = match source with
+ | CL_CONST sp -> Universes.fresh_global_instance env (ConstRef sp)
| _ -> error_not_transparent source in
- let c = match constant_opt_value env (destConst vs) with
+ let c = match constant_opt_value_in env (destConst vs) with
| Some c -> c
| None -> error_not_transparent source in
let lams,t = decompose_lam_assum c in
let val_f =
it_mkLambda_or_LetIn
- (mkLambda (Name (id_of_string "x"),
+ (mkLambda (Name Namegen.default_dependent_ident,
applistc vs (extended_rel_list 0 lams),
mkRel 1))
lams
@@ -212,17 +218,17 @@ let build_id_coercion idf_opt source =
match idf_opt with
| Some idf -> idf
| None ->
- let cl,_ = find_class_type Evd.empty t in
- id_of_string ("Id_"^(ident_key_of_class source)^"_"^
+ let cl,u,_ = find_class_type Evd.empty t in
+ Id.of_string ("Id_"^(ident_key_of_class source)^"_"^
(ident_key_of_class cl))
in
let constr_entry = (* Cast is necessary to express [val_f] is identity *)
DefinitionEntry
- { const_entry_body = mkCast (val_f, DEFAULTcast, typ_f);
- const_entry_secctx = None;
- const_entry_type = Some typ_f;
- const_entry_opaque = false } in
- let kn = declare_constant idf (constr_entry,IsDefinition IdentityCoercion) in
+ (definition_entry ~types:typ_f ~poly ~univs:(Univ.ContextSet.to_context ctx)
+ ~inline:true (mkCast (val_f, DEFAULTcast, typ_f)))
+ in
+ let decl = (constr_entry, IsDefinition IdentityCoercion) in
+ let kn = declare_constant idf decl in
ConstRef kn
let check_source = function
@@ -240,14 +246,14 @@ booleen "coercion identite'?"
lorque source est None alors target est None aussi.
*)
-let add_new_coercion_core coef stre source target isid =
+let add_new_coercion_core coef stre poly source target isid =
check_source source;
- let t = Global.type_of_global coef in
+ let t = Global.type_of_global_unsafe coef in
if coercion_exists coef then raise (CoercionError AlreadyExists);
let tg,lp = prods_of t in
let llp = List.length lp in
- if llp = 0 then raise (CoercionError NotAFunction);
- let (cls,lvs,ind) =
+ if Int.equal llp 0 then raise (CoercionError NotAFunction);
+ let (cls,us,lvs,ind) =
try
get_source lp source
with Not_found ->
@@ -255,7 +261,7 @@ let add_new_coercion_core coef stre source target isid =
in
check_source (Some cls);
if not (uniform_cond (llp-ind) lvs) then
- msg_warn (Pp.string_of_ppcmds (explain_coercion_error coef NotUniform));
+ msg_warning (explain_coercion_error coef NotUniform);
let clt =
try
get_target tg ind
@@ -265,38 +271,55 @@ let add_new_coercion_core coef stre source target isid =
check_target clt target;
check_arity cls;
check_arity clt;
- let stre' = get_strength stre coef cls clt in
- declare_coercion coef stre' ~isid ~src:cls ~target:clt ~params:(List.length lvs)
+ let local = match get_strength stre coef cls clt with
+ | `LOCAL -> true
+ | `GLOBAL -> false
+ in
+ declare_coercion coef ~local ~isid ~src:cls ~target:clt ~params:(List.length lvs)
+
-let try_add_new_coercion_core ref b c d e =
- try add_new_coercion_core ref b c d e
+let try_add_new_coercion_core ref ~local c d e f =
+ try add_new_coercion_core ref (loc_of_bool local) c d e f
with CoercionError e ->
errorlabstrm "try_add_new_coercion_core"
(explain_coercion_error ref e ++ str ".")
-let try_add_new_coercion ref stre =
- try_add_new_coercion_core ref stre None None false
+let try_add_new_coercion ref ~local poly =
+ try_add_new_coercion_core ref ~local poly None None false
-let try_add_new_coercion_subclass cl stre =
- let coe_ref = build_id_coercion None cl in
- try_add_new_coercion_core coe_ref stre (Some cl) None true
+let try_add_new_coercion_subclass cl ~local poly =
+ let coe_ref = build_id_coercion None cl poly in
+ try_add_new_coercion_core coe_ref ~local poly (Some cl) None true
-let try_add_new_coercion_with_target ref stre ~source ~target =
- try_add_new_coercion_core ref stre (Some source) (Some target) false
+let try_add_new_coercion_with_target ref ~local poly ~source ~target =
+ try_add_new_coercion_core ref ~local poly (Some source) (Some target) false
-let try_add_new_identity_coercion id stre ~source ~target =
- let ref = build_id_coercion (Some id) source in
- try_add_new_coercion_core ref stre (Some source) (Some target) true
+let try_add_new_identity_coercion id ~local poly ~source ~target =
+ let ref = build_id_coercion (Some id) source poly in
+ try_add_new_coercion_core ref ~local poly (Some source) (Some target) true
-let try_add_new_coercion_with_source ref stre ~source =
- try_add_new_coercion_core ref stre (Some source) None false
+let try_add_new_coercion_with_source ref ~local poly ~source =
+ try_add_new_coercion_core ref ~local poly (Some source) None false
-let add_coercion_hook stre ref =
- try_add_new_coercion ref stre;
- Flags.if_verbose message
- (string_of_qualid (shortest_qualid_of_global Idset.empty ref)
- ^ " is now a coercion")
+let add_coercion_hook poly local ref =
+ let stre = match local with
+ | Local -> true
+ | Global -> false
+ | Discharge -> assert false
+ in
+ let () = try_add_new_coercion ref stre poly in
+ let msg = pr_global_env Id.Set.empty ref ++ str " is now a coercion" in
+ Flags.if_verbose msg_info msg
+
+let add_coercion_hook poly = Lemmas.mk_hook (add_coercion_hook poly)
-let add_subclass_hook stre ref =
+let add_subclass_hook poly local ref =
+ let stre = match local with
+ | Local -> true
+ | Global -> false
+ | Discharge -> assert false
+ in
let cl = class_of_global ref in
- try_add_new_coercion_subclass cl stre
+ try_add_new_coercion_subclass cl stre poly
+
+let add_subclass_hook poly = Lemmas.mk_hook (add_subclass_hook poly)
diff --git a/toplevel/class.mli b/toplevel/class.mli
index d2e12de6..bd6c7a6d 100644
--- a/toplevel/class.mli
+++ b/toplevel/class.mli
@@ -1,49 +1,48 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Names
-open Term
open Classops
-open Declare
-open Libnames
-open Decl_kinds
-open Nametab
+open Globnames
(** Classes and coercions. *)
(** [try_add_new_coercion_with_target ref s src tg] declares [ref] as a coercion
from [src] to [tg] *)
-val try_add_new_coercion_with_target : global_reference -> locality ->
+val try_add_new_coercion_with_target : global_reference -> local:bool ->
+ Decl_kinds.polymorphic ->
source:cl_typ -> target:cl_typ -> unit
(** [try_add_new_coercion ref s] declares [ref], assumed to be of type
[(x1:T1)...(xn:Tn)src->tg], as a coercion from [src] to [tg] *)
-val try_add_new_coercion : global_reference -> locality -> unit
+val try_add_new_coercion : global_reference -> local:bool ->
+ Decl_kinds.polymorphic -> unit
(** [try_add_new_coercion_subclass cst s] expects that [cst] denotes a
transparent constant which unfolds to some class [tg]; it declares
an identity coercion from [cst] to [tg], named something like
["Id_cst_tg"] *)
-val try_add_new_coercion_subclass : cl_typ -> locality -> unit
+val try_add_new_coercion_subclass : cl_typ -> local:bool ->
+ Decl_kinds.polymorphic -> unit
(** [try_add_new_coercion_with_source ref s src] declares [ref] as a coercion
from [src] to [tg] where the target is inferred from the type of [ref] *)
-val try_add_new_coercion_with_source : global_reference -> locality ->
- source:cl_typ -> unit
+val try_add_new_coercion_with_source : global_reference -> local:bool ->
+ Decl_kinds.polymorphic -> source:cl_typ -> unit
(** [try_add_new_identity_coercion id s src tg] enriches the
environment with a new definition of name [id] declared as an
identity coercion from [src] to [tg] *)
-val try_add_new_identity_coercion : identifier -> locality ->
- source:cl_typ -> target:cl_typ -> unit
+val try_add_new_identity_coercion : Id.t -> local:bool ->
+ Decl_kinds.polymorphic -> source:cl_typ -> target:cl_typ -> unit
-val add_coercion_hook : Tacexpr.declaration_hook
+val add_coercion_hook : Decl_kinds.polymorphic -> unit Lemmas.declaration_hook
-val add_subclass_hook : Tacexpr.declaration_hook
+val add_subclass_hook : Decl_kinds.polymorphic -> unit Lemmas.declaration_hook
val class_of_global : global_reference -> cl_typ
diff --git a/toplevel/classes.ml b/toplevel/classes.ml
index 3843ea83..f44ac367 100644
--- a/toplevel/classes.ml
+++ b/toplevel/classes.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,80 +8,77 @@
(*i*)
open Names
-open Decl_kinds
open Term
-open Sign
-open Entries
-open Evd
+open Vars
open Environ
open Nametab
-open Mod_subst
+open Errors
open Util
open Typeclasses_errors
open Typeclasses
open Libnames
+open Globnames
open Constrintern
-open Glob_term
-open Topconstr
+open Constrexpr
(*i*)
open Decl_kinds
open Entries
+let refine_instance = ref true
+
+let _ = Goptions.declare_bool_option {
+ Goptions.optsync = true;
+ Goptions.optdepr = false;
+ Goptions.optname = "definition of instances by refining";
+ Goptions.optkey = ["Refine";"Instance";"Mode"];
+ Goptions.optread = (fun () -> !refine_instance);
+ Goptions.optwrite = (fun b -> refine_instance := b)
+}
+
let typeclasses_db = "typeclass_instances"
let set_typeclass_transparency c local b =
- Auto.add_hints local [typeclasses_db]
- (Auto.HintsTransparencyEntry ([c], b))
+ Hints.add_hints local [typeclasses_db]
+ (Hints.HintsTransparencyEntry ([c], b))
let _ =
- Typeclasses.register_add_instance_hint
- (fun inst local pri ->
- let path =
- try Auto.PathHints [global_of_constr inst]
- with e when Errors.noncritical e -> Auto.PathAny
+ Hook.set Typeclasses.add_instance_hint_hook
+ (fun inst path local pri poly ->
+ let inst' = match inst with IsConstr c -> Hints.IsConstr (c, Univ.ContextSet.empty)
+ | IsGlobal gr -> Hints.IsGlobRef gr
in
Flags.silently (fun () ->
- Auto.add_hints local [typeclasses_db]
- (Auto.HintsResolveEntry
- [pri, false, path, inst])) ());
- Typeclasses.register_set_typeclass_transparency set_typeclass_transparency;
- Typeclasses.register_classes_transparent_state
- (fun () -> Auto.Hint_db.transparent_state (Auto.searchtable_map typeclasses_db))
-
-let declare_class g =
- match global g with
- | ConstRef x -> Typeclasses.add_constant_class x
- | IndRef x -> Typeclasses.add_inductive_class x
- | _ -> user_err_loc (loc_of_reference g, "declare_class",
- Pp.str"Unsupported class type, only constants and inductives are allowed")
-
+ Hints.add_hints local [typeclasses_db]
+ (Hints.HintsResolveEntry
+ [pri, 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))
+
(** TODO: add subinstances *)
-let existing_instance glob g =
+let existing_instance glob g pri =
let c = global g in
- let instance = Typing.type_of (Global.env ()) Evd.empty (constr_of_global c) 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, _)) -> add_instance (new_instance tc None glob c)
+ | Some (_, ((tc,u), _)) -> add_instance (new_instance tc pri 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.")
let mismatched_params env n m = mismatched_ctx_inst env Parameters n m
let mismatched_props env n m = mismatched_ctx_inst env Properties n m
-type binder_list = (identifier located * bool * constr_expr) list
-
(* Declare everything in the parameters as implicit, and the class instance as well *)
-open Topconstr
-
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
let c', l =
match b with
- | None -> interp_casted_constr_evars evars env (List.hd l) t', List.tl l
+ | None -> interp_casted_constr_evars env evars (List.hd l) t', List.tl l
| Some b -> substl subst b, l
in
let d = na, Some c', t' in
@@ -89,11 +86,9 @@ let type_ctx_instance evars env ctx inst subst =
| [] -> subst
in aux (subst, []) inst (List.rev ctx)
-let refine_ref = ref (fun _ -> assert(false))
-
let id_of_class cl =
match cl.cl_impl with
- | ConstRef kn -> let _,_,l = repr_con kn in id_of_label l
+ | ConstRef kn -> let _,_,l = repr_con kn in Label.to_id l
| IndRef (kn,i) ->
let mip = (Environ.lookup_mind kn (Global.env ())).Declarations.mind_packets in
mip.(0).Declarations.mind_typename
@@ -101,55 +96,53 @@ let id_of_class cl =
open Pp
-let ($$) g f = fun x -> g (f x)
-
let instance_hook k pri global imps ?hook cst =
Impargs.maybe_declare_manual_implicits false cst ~enriching:false imps;
Typeclasses.declare_instance pri (not global) cst;
(match hook with Some h -> h cst | None -> ())
-let declare_instance_constant k pri global imps ?hook id term termtype =
- let cdecl =
- let kind = IsDefinition Instance in
- let entry =
- { const_entry_body = term;
- const_entry_secctx = None;
- const_entry_type = Some termtype;
- const_entry_opaque = false }
- in DefinitionEntry entry, kind
+let declare_instance_constant k pri global imps ?hook id poly uctx term termtype =
+ let kind = IsDefinition Instance in
+ let entry =
+ Declare.definition_entry ~types:termtype ~poly ~univs:uctx term
in
+ let cdecl = (DefinitionEntry entry, kind) in
let kn = Declare.declare_constant id cdecl in
Declare.definition_message id;
instance_hook k pri global imps ?hook (ConstRef kn);
id
-let new_instance ?(abstract=false) ?(global=false) ctx (instid, bk, cl) props
+let new_instance ?(abstract=false) ?(global=false) poly ctx (instid, bk, cl) props
?(generalize=true)
- ?(tac:Proof_type.tactic option) ?(hook:(global_reference -> unit) option) pri =
+ ?(tac:unit Proofview.tactic option) ?hook pri =
let env = Global.env() in
- let evars = ref Evd.empty in
+ let evars = ref (Evd.from_env env) in
let tclass, ids =
match bk with
| Implicit ->
- Implicit_quantifiers.implicit_application Idset.empty ~allow_partial:false
+ Implicit_quantifiers.implicit_application Id.Set.empty ~allow_partial:false
(fun avoid (clname, (id, _, t)) ->
match clname with
| Some (cl, b) ->
- let t = CHole (Util.dummy_loc, None) in
+ let t = CHole (Loc.ghost, None, Misctypes.IntroAnonymous, None) in
t, avoid
| None -> failwith ("new instance: under-applied typeclass"))
cl
- | Explicit -> cl, Idset.empty
+ | Explicit -> cl, Id.Set.empty
+ in
+ let tclass =
+ if generalize then CGeneralization (Loc.ghost, Implicit, Some AbsPi, tclass)
+ else tclass
in
- let tclass = if generalize then CGeneralization (dummy_loc, Implicit, Some AbsPi, tclass) else tclass in
- let k, cty, ctx', ctx, len, imps, subst =
- let impls, ((env', ctx), imps) = interp_context_evars evars env ctx in
- let c', imps' = interp_type_evars_impls ~impls ~evdref:evars ~fail_evar:false env' tclass in
+ let k, u, cty, ctx', ctx, len, imps, subst =
+ let impls, ((env', ctx), imps) = interp_context_evars env evars ctx in
+ let c', imps' = interp_type_evars_impls ~impls env' evars tclass in
let len = List.length ctx in
let imps = imps @ Impargs.lift_implicits len imps' in
let ctx', c = decompose_prod_assum c' in
let ctx'' = ctx' @ ctx in
- let cl, args = Typeclasses.dest_class_app (push_rel_context ctx'' env) c in
+ let k, args = Typeclasses.dest_class_app (push_rel_context ctx'' env) c in
+ let cl, u = Typeclasses.typeclass_univ_instance k in
let _, args =
List.fold_right (fun (na, b, t) (args, args') ->
match b with
@@ -157,7 +150,7 @@ let new_instance ?(abstract=false) ?(global=false) ctx (instid, bk, cl) props
| Some b -> (args, substl args' b :: args'))
(snd cl.cl_context) (args, [])
in
- cl, c', ctx', ctx, len, imps, args
+ cl, u, c', ctx', ctx, len, imps, args
in
let id =
match snd instid with
@@ -172,158 +165,217 @@ let new_instance ?(abstract=false) ?(global=false) ctx (instid, bk, cl) props
in
let env' = push_rel_context ctx env in
evars := Evarutil.nf_evar_map !evars;
- evars := resolve_typeclasses env !evars;
- let sigma = !evars in
- let subst = List.map (Evarutil.nf_evar sigma) subst in
+ evars := resolve_typeclasses ~filter:Typeclasses.no_goals ~fail:true env !evars;
+ let subst = List.map (Evarutil.nf_evar !evars) subst in
if abstract then
begin
- if not (Lib.is_modtype ()) then
- error "Declare Instance while not in Module Type.";
- let _, ty_constr = instance_constructor k (List.rev subst) in
+ let subst = List.fold_left2
+ (fun subst' s (_, b, _) -> if Option.is_empty b then s :: subst' else subst')
+ [] subst (snd k.cl_context)
+ in
+ let (_, ty_constr) = instance_constructor (k,u) subst in
+ let nf, subst = Evarutil.e_nf_evars_and_universes evars in
let termtype =
let t = it_mkProd_or_LetIn ty_constr (ctx' @ ctx) in
- Evarutil.nf_evar !evars t
+ nf t
in
Evarutil.check_evars env Evd.empty !evars termtype;
+ let ctx = Evd.universe_context !evars in
let cst = Declare.declare_constant ~internal:Declare.KernelSilent id
(Entries.ParameterEntry
- (None,termtype,None), Decl_kinds.IsAssumption Decl_kinds.Logical)
+ (None,poly,(termtype,ctx),None), Decl_kinds.IsAssumption Decl_kinds.Logical)
in instance_hook k None global imps ?hook (ConstRef cst); id
end
- else
- begin
- let props =
- match props with
- | Some (CRecord (loc, _, fs)) ->
- if List.length fs > List.length k.cl_props then
- mismatched_props env' (List.map snd fs) k.cl_props;
- Some (Inl fs)
- | Some t -> Some (Inr t)
- | None -> None
- in
- let subst =
- match props with
- | None -> if k.cl_props = [] then Some (Inl subst) else None
- | Some (Inr term) ->
- let c = interp_casted_constr_evars evars env' term cty in
- Some (Inr (c, subst))
- | Some (Inl props) ->
- let get_id =
- function
- | Ident id' -> id'
- | _ -> errorlabstrm "new_instance" (Pp.str "Only local structures are handled")
- in
- let props, rest =
- List.fold_left
- (fun (props, rest) (id,b,_) ->
- if b = None then
- try
- let (loc_mid, c) = List.find (fun (id', _) -> Name (snd (get_id id')) = id) rest in
- let rest' = List.filter (fun (id', _) -> Name (snd (get_id id')) <> id) rest in
- let (loc, mid) = get_id loc_mid in
- List.iter (fun (n, _, x) ->
- if n = Name mid then
- Option.iter (fun x -> Dumpglob.add_glob loc (ConstRef x)) x)
- k.cl_projs;
- c :: props, rest'
- with Not_found ->
- (CHole (Util.dummy_loc, None) :: props), rest
- else props, rest)
- ([], props) k.cl_props
- in
- if rest <> [] then
- unbound_method env' k.cl_impl (get_id (fst (List.hd rest)))
- else
- Some (Inl (type_ctx_instance evars (push_rel_context ctx' env') k.cl_props props subst))
- in
+ else (
+ let props =
+ match props with
+ | Some (true, CRecord (loc, _, fs)) ->
+ if List.length fs > List.length k.cl_props then
+ mismatched_props env' (List.map snd fs) k.cl_props;
+ Some (Inl fs)
+ | Some (_, t) -> Some (Inr t)
+ | None ->
+ if Flags.is_program_mode () then Some (Inl [])
+ else None
+ in
+ let subst =
+ match props with
+ | None -> if List.is_empty k.cl_props then Some (Inl subst) else None
+ | Some (Inr term) ->
+ let c = interp_casted_constr_evars env' evars term cty in
+ Some (Inr (c, subst))
+ | Some (Inl props) ->
+ let get_id =
+ function
+ | Ident id' -> id'
+ | _ -> errorlabstrm "new_instance" (Pp.str "Only local structures are handled")
+ in
+ let props, rest =
+ List.fold_left
+ (fun (props, rest) (id,b,_) ->
+ if Option.is_empty b then
+ try
+ let is_id (id', _) = match id, get_id id' with
+ | Name id, (_, id') -> Id.equal id id'
+ | Anonymous, _ -> false
+ in
+ let (loc_mid, c) =
+ List.find is_id rest
+ in
+ let rest' =
+ List.filter (fun v -> not (is_id v)) rest
+ in
+ let (loc, mid) = get_id loc_mid in
+ List.iter (fun (n, _, x) ->
+ if Name.equal n (Name mid) then
+ Option.iter (fun x -> Dumpglob.add_glob loc (ConstRef x)) x)
+ k.cl_projs;
+ c :: props, rest'
+ with Not_found ->
+ (CHole (Loc.ghost, Some Evar_kinds.GoalEvar, Misctypes.IntroAnonymous, None) :: props), rest
+ else props, rest)
+ ([], props) k.cl_props
+ in
+ match rest with
+ | (n, _) :: _ ->
+ unbound_method env' k.cl_impl (get_id n)
+ | _ ->
+ Some (Inl (type_ctx_instance evars (push_rel_context ctx' env')
+ k.cl_props props subst))
+ in
+ let term, termtype =
+ match subst with
+ | None -> let termtype = it_mkProd_or_LetIn cty ctx in
+ None, termtype
+ | Some (Inl subst) ->
+ let subst = List.fold_left2
+ (fun subst' s (_, b, _) -> if Option.is_empty b then s :: subst' else subst')
+ [] subst (k.cl_props @ snd k.cl_context)
+ in
+ let (app, ty_constr) = instance_constructor (k,u) subst in
+ let termtype = it_mkProd_or_LetIn ty_constr (ctx' @ ctx) in
+ let term = Termops.it_mkLambda_or_LetIn (Option.get app) (ctx' @ ctx) in
+ Some term, termtype
+ | Some (Inr (def, subst)) ->
+ let termtype = it_mkProd_or_LetIn cty ctx in
+ let term = Termops.it_mkLambda_or_LetIn def ctx in
+ Some term, termtype
+ in
+ let _ =
evars := Evarutil.nf_evar_map !evars;
- let term, termtype =
- match subst with
- | None -> let termtype = it_mkProd_or_LetIn cty ctx in
- None, termtype
- | Some (Inl subst) ->
- let subst = List.fold_left2
- (fun subst' s (_, b, _) -> if b = None then s :: subst' else subst')
- [] subst (k.cl_props @ snd k.cl_context)
+ evars := Typeclasses.resolve_typeclasses ~filter:Typeclasses.no_goals_or_obligations ~fail:true
+ env !evars;
+ (* Try resolving fields that are typeclasses automatically. *)
+ evars := Typeclasses.resolve_typeclasses ~filter:Typeclasses.all_evars ~fail:false
+ env !evars
+ in
+ let _ = evars := Evarutil.nf_evar_map_undefined !evars in
+ let evm, nf = Evarutil.nf_evar_map_universes !evars in
+ let termtype = nf termtype in
+ let _ = (* Check that the type is free of evars now. *)
+ Evarutil.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 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
+ 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)
in
- let app, ty_constr = instance_constructor k subst in
- let termtype = it_mkProd_or_LetIn ty_constr (ctx' @ ctx) in
- let term = Termops.it_mkLambda_or_LetIn (Option.get app) (ctx' @ ctx) in
- Some term, termtype
- | Some (Inr (def, subst)) ->
- let termtype = it_mkProd_or_LetIn cty ctx in
- let term = Termops.it_mkLambda_or_LetIn def ctx in
- Some term, termtype
- in
- let _ =
- evars := Evarutil.nf_evar_map !evars;
- evars := Typeclasses.resolve_typeclasses ~filter:Typeclasses.no_goals ~fail:true
- env !evars;
- (* Try resolving fields that are typeclasses automatically. *)
- evars := Typeclasses.resolve_typeclasses ~filter:Typeclasses.all_evars ~fail:false
- env !evars
- in
- let termtype = Evarutil.nf_evar !evars termtype in
- let term = Option.map (Evarutil.nf_evar !evars) term in
- let evm = undefined_evars !evars in
- Evarutil.check_evars env Evd.empty !evars termtype;
- if Evd.is_empty evm && term <> None then
- declare_instance_constant k pri global imps ?hook id (Option.get term) termtype
- else begin
- let kind = Decl_kinds.Global, Decl_kinds.DefinitionBody Decl_kinds.Instance in
- Flags.silently (fun () ->
- Lemmas.start_proof id kind termtype (fun _ -> instance_hook k pri global imps ?hook);
- if term <> None then
- Pfedit.by (!refine_ref (evm, Option.get term))
+ let obls, constr, typ =
+ match term with
+ | Some t ->
+ let obls, _, constr, typ =
+ Obligations.eterm_obligations env id evm 0 t termtype
+ in obls, Some constr, typ
+ | None -> [||], None, termtype
+ in
+ let hook = Lemmas.mk_hook hook in
+ let ctx = Evd.evar_universe_context evm in
+ ignore (Obligations.add_definition id ?term:constr
+ typ ctx ~kind:(Global,poly,Instance) ~hook obls);
+ id
+ else
+ (Flags.silently
+ (fun () ->
+ (* spiwack: it is hard to reorder the actions to do
+ the pretyping after the proof has opened. As a
+ consequence, we use the low-level primitives to code
+ the refinement manually.*)
+ let gls = List.rev (Evd.future_goals evm) in
+ let evm = Evd.reset_future_goals evm in
+ Lemmas.start_proof id kind evm termtype
+ (Lemmas.mk_hook
+ (fun _ -> instance_hook k pri global imps ?hook));
+ (* spiwack: I don't know what to do with the status here. *)
+ if not (Option.is_empty term) then
+ let init_refine =
+ Tacticals.New.tclTHENLIST [
+ Proofview.Refine.refine (fun evm -> evm, Option.get term);
+ Proofview.Unsafe.tclNEWGOALS gls;
+ Tactics.New.reduce_after_refine;
+ ]
+ in
+ ignore (Pfedit.by init_refine)
else if Flags.is_auto_intros () then
- Pfedit.by (Refiner.tclDO len Tactics.intro);
- (match tac with Some tac -> Pfedit.by tac | None -> ())) ();
- Flags.if_verbose (msg $$ Printer.pr_open_subgoals) ();
- id
- end
- end
+ ignore (Pfedit.by (Tacticals.New.tclDO len Tactics.intro));
+ (match tac with Some tac -> ignore (Pfedit.by tac) | None -> ())) ();
+ id)
+ end
+ else Errors.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 -> raise (Invalid_argument "named_of_rel_context") | Name id -> id in
+ 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
(mkVar id :: subst, d :: ctx))
l ([], [])
in ctx
-let string_of_global r =
- string_of_qualid (Nametab.shortest_qualid_of_global Idset.empty r)
-
-let context l =
+let context poly l =
let env = Global.env() in
let evars = ref Evd.empty in
- let _, ((env', fullctx), impls) = interp_context_evars evars env l in
- let fullctx = Evarutil.nf_rel_context_evar !evars fullctx 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
- List.iter (fun (n, b, t) -> Option.iter ce b; ce t) fullctx;
+ let () = List.iter (fun (n, b, t) -> Option.iter ce b; ce t) fullctx in
let ctx =
try named_of_rel_context fullctx
with e when Errors.noncritical e ->
error "Anonymous variables not allowed in contexts."
in
- let fn (id, _, t) =
+ let uctx = Evd.universe_context_set !evars in
+ let fn status (id, b, t) =
if Lib.is_modtype () && not (Lib.sections_are_opened ()) then
- let cst = Declare.declare_constant ~internal:Declare.KernelSilent id
- (ParameterEntry (None,t,None), IsAssumption Logical)
- in
+ let uctx = Univ.ContextSet.to_context uctx in
+ let decl = (ParameterEntry (None,poly,(t,uctx),None), IsAssumption Logical) in
+ let cst = Declare.declare_constant ~internal:Declare.KernelSilent id decl in
match class_of_constr t with
- | Some (rels, (tc, args) as _cl) ->
- add_instance (Typeclasses.new_instance tc None false (ConstRef cst))
+ | Some (rels, ((tc,_), args) as _cl) ->
+ add_instance (Typeclasses.new_instance tc None false (*FIXME*)
+ poly (ConstRef cst));
+ status
(* declare_subclasses (ConstRef cst) cl *)
- | None -> ()
- else (
- let impl = List.exists
- (fun (x,_) ->
- match x with ExplByPos (_, Some id') -> id = id' | _ -> false) impls
+ | None -> status
+ else
+ let test (x, _) = match x with
+ | ExplByPos (_, Some id') -> Id.equal id id'
+ | _ -> false
+ in
+ let impl = List.exists test impls in
+ let decl = (Discharge, poly, Definitional) in
+ let nstatus =
+ pi3 (Command.declare_assumption false decl (t, uctx) [] impl
+ Vernacexpr.NoInline (Loc.ghost, id))
in
- Command.declare_assumption false (Local (* global *), Definitional) t
- [] impl (* implicit *) None (* inline *) (dummy_loc, id))
- in List.iter fn (List.rev ctx)
-
+ status && nstatus
+ in List.fold_left fn true (List.rev ctx)
diff --git a/toplevel/classes.mli b/toplevel/classes.mli
index 97b363c2..0a351d3c 100644
--- a/toplevel/classes.mli
+++ b/toplevel/classes.mli
@@ -1,23 +1,17 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Names
-open Decl_kinds
-open Term
-open Sign
+open Context
open Evd
open Environ
-open Nametab
-open Mod_subst
-open Topconstr
-open Util
+open Constrexpr
open Typeclasses
-open Implicit_quantifiers
open Libnames
(** Errors *)
@@ -26,36 +20,36 @@ val mismatched_params : env -> constr_expr list -> rel_context -> 'a
val mismatched_props : env -> constr_expr list -> rel_context -> 'a
-(** Post-hoc class declaration. *)
-
-val declare_class : reference -> unit
-
(** Instance declaration *)
-val existing_instance : bool -> reference -> unit
+val existing_instance : bool -> reference -> int option -> unit
+(** globality, reference, priority *)
val declare_instance_constant :
typeclass ->
int option -> (** priority *)
bool -> (** globality *)
Impargs.manual_explicitation list -> (** implicits *)
- ?hook:(Libnames.global_reference -> unit) ->
- identifier -> (** name *)
- Term.constr -> (** body *)
+ ?hook:(Globnames.global_reference -> unit) ->
+ Id.t -> (** name *)
+ bool -> (* polymorphic *)
+ Univ.universe_context -> (* Universes *)
+ Constr.t -> (** body *)
Term.types -> (** type *)
- Names.identifier
+ Names.Id.t
val new_instance :
?abstract:bool -> (** Not abstract by default. *)
?global:bool -> (** Not global by default. *)
+ Decl_kinds.polymorphic ->
local_binder list ->
typeclass_constraint ->
- constr_expr option ->
+ (bool * constr_expr) option ->
?generalize:bool ->
- ?tac:Proof_type.tactic ->
- ?hook:(Libnames.global_reference -> unit) ->
+ ?tac:unit Proofview.tactic ->
+ ?hook:(Globnames.global_reference -> unit) ->
int option ->
- identifier
+ Id.t
(** Setting opacity *)
@@ -63,12 +57,10 @@ val set_typeclass_transparency : evaluable_global_reference -> bool -> bool -> u
(** For generation on names based on classes only *)
-val id_of_class : typeclass -> identifier
+val id_of_class : typeclass -> Id.t
(** Context command *)
-val context : local_binder list -> unit
-
-(** Forward ref for refine *)
-
-val refine_ref : (open_constr -> Proof_type.tactic) ref
+(** returns [false] if, for lack of section, it declares an assumption
+ (unless in a module type). *)
+val context : Decl_kinds.polymorphic -> local_binder list -> bool
diff --git a/toplevel/command.ml b/toplevel/command.ml
index 6e1fb49c..9cb3bb86 100644
--- a/toplevel/command.ml
+++ b/toplevel/command.ml
@@ -1,15 +1,18 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Pp
+open Errors
open Util
open Flags
open Term
+open Vars
+open Context
open Termops
open Entries
open Environ
@@ -17,7 +20,10 @@ open Redexpr
open Declare
open Names
open Libnames
+open Globnames
open Nameops
+open Constrexpr
+open Constrexpr_ops
open Topconstr
open Constrintern
open Nametab
@@ -28,11 +34,15 @@ open Decl_kinds
open Pretyping
open Evarutil
open Evarconv
-open Notation
open Indschemes
+open Misctypes
+open Vernacexpr
+
+let do_universe l = Declare.do_universe l
+let do_constraint l = Declare.do_constraint l
let rec under_binders env f n c =
- if n = 0 then f env Evd.empty c else
+ if Int.equal n 0 then snd (f env Evd.empty c) else
match kind_of_term c with
| Lambda (x,t,c) ->
mkLambda (x,t,under_binders (push_rel (x,None,t) env) f (n-1) c)
@@ -43,14 +53,14 @@ let rec under_binders env f n c =
let rec complete_conclusion a cs = function
| CProdN (loc,bl,c) -> CProdN (loc,bl,complete_conclusion a cs c)
| CLetIn (loc,b,t,c) -> CLetIn (loc,b,t,complete_conclusion a cs c)
- | CHole (loc, k) ->
+ | CHole (loc, k, _, _) ->
let (has_no_args,name,params) = a in
if not has_no_args then
user_err_loc (loc,"",
strbrk"Cannot infer the non constant arguments of the conclusion of "
++ pr_id cs ++ str ".");
- let args = List.map (fun id -> CRef(Ident(loc,id))) params in
- CAppExpl (loc,(None,Ident(loc,name)),List.rev args)
+ let args = List.map (fun id -> CRef(Ident(loc,id),None)) params in
+ CAppExpl (loc,(None,Ident(loc,name),None),List.rev args)
| c -> c
(* Commands of the interface *)
@@ -60,154 +70,254 @@ let rec complete_conclusion a cs = function
let red_constant_entry n ce = function
| None -> ce
| Some red ->
- let body = ce.const_entry_body in
- { ce with const_entry_body =
- under_binders (Global.env()) (fst (reduction_of_red_expr red)) n body }
-
-let interp_definition bl red_option c ctypopt =
+ let proof_out = ce.const_entry_body in
+ let env = Global.env () in
+ { ce with const_entry_body = Future.chain ~greedy:true ~pure:true proof_out
+ (fun ((body,ctx),eff) ->
+ (under_binders env
+ (fst (reduction_of_red_expr env red)) n body,ctx),eff) }
+
+let interp_definition bl p red_option c ctypopt =
let env = Global.env() in
- let evdref = ref Evd.empty in
- let impls, ((env_bl, ctx), imps1) = interp_context_evars evdref env bl in
+ let evdref = ref (Evd.from_env env) in
+ let impls, ((env_bl, ctx), imps1) = interp_context_evars env evdref bl in
let nb_args = List.length ctx in
let imps,ce =
match ctypopt with
None ->
- let c, imps2 = interp_constr_evars_impls ~impls ~evdref ~fail_evar:false env_bl c in
- let body = nf_evar !evdref (it_mkLambda_or_LetIn c ctx) in
- check_evars env Evd.empty !evdref body;
- imps1@(Impargs.lift_implicits nb_args imps2),
- { const_entry_body = body;
- const_entry_secctx = None;
- const_entry_type = None;
- const_entry_opaque = false }
+ let subst = evd_comb0 Evd.nf_univ_variables evdref in
+ let ctx = map_rel_context (Vars.subst_univs_constr subst) ctx in
+ let env_bl = push_rel_context ctx env in
+ let c, imps2 = interp_constr_evars_impls ~impls env_bl evdref c in
+ let nf,subst = Evarutil.e_nf_evars_and_universes evdref in
+ let body = nf (it_mkLambda_or_LetIn c ctx) in
+ let vars = Universes.universes_of_constr body in
+ let ctx = Universes.restrict_universe_context
+ (Evd.universe_context_set !evdref) vars in
+ imps1@(Impargs.lift_implicits nb_args imps2),
+ definition_entry ~univs:(Univ.ContextSet.to_context ctx) ~poly:p body
| Some ctyp ->
- let ty, impsty = interp_type_evars_impls ~impls ~evdref ~fail_evar:false env_bl ctyp in
- let c, imps2 = interp_casted_constr_evars_impls ~impls ~evdref ~fail_evar:false env_bl c ty in
- let body = nf_evar !evdref (it_mkLambda_or_LetIn c ctx) in
- let typ = nf_evar !evdref (it_mkProd_or_LetIn ty ctx) in
- check_evars env Evd.empty !evdref body;
- check_evars env Evd.empty !evdref typ;
- (* Check that all implicit arguments inferable from the term is inferable from the type *)
- if not (try List.for_all (fun (key,va) -> List.assoc key impsty = va) imps2 with Not_found -> false)
- then warn (str "Implicit arguments declaration relies on type." ++
- spc () ++ str "The term declares more implicits than the type here.");
+ 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 env_bl = push_rel_context ctx env in
+ let c, imps2 = interp_casted_constr_evars_impls ~impls env_bl evdref c ty in
+ let nf, subst = Evarutil.e_nf_evars_and_universes evdref in
+ let body = nf (it_mkLambda_or_LetIn c ctx) in
+ let typ = nf (it_mkProd_or_LetIn ty ctx) in
+ let beq b1 b2 = if b1 then b2 else not b2 in
+ let impl_eq (x,y,z) (x',y',z') = beq x x' && beq y y' && beq z z' in
+ (* Check that all implicit arguments inferable from the term
+ are inferable from the type *)
+ let chk (key,va) =
+ impl_eq (List.assoc_f Pervasives.(=) key impsty) va (* FIXME *)
+ in
+ if not (try List.for_all chk imps2 with Not_found -> false)
+ then msg_warning
+ (strbrk "Implicit arguments declaration relies on type." ++ spc () ++
+ strbrk "The term declares more implicits than the type here.");
+ let vars = Univ.LSet.union (Universes.universes_of_constr body)
+ (Universes.universes_of_constr typ) in
+ let ctx = Universes.restrict_universe_context
+ (Evd.universe_context_set !evdref) vars in
imps1@(Impargs.lift_implicits nb_args impsty),
- { const_entry_body = body;
- const_entry_secctx = None;
- const_entry_type = Some typ;
- const_entry_opaque = false }
+ definition_entry ~types:typ ~poly:p
+ ~univs:(Univ.ContextSet.to_context ctx) body
in
- red_constant_entry (rel_context_length ctx) ce red_option, imps
+ red_constant_entry (rel_context_length ctx) ce red_option, !evdref, imps
+
+let check_definition (ce, evd, imps) =
+ check_evars_are_solved (Global.env ()) evd (Evd.empty,evd);
+ ce
+
+let get_locality id = 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
+ true
+| Local -> true
+| Global -> false
let declare_global_definition ident ce local k imps =
- let kn = declare_constant ident (DefinitionEntry ce,IsDefinition k) in
+ let local = get_locality ident local in
+ let kn = declare_constant ident ~local (DefinitionEntry ce, IsDefinition k) in
let gr = ConstRef kn in
- maybe_declare_manual_implicits false gr imps;
- if local = Local && Flags.is_verbose() then
- msg_warning (pr_id ident ++ str" is declared as a global definition");
- definition_message ident;
- Autoinstance.search_declaration (ConstRef kn);
- gr
+ let () = maybe_declare_manual_implicits false gr imps in
+ let () = definition_message ident in
+ gr
let declare_definition_hook = ref ignore
let set_declare_definition_hook = (:=) declare_definition_hook
let get_declare_definition_hook () = !declare_definition_hook
-let declare_definition ident (local,k) ce imps hook =
- !declare_definition_hook ce;
+let declare_definition ident (local, p, k) ce imps hook =
+ let () = !declare_definition_hook ce in
let r = match local with
- | Local when Lib.sections_are_opened () ->
- let c =
- SectionLocalDef(ce.const_entry_body ,ce.const_entry_type,false) in
- let _ = declare_variable ident (Lib.cwd(),c,IsDefinition k) in
- definition_message ident;
- if Pfedit.refining () then
- Flags.if_warn msg_warning
- (str"Local definition " ++ pr_id ident ++
- str" is not visible from current goals");
- VarRef ident
- | (Global|Local) ->
- declare_global_definition ident ce local k imps in
- hook local r
+ | Discharge when Lib.sections_are_opened () ->
+ let c = SectionLocalDef ce in
+ let _ = declare_variable ident (Lib.cwd(), c, IsDefinition k) in
+ let () = definition_message ident in
+ let gr = VarRef ident in
+ let () = maybe_declare_manual_implicits false gr imps in
+ let () = if Pfedit.refining () then
+ let msg = strbrk "Section definition " ++
+ pr_id ident ++ strbrk " is not visible from current goals" in
+ msg_warning msg
+ in
+ gr
+ | Discharge | Local | Global ->
+ declare_global_definition ident ce local k imps in
+ Lemmas.call_hook (Future.fix_exn_of ce.Entries.const_entry_body) hook local r
+
+let _ = Obligations.declare_definition_ref := declare_definition
+
+let do_definition ident k bl red_option c ctypopt hook =
+ let (ce, evd, imps as def) = interp_definition bl (pi2 k) red_option c ctypopt in
+ if Flags.is_program_mode () then
+ let env = Global.env () in
+ let (c,ctx), sideff = Future.force ce.const_entry_body in
+ assert(Declareops.side_effects_is_empty sideff);
+ assert(Univ.ContextSet.is_empty ctx);
+ let typ = match ce.const_entry_type with
+ | Some t -> t
+ | None -> Retyping.get_type_of env evd c
+ in
+ Obligations.check_evars env evd;
+ let obls, _, c, cty =
+ Obligations.eterm_obligations env ident evd 0 c typ
+ in
+ let ctx = Evd.evar_universe_context evd in
+ ignore(Obligations.add_definition
+ ident ~term:c cty ctx ~implicits:imps ~kind:k ~hook obls)
+ else let ce = check_definition def in
+ ignore(declare_definition ident k ce imps
+ (Lemmas.mk_hook
+ (fun l r -> Lemmas.call_hook (fun exn -> exn) hook l r;r)))
(* 2| Variable/Hypothesis/Parameter/Axiom declarations *)
-let declare_assumption is_coe (local,kind) c imps impl nl (_,ident) =
- let r = match local with
- | Local when Lib.sections_are_opened () ->
- let _ =
- declare_variable ident
- (Lib.cwd(), SectionLocalAssum (c,impl), IsAssumption kind) in
- assumption_message ident;
- if is_verbose () & Pfedit.refining () then
- msgerrnl (str"Warning: Variable " ++ pr_id ident ++
- str" is not visible from current goals");
- let r = VarRef ident in
- Typeclasses.declare_instance None true r; r
- | (Global|Local) ->
- let kn =
- declare_constant ident
- (ParameterEntry (None,c,nl), IsAssumption kind) in
- let gr = ConstRef kn in
- maybe_declare_manual_implicits false gr imps;
- assumption_message ident;
- if local=Local & Flags.is_verbose () then
- msg_warning (pr_id ident ++ str" is declared as a parameter" ++
- str" because it is at a global level");
- Autoinstance.search_declaration (ConstRef kn);
- Typeclasses.declare_instance None false gr;
- gr
+let declare_assumption is_coe (local,p,kind) (c,ctx) imps impl nl (_,ident) = match local with
+| Discharge when Lib.sections_are_opened () ->
+ let decl = (Lib.cwd(), SectionLocalAssum ((c,ctx),p,impl), IsAssumption kind) in
+ let _ = declare_variable ident decl in
+ let () = assumption_message ident in
+ let () =
+ if is_verbose () && Pfedit.refining () then
+ msg_warning (str"Variable" ++ spc () ++ pr_id ident ++
+ strbrk " is not visible from current goals")
in
- if is_coe then Class.try_add_new_coercion r local
-
-let declare_assumptions_hook = ref ignore
-let set_declare_assumptions_hook = (:=) declare_assumptions_hook
+ let r = VarRef ident in
+ let () = Typeclasses.declare_instance None true r in
+ let () = if is_coe then Class.try_add_new_coercion r ~local:true false in
+ (r,Univ.Instance.empty,true)
+
+| Global | Local | Discharge ->
+ let local = get_locality ident local in
+ let inl = match nl with
+ | NoInline -> None
+ | DefaultInline -> Some (Flags.get_inline_level())
+ | InlineAt i -> Some i
+ in
+ let ctx = Univ.ContextSet.to_context ctx in
+ let decl = (ParameterEntry (None,p,(c,ctx),inl), IsAssumption kind) in
+ let kn = declare_constant ident ~local decl in
+ let gr = ConstRef kn in
+ let () = maybe_declare_manual_implicits false gr imps in
+ let () = assumption_message ident in
+ let () = Typeclasses.declare_instance None false gr in
+ let () = if is_coe then Class.try_add_new_coercion gr local p in
+ let inst =
+ if p (* polymorphic *) then Univ.UContext.instance ctx
+ else Univ.Instance.empty
+ in
+ (gr,inst,Lib.is_modtype_strict ())
-let interp_assumption bl c =
+let interp_assumption evdref env bl c =
let c = prod_constr_expr c bl in
- let env = Global.env () in
- interp_type_evars_impls env c
+ let ty, impls = interp_type_evars_impls env evdref c in
+ let evd, nf = nf_evars_and_universes !evdref in
+ let ctx = Evd.universe_context_set evd in
+ ((nf ty, ctx), impls)
let declare_assumptions idl is_coe k c imps impl_is_on nl =
- !declare_assumptions_hook c;
- List.iter (declare_assumption is_coe k c imps impl_is_on nl) idl
+ let refs, status =
+ List.fold_left (fun (refs,status) id ->
+ let ref',u',status' = declare_assumption is_coe k c imps impl_is_on nl id in
+ (ref',u')::refs, status' && status) ([],true) idl in
+ List.rev refs, status
+
+let do_assumptions (_, poly, _ as kind) nl l =
+ let env = Global.env () in
+ let evdref = ref (Evd.from_env env) in
+ let l =
+ if poly then
+ (* Separate declarations so that A B : Type puts A and B in different levels. *)
+ List.fold_right (fun (is_coe,(idl,c)) acc ->
+ List.fold_right (fun id acc ->
+ (is_coe, ([id], c)) :: acc) idl acc)
+ l []
+ else l
+ in
+ let _,l = List.fold_map (fun env (is_coe,(idl,c)) ->
+ let (t,ctx),imps = interp_assumption evdref env [] c in
+ let env =
+ push_named_context (List.map (fun (_,id) -> (id,None,t)) idl) env in
+ (env,((is_coe,idl),t,(ctx,imps))))
+ env l
+ in
+ let evd = solve_remaining_evars all_and_fail_flags env !evdref (Evd.empty,!evdref) 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)) ->
+ 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)
(* 3a| Elimination schemes for mutual inductive definitions *)
(* 3b| Mutual inductive definitions *)
-let push_named_types env idl tl =
- List.fold_left2 (fun env id t -> Environ.push_named (id,None,t) env)
- env idl tl
-
let push_types env idl tl =
List.fold_left2 (fun env id t -> Environ.push_rel (Name id,None,t) env)
env idl tl
type structured_one_inductive_expr = {
- ind_name : identifier;
+ ind_name : Id.t;
ind_arity : constr_expr;
- ind_lc : (identifier * constr_expr) list
+ ind_lc : (Id.t * constr_expr) list
}
type structured_inductive_expr =
local_binder list * structured_one_inductive_expr list
-let minductive_message = function
+let minductive_message warn = function
| [] -> error "No inductive definition."
- | [x] -> (pr_id x ++ str " is defined")
+ | [x] -> (pr_id x ++ str " is defined" ++
+ if warn then str " as a non-primitive record" else mt())
| l -> hov 0 (prlist_with_sep pr_comma pr_id l ++
spc () ++ str "are defined")
let check_all_names_different indl =
let ind_names = List.map (fun ind -> ind.ind_name) indl in
- let cstr_names = list_map_append (fun ind -> List.map fst ind.ind_lc) indl in
- let l = list_duplicates ind_names in
- if l <> [] then raise (InductiveError (SameNamesTypes (List.hd l)));
- let l = list_duplicates cstr_names in
- if l <> [] then raise (InductiveError (SameNamesConstructors (List.hd l)));
- let l = list_intersect ind_names cstr_names in
- if l <> [] then raise (InductiveError (SameNamesOverlap l))
+ let cstr_names = List.map_append (fun ind -> List.map fst ind.ind_lc) indl in
+ let l = List.duplicates Id.equal ind_names in
+ let () = match l with
+ | [] -> ()
+ | t :: _ -> raise (InductiveError (SameNamesTypes t))
+ in
+ let l = List.duplicates Id.equal cstr_names in
+ let () = match l with
+ | [] -> ()
+ | c :: _ -> raise (InductiveError (SameNamesConstructors (List.hd l)))
+ in
+ let l = List.intersect Id.equal ind_names cstr_names in
+ match l with
+ | [] -> ()
+ | _ -> raise (InductiveError (SameNamesOverlap l))
let mk_mltype_data evdref env assums arity indname =
let is_ml_type = is_sort env !evdref arity in
@@ -217,39 +327,176 @@ let prepare_param = function
| (na,None,t) -> out_name na, LocalAssum t
| (na,Some b,_) -> out_name na, LocalDef b
-let interp_ind_arity evdref env ind =
- interp_type_evars_impls ~evdref env ind.ind_arity
+(** Make the arity conclusion flexible to avoid generating an upper bound universe now,
+ only if the universe does not appear anywhere else.
+ This is really a hack to stay compatible with the semantics of template polymorphic
+ inductives which are recognized when a "Type" appears at the end of the conlusion in
+ the source syntax. *)
+
+let rec check_anonymous_type ind =
+ let open Glob_term in
+ match ind with
+ | GSort (_, GType []) -> true
+ | GProd (_, _, _, _, e)
+ | GLetIn (_, _, _, e)
+ | GLambda (_, _, _, _, e)
+ | GApp (_, e, _)
+ | GCast (_, e, _) -> check_anonymous_type e
+ | _ -> false
+
+let make_conclusion_flexible evdref ty poly =
+ if poly && isArity ty then
+ let _, concl = destArity ty in
+ match concl with
+ | Type u ->
+ (match Univ.universe_level u with
+ | Some u ->
+ evdref := Evd.make_flexible_variable !evdref true u
+ | None -> ())
+ | _ -> ()
+ else ()
+
+let is_impredicative env u =
+ u = Prop Null ||
+ (engagement env = Some Declarations.ImpredicativeSet && u = Prop Pos)
+
+let interp_ind_arity env evdref ind =
+ let c = intern_gen IsType env ind.ind_arity in
+ let imps = Implicit_quantifiers.implicits_of_glob_constr ~with_products:true c in
+ let t, impls = understand_tcc_evars env evdref ~expected_type:IsType c, imps in
+ let pseudo_poly = check_anonymous_type c in
+ let () = if not (Reduction.is_arity env t) then
+ user_err_loc (constr_loc ind.ind_arity, "", str "Not an arity")
+ in
+ t, pseudo_poly, impls
let interp_cstrs evdref env impls mldata arity ind =
let cnames,ctyps = List.split ind.ind_lc in
(* Complete conclusions of constructor types if given in ML-style syntax *)
let ctyps' = List.map2 (complete_conclusion mldata) cnames ctyps in
(* Interpret the constructor types *)
- let ctyps'', cimpls = List.split (List.map (interp_type_evars_impls ~evdref env ~impls) ctyps') in
+ let ctyps'', cimpls = List.split (List.map (interp_type_evars_impls evdref env ~impls) ctyps') in
(cnames, ctyps'', cimpls)
-let interp_mutual_inductive (paramsl,indl) notations finite =
+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)))
+ in
+ let u = univ_of_sort s in
+ (Univ.sup u lev, push_rel d env))
+ sign (Univ.type0m_univ,env))
+
+let sup_list min = List.fold_left Univ.sup min
+
+let extract_level env evd min tys =
+ let sorts = List.map (fun ty ->
+ let ctx, concl = Reduction.dest_prod_assum env ty in
+ sign_level env evd ((Anonymous, None, concl) :: ctx)) tys
+ in sup_list min sorts
+
+let inductive_levels env evdref poly arities inds =
+ let destarities = List.map (Reduction.dest_arity env) arities in
+ let levels = List.map (fun (ctx,a) ->
+ if a = Prop Null then None
+ else Some (univ_of_sort a)) destarities
+ in
+ let cstrs_levels, min_levels, sizes =
+ CList.split3
+ (List.map2 (fun (_,tys,_) (ctx,du) ->
+ let len = List.length tys in
+ let minlev =
+ if len > 1 && not (is_impredicative env du) then
+ Univ.type0_univ
+ else Univ.type0m_univ
+ in
+ let clev = extract_level env !evdref minlev tys in
+ (clev, minlev, len)) inds destarities)
+ in
+ (* Take the transitive closure of the system of constructors *)
+ (* level constraints and remove the recursive dependencies *)
+ let levels' = Universes.solve_constraints_system (Array.of_list levels)
+ (Array.of_list cstrs_levels) (Array.of_list min_levels)
+ in
+ let evd =
+ CList.fold_left3 (fun evd cu (ctx,du) len ->
+ if is_impredicative env du then
+ (** Any product is allowed here. *)
+ evd
+ else (** If in a predicative sort, or asked to infer the type,
+ we take the max of:
+ - indices (if in indices-matter mode)
+ - constructors
+ - Type(1) if there is more than 1 constructor
+ *)
+ let evd =
+ (** Indices contribute. *)
+ if Indtypes.is_indices_matter () && List.length ctx > 0 then (
+ let ilev = sign_level env !evdref ctx in
+ Evd.set_leq_sort env evd (Type ilev) du)
+ else evd
+ in
+ (** Constructors contribute. *)
+ let evd =
+ if Sorts.is_set du then
+ if not (Evd.check_leq evd cu Univ.type0_univ) then
+ raise (Indtypes.InductiveError Indtypes.LargeNonPropInductiveNotInType)
+ else evd
+ else
+ Evd.set_leq_sort env evd (Type cu) du
+ in
+ let evd =
+ if len >= 2 && Univ.is_type0m_univ cu then
+ (** "Polymorphic" type constraint and more than one constructor,
+ should not land in Prop. Add constraint only if it would
+ land in Prop directly (no informative arguments as well). *)
+ Evd.set_leq_sort env evd (Prop Pos) du
+ else evd
+ in evd)
+ !evdref (Array.to_list levels') destarities sizes
+ in evdref := evd; arities
+
+let check_named (loc, na) = match na with
+| Name _ -> ()
+| Anonymous ->
+ let msg = str "Parameters must be named." in
+ user_err_loc (loc, "", msg)
+
+
+let check_param = function
+| LocalRawDef (na, _) -> check_named na
+| LocalRawAssum (nas, Default _, _) -> List.iter check_named nas
+| LocalRawAssum (nas, Generalized _, _) -> ()
+
+let interp_mutual_inductive (paramsl,indl) notations poly prv finite =
check_all_names_different indl;
+ List.iter check_param paramsl;
let env0 = Global.env() in
- let evdref = ref Evd.empty in
+ let evdref = ref Evd.(from_env env0) in
let _, ((env_params, ctx_params), userimpls) =
- interp_context_evars evdref env0 paramsl
+ interp_context_evars env0 evdref paramsl
in
let indnames = List.map (fun ind -> ind.ind_name) indl in
(* Names of parameters as arguments of the inductive type (defs removed) *)
- let assums = List.filter(fun (_,b,_) -> b=None) ctx_params in
+ let assums = List.filter(fun (_,b,_) -> Option.is_empty b) ctx_params in
let params = List.map (fun (na,_,_) -> out_name na) assums in
(* Interpret the arities *)
- let arities = List.map (interp_ind_arity evdref env_params) indl in
- let fullarities = List.map (fun (c, _) -> it_mkProd_or_LetIn c ctx_params) arities in
+ let arities = List.map (interp_ind_arity env_params evdref) indl in
+
+ let fullarities = List.map (fun (c, _, _) -> it_mkProd_or_LetIn c ctx_params) arities in
let env_ar = push_types env0 indnames fullarities in
let env_ar_params = push_rel_context ctx_params env_ar in
(* Compute interpretation metadatas *)
- let indimpls = List.map (fun (_, impls) -> userimpls @ lift_implicits (rel_context_nhyps ctx_params) impls) arities in
- let arities = List.map fst arities in
+ let indimpls = List.map (fun (_, _, impls) -> userimpls @
+ lift_implicits (rel_context_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
@@ -258,29 +505,38 @@ let interp_mutual_inductive (paramsl,indl) notations finite =
(* Temporary declaration of notations and scopes *)
List.iter (Metasyntax.set_notation_for_interpretation impls) notations;
(* Interpret the constructor types *)
- list_map3 (interp_cstrs evdref env_ar_params impls) mldatas arities indl)
+ List.map3 (interp_cstrs env_ar_params evdref impls) mldatas arities indl)
() in
- (* Instantiate evars and check all are resolved *)
- let evd = consider_remaining_unif_problems env_params !evdref in
- let evd = Typeclasses.resolve_typeclasses ~filter:(Typeclasses.no_goals) ~fail:true env_params evd in
- let sigma = evd in
- let constructors = List.map (fun (idl,cl,impsl) -> (idl,List.map (nf_evar sigma) cl,impsl)) constructors in
- let ctx_params = Sign.map_rel_context (nf_evar sigma) ctx_params in
- let arities = List.map (nf_evar sigma) arities in
+ (* Try further to solve evars, and instantiate them *)
+ let sigma = solve_remaining_evars all_and_fail_flags env_params !evdref (Evd.empty,!evdref) in
+ evdref := sigma;
+ (* Compute renewed arities *)
+ let nf,_ = e_nf_evars_and_universes evdref in
+ let arities = List.map nf arities in
+ let constructors = List.map (fun (idl,cl,impsl) -> (idl,List.map nf cl,impsl)) constructors in
+ let _ = List.iter2 (fun ty poly -> make_conclusion_flexible evdref ty poly) arities aritypoly in
+ let arities = inductive_levels env_ar_params evdref poly arities constructors in
+ let nf',_ = e_nf_evars_and_universes evdref in
+ let nf x = nf' (nf x) in
+ let arities = List.map nf' arities in
+ let constructors = List.map (fun (idl,cl,impsl) -> (idl,List.map nf' cl,impsl)) constructors in
+ let ctx_params = map_rel_context nf ctx_params in
+ let evd = !evdref in
List.iter (check_evars env_params Evd.empty evd) arities;
- Sign.iter_rel_context (check_evars env0 Evd.empty evd) ctx_params;
+ iter_rel_context (check_evars env0 Evd.empty evd) ctx_params;
List.iter (fun (_,ctyps,_) ->
List.iter (check_evars env_ar_params Evd.empty evd) ctyps)
constructors;
(* Build the inductive entries *)
- let entries = list_map3 (fun ind arity (cnames,ctypes,cimpls) -> {
+ let entries = List.map4 (fun ind arity template (cnames,ctypes,cimpls) -> {
mind_entry_typename = ind.ind_name;
mind_entry_arity = arity;
+ mind_entry_template = template;
mind_entry_consnames = cnames;
mind_entry_lc = ctypes
- }) indl arities constructors in
+ }) indl arities aritypoly constructors in
let impls =
let len = rel_context_nhyps ctx_params in
List.map2 (fun indimpls (_,_,cimpls) ->
@@ -289,24 +545,17 @@ let interp_mutual_inductive (paramsl,indl) notations finite =
in
(* Build the mutual inductive entry *)
{ mind_entry_params = List.map prepare_param ctx_params;
- mind_entry_record = false;
+ mind_entry_record = None;
mind_entry_finite = finite;
- mind_entry_inds = entries },
+ mind_entry_inds = entries;
+ mind_entry_polymorphic = poly;
+ mind_entry_private = if prv then Some false else None;
+ mind_entry_universes = Evd.universe_context evd },
impls
(* Very syntactical equality *)
-let eq_local_binder d1 d2 = match d1,d2 with
- | LocalRawAssum (nal1,k1,c1), LocalRawAssum (nal2,k2,c2) ->
- List.length nal1 = List.length nal2 && k1 = k2 &&
- List.for_all2 (fun (_,na1) (_,na2) -> na1 = na2) nal1 nal2 &&
- Constrextern.is_same_type c1 c2
- | LocalRawDef ((_,id1),c1), LocalRawDef ((_,id2),c2) ->
- id1 = id2 && Constrextern.is_same_type c1 c2
- | _ ->
- false
-
let eq_local_binders bl1 bl2 =
- List.length bl1 = List.length bl2 && List.for_all2 eq_local_binder bl1 bl2
+ List.equal local_binder_eq bl1 bl2
let extract_coercions indl =
let mkqid (_,((_,id),_)) = qualid_of_ident id in
@@ -316,7 +565,7 @@ let extract_coercions indl =
let extract_params indl =
let paramsl = List.map (fun (_,params,_,_) -> params) indl in
match paramsl with
- | [] -> anomaly "empty list of inductive types"
+ | [] -> anomaly (Pp.str "empty list of inductive types")
| params::paramsl ->
if not (List.for_all (eq_local_binders params) paramsl) then error
"Parameters should be syntactically the same for each inductive type.";
@@ -325,7 +574,7 @@ let extract_params indl =
let extract_inductive indl =
List.map (fun ((_,indname),_,ar,lc) -> {
ind_name = indname;
- ind_arity = Option.cata (fun x -> x) (CSort (dummy_loc, Glob_term.GType None)) ar;
+ ind_arity = Option.cata (fun x -> x) (CSort (Loc.ghost,GType [])) ar;
ind_lc = List.map (fun (_,((_,id),t)) -> (id,t)) lc
}) indl
@@ -336,43 +585,64 @@ let extract_mutual_inductive_declaration_components indl =
let indl = extract_inductive indl in
(params,indl), coes, List.flatten ntnl
-let declare_mutual_inductive_with_eliminations isrecord mie impls =
+let is_recursive mie =
+ let rec is_recursive_constructor lift typ =
+ match Term.kind_of_term typ with
+ | Prod (_,arg,rest) ->
+ Termops.dependent (mkRel lift) arg ||
+ is_recursive_constructor (lift+1) rest
+ | LetIn (na,b,t,rest) -> is_recursive_constructor (lift+1) rest
+ | _ -> false
+ in
+ match mie.mind_entry_inds with
+ | [ind] ->
+ let nparams = List.length mie.mind_entry_params in
+ List.exists (fun t -> is_recursive_constructor (nparams+1) t) ind.mind_entry_lc
+ | _ -> false
+
+let declare_mutual_inductive_with_eliminations mie impls =
+ (* spiwack: raises an error if the structure is supposed to be non-recursive,
+ but isn't *)
+ begin match mie.mind_entry_finite with
+ | BiFinite when is_recursive mie ->
+ if Option.has_some mie.mind_entry_record then
+ error ("Records declared with the keywords Record or Structure cannot be recursive." ^
+ "You can, however, define recursive records using the Inductive or CoInductive command.")
+ else
+ error ("Types declared with the keyword Variant cannot be recursive. Recursive types are defined with the Inductive and CoInductive command.")
+ | _ -> ()
+ end;
let names = List.map (fun e -> e.mind_entry_typename) mie.mind_entry_inds in
- let (_,kn) = declare_mind isrecord mie in
+ let (_, kn), prim = declare_mind mie in
let mind = Global.mind_of_delta_kn kn in
- list_iter_i (fun i (indimpls, constrimpls) ->
+ List.iteri (fun i (indimpls, constrimpls) ->
let ind = (mind,i) in
- Autoinstance.search_declaration (IndRef ind);
maybe_declare_manual_implicits false (IndRef ind) indimpls;
- list_iter_i
+ List.iteri
(fun j impls ->
-(* Autoinstance.search_declaration (ConstructRef (ind,j));*)
maybe_declare_manual_implicits false (ConstructRef (ind, succ j)) impls)
constrimpls)
impls;
- if_verbose ppnl (minductive_message names);
- declare_default_schemes mind;
+ 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 mie.mind_entry_private == None
+ then declare_default_schemes mind;
mind
-open Vernacexpr
-
type one_inductive_impls =
Impargs.manual_explicitation list (* for inds *)*
Impargs.manual_explicitation list list (* for constrs *)
-type one_inductive_expr =
- lident * local_binder list * constr_expr option * constructor_expr list
-
-let do_mutual_inductive indl finite =
+let do_mutual_inductive indl poly prv finite =
let indl,coes,ntns = extract_mutual_inductive_declaration_components indl in
(* Interpret the types *)
- let mie,impls = interp_mutual_inductive indl ntns finite in
+ let mie,impls = interp_mutual_inductive indl ntns poly prv finite in
(* Declare the mutual inductive block with its associated schemes *)
- ignore (declare_mutual_inductive_with_eliminations UserVerbose mie impls);
+ ignore (declare_mutual_inductive_with_eliminations mie impls);
(* Declare the possible notations of inductive types *)
List.iter Metasyntax.add_notation_interpretation ntns;
(* Declare the coercions *)
- List.iter (fun qid -> Class.try_add_new_coercion (locate qid) Global) coes
+ List.iter (fun qid -> Class.try_add_new_coercion (locate qid) false poly) coes
(* 3c| Fixpoints and co-fixpoints *)
@@ -389,46 +659,47 @@ let do_mutual_inductive indl finite =
partial_order : ('a * 'a list) list -> ('a * ('a,'a list) union) list
*)
-let rec partial_order = function
+let rec partial_order cmp = function
| [] -> []
| (x,xge)::rest ->
let rec browse res xge' = function
| [] ->
let res = List.map (function
- | (z, Inr zge) when List.mem x zge -> (z, Inr (list_union zge xge'))
+ | (z, Inr zge) when List.mem_f cmp x zge ->
+ (z, Inr (List.union cmp zge xge'))
| r -> r) res in
(x,Inr xge')::res
| y::xge ->
let rec link y =
- try match List.assoc y res with
+ try match List.assoc_f cmp y res with
| Inl z -> link z
| Inr yge ->
- if List.mem x yge then
- let res = List.remove_assoc y res in
+ if List.mem_f cmp x yge then
+ let res = List.remove_assoc_f cmp y res in
let res = List.map (function
| (z, Inl t) ->
- if t = y then (z, Inl x) else (z, Inl t)
+ if cmp t y then (z, Inl x) else (z, Inl t)
| (z, Inr zge) ->
- if List.mem y zge then
- (z, Inr (list_add_set x (list_remove y zge)))
+ if List.mem_f cmp y zge then
+ (z, Inr (List.add_set cmp x (List.remove cmp y zge)))
else
(z, Inr zge)) res in
- browse ((y,Inl x)::res) xge' (list_union xge (list_remove x yge))
+ browse ((y,Inl x)::res) xge' (List.union cmp xge (List.remove cmp x yge))
else
- browse res (list_add_set y (list_union xge' yge)) xge
- with Not_found -> browse res (list_add_set y xge') xge
+ browse res (List.add_set cmp y (List.union cmp xge' yge)) xge
+ with Not_found -> browse res (List.add_set cmp y xge') xge
in link y
- in browse (partial_order rest) [] xge
+ in browse (partial_order cmp rest) [] xge
let non_full_mutual_message x xge y yge isfix rest =
let reason =
- if List.mem x yge then
- string_of_id y^" depends on "^string_of_id x^" but not conversely"
- else if List.mem y xge then
- string_of_id x^" depends on "^string_of_id y^" but not conversely"
+ if Id.List.mem x yge then
+ Id.to_string y^" depends on "^Id.to_string x^" but not conversely"
+ else if Id.List.mem y xge then
+ Id.to_string x^" depends on "^Id.to_string y^" but not conversely"
else
- string_of_id y^" and "^string_of_id x^" are not mutually dependent" in
- let e = if rest <> [] then "e.g.: "^reason else reason in
+ Id.to_string y^" and "^Id.to_string x^" are not mutually dependent" in
+ let e = if List.is_empty rest then reason else "e.g.: "^reason in
let k = if isfix then "fixpoint" else "cofixpoint" in
let w =
if isfix
@@ -441,52 +712,45 @@ let check_mutuality env isfix fixl =
let names = List.map fst fixl in
let preorder =
List.map (fun (id,def) ->
- (id, List.filter (fun id' -> id<>id' & occur_var env id' def) names))
+ (id, List.filter (fun id' -> not (Id.equal id id') && occur_var env id' def) names))
fixl in
- let po = partial_order preorder in
+ let po = partial_order Id.equal preorder in
match List.filter (function (_,Inr _) -> true | _ -> false) po with
| (x,Inr xge)::(y,Inr yge)::rest ->
- if_warn msg_warning (non_full_mutual_message x xge y yge isfix rest)
+ msg_warning (non_full_mutual_message x xge y yge isfix rest)
| _ -> ()
type structured_fixpoint_expr = {
- fix_name : identifier;
- fix_annot : identifier located option;
+ fix_name : Id.t;
+ fix_annot : Id.t Loc.located option;
fix_binders : local_binder list;
fix_body : constr_expr option;
fix_type : constr_expr
}
-let interp_fix_context evdref env isfix fix =
+let interp_fix_context env evdref isfix fix =
let before, after = if isfix then split_at_annot fix.fix_binders fix.fix_annot else [], fix.fix_binders in
- let impl_env, ((env', ctx), imps) = interp_context_evars evdref env before in
- let impl_env', ((env'', ctx'), imps') = interp_context_evars ~impl_env evdref env' after in
+ let impl_env, ((env', ctx), imps) = interp_context_evars env evdref before in
+ let impl_env', ((env'', ctx'), imps') = interp_context_evars ~impl_env env' evdref after in
let annot = Option.map (fun _ -> List.length (assums_of_rel_context ctx)) fix.fix_annot in
((env'', ctx' @ ctx), (impl_env',imps @ imps'), annot)
let interp_fix_ccl evdref impls (env,_) fix =
- interp_type_evars_impls ~impls ~evdref ~fail_evar:false env fix.fix_type
+ interp_type_evars_impls ~impls env evdref fix.fix_type
-let interp_fix_body evdref env_rec impls (_,ctx) fix ccl =
+let interp_fix_body env_rec evdref impls (_,ctx) fix ccl =
Option.map (fun body ->
let env = push_rel_context ctx env_rec in
- let body = interp_casted_constr_evars evdref env ~impls body ccl in
+ let body = interp_casted_constr_evars env evdref ~impls body ccl in
it_mkLambda_or_LetIn body ctx) fix.fix_body
let build_fix_type (_,ctx) ccl = it_mkProd_or_LetIn ccl ctx
-let declare_fix kind f def t imps =
- let ce = {
- const_entry_body = def;
- const_entry_secctx = None;
- const_entry_type = Some t;
- const_entry_opaque = false }
- in
- let kn = declare_constant f (DefinitionEntry ce,IsDefinition kind) in
- let gr = ConstRef kn in
- Autoinstance.search_declaration (ConstRef kn);
- maybe_declare_manual_implicits false gr imps;
- gr
+let declare_fix (_,poly,_ as kind) ctx f ((def,_),eff) t imps =
+ let ce = definition_entry ~types:t ~poly ~univs:ctx ~eff def in
+ declare_definition f kind ce imps (Lemmas.mk_hook (fun _ r -> r))
+
+let _ = Obligations.declare_fix_ref := declare_fix
let prepare_recursive_declaration fixnames fixtypes fixdefs =
let defs = List.map (subst_vars (List.rev fixnames)) fixdefs in
@@ -504,27 +768,236 @@ let compute_possible_guardness_evidences (ids,_,na) =
but doing it properly involves delta-reduction, and it finally
doesn't seem to worth the effort (except for huge mutual
fixpoints ?) *)
- interval 0 (List.length ids - 1)
+ List.interval 0 (List.length ids - 1)
type recursive_preentry =
- identifier list * constr option list * types list
+ Id.t list * constr option list * types list
+
+(* Wellfounded definition *)
+
+open Coqlib
+
+let contrib_name = "Program"
+let subtac_dir = [contrib_name]
+let fixsub_module = subtac_dir @ ["Wf"]
+let tactics_module = subtac_dir @ ["Tactics"]
+
+let init_reference dir s () = Coqlib.gen_reference "Command" dir s
+let init_constant dir s () = Coqlib.gen_constant "Command" dir s
+let make_ref l s = init_reference l s
+let fix_proto = init_constant tactics_module "fix_proto"
+let fix_sub_ref = make_ref fixsub_module "Fix_sub"
+let measure_on_R_ref = make_ref fixsub_module "MR"
+let well_founded = init_constant ["Init"; "Wf"] "well_founded"
+let mkSubset name typ prop =
+ mkApp (Universes.constr_of_global (delayed_force build_sigma).typ,
+ [| typ; mkLambda (name, typ, prop) |])
+let sigT = Lazy.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 ->
+ let ty, tys, (k, constr) =
+ List.fold_left
+ (fun (ty, tys, (k, constr)) (n, b, t) ->
+ let pred = mkLambda (n, t, ty) in
+ let ty = Universes.constr_of_global (Lazy.force sigT).typ in
+ let intro = Universes.constr_of_global (Lazy.force sigT).intro in
+ let sigty = mkApp (ty, [|t; pred|]) in
+ let intro = mkApp (intro, [|lift k t; lift k pred; mkRel k; constr|]) in
+ (sigty, pred :: tys, (succ k, intro)))
+ (t, [], (2, mkRel 1)) tl
+ in
+ let (last, subst) = List.fold_right2
+ (fun pred (n, b, t) (prev, subst) ->
+ 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))
+ (List.rev tys) tl (mkRel 1, [])
+ in ty, ((n, Some 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
+
+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
+
+let build_wellfounded (recname,n,bl,arityc,body) r measure notation =
+ Coqlib.check_required_library ["Coq";"Program";"Wf"];
+ let env = Global.env() in
+ let evdref = ref (Evd.from_env env) in
+ let _, ((env', binders_rel), impls) = interp_context_evars env evdref bl in
+ let len = List.length binders_rel in
+ let top_env = push_rel_context binders_rel env in
+ let top_arity = interp_type_evars top_env evdref arityc in
+ let full_arity = it_mkProd_or_LetIn top_arity binders_rel in
+ let argtyp, letbinders, make = telescope binders_rel in
+ let argname = Id.of_string "recarg" in
+ let arg = (Name argname, None, 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.type_of env !evdref rel in
+ let relargty =
+ let error () =
+ user_err_loc (constr_loc r,
+ "Command.build_wellfounded",
+ Printer.pr_constr_env env !evdref rel ++ str " is not an homogeneous binary relation.")
+ in
+ try
+ let ctx, ar = Reductionops.splay_prod_n env !evdref 2 relty in
+ match ctx, kind_of_term ar with
+ | [(_, None, t); (_, None, u)], Sort (Prop Null)
+ when Reductionops.is_conv env !evdref t u -> t
+ | _, _ -> error ()
+ with e when Errors.noncritical e -> error ()
+ in
+ let measure = interp_casted_constr_evars binders_env evdref measure relargty in
+ let wf_rel, wf_rel_fun, measure_fn =
+ let measure_body, measure =
+ it_mkLambda_or_LetIn measure letbinders,
+ it_mkLambda_or_LetIn measure binders
+ in
+ let comb = Universes.constr_of_global (delayed_force measure_on_R_ref) in
+ let wf_rel = mkApp (comb, [| argtyp; relargty; rel; measure |]) in
+ let wf_rel_fun x y =
+ mkApp (rel, [| subst1 x measure_body;
+ subst1 y measure_body |])
+ in wf_rel, wf_rel_fun, measure
+ in
+ let wf_proof = mkApp (delayed_force well_founded, [| argtyp ; wf_rel |]) in
+ let argid' = Id.of_string (Id.to_string argname ^ "'") in
+ let wfarg len = (Name argid', None,
+ mkSubset (Name argid') argtyp
+ (wf_rel_fun (mkRel 1) (mkRel (len + 1))))
+ in
+ let intern_bl = wfarg 1 :: [arg] in
+ let _intern_env = push_rel_context intern_bl env in
+ let proj = (*FIXME*)Universes.constr_of_global (delayed_force build_sigma).Coqlib.proj1 in
+ let wfargpred = mkLambda (Name argid', argtyp, wf_rel_fun (mkRel 1) (mkRel 3)) in
+ let projection = (* in wfarg :: arg :: before *)
+ mkApp (proj, [| argtyp ; wfargpred ; mkRel 1 |])
+ in
+ let top_arity_let = it_mkLambda_or_LetIn top_arity letbinders in
+ let intern_arity = substl [projection] top_arity_let in
+ (* substitute the projection of wfarg for something,
+ now intern_arity is in wfarg :: arg *)
+ let intern_fun_arity_prod = it_mkProd_or_LetIn intern_arity [wfarg 1] in
+ let intern_fun_binder = (Name (add_suffix recname "'"), None, 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 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)
+ 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 (r, l, impls, scopes) =
+ Constrintern.compute_internalization_data env
+ Constrintern.Recursive full_arity impls
+ in
+ let newimpls = Id.Map.singleton recname
+ (r, l, impls @ [(Some (Id.of_string "recproof", Impargs.Manual, (true, false)))],
+ scopes @ [None]) in
+ interp_casted_constr_evars (push_rel_context ctx env) evdref
+ ~impls:newimpls body (lift 1 top_arity)
+ in
+ let intern_body_lam = it_mkLambda_or_LetIn intern_body (curry_fun :: lift_lets @ fun_bl) in
+ let prop = mkLambda (Name argname, argtyp, top_arity_let) in
+ let def =
+ mkApp (Universes.constr_of_global (delayed_force fix_sub_ref),
+ [| argtyp ; wf_rel ;
+ Evarutil.e_new_evar env evdref
+ ~src:(Loc.ghost, Evar_kinds.QuestionMark (Evar_kinds.Define false)) wf_proof;
+ prop |])
+ in
+ let def = Typing.solve_evars env evdref def in
+ let _ = evdref := Evarutil.nf_evar_map !evdref in
+ let def = mkApp (def, [|intern_body_lam|]) in
+ let binders_rel = nf_evar_context !evdref binders_rel in
+ let binders = nf_evar_context !evdref binders in
+ let top_arity = Evarutil.nf_evar !evdref top_arity in
+ let hook, recname, typ =
+ if List.length binders_rel > 1 then
+ let name = add_suffix recname "_func" in
+ let hook l gr =
+ let body = it_mkLambda_or_LetIn (mkApp (Universes.constr_of_global gr, [|make|])) binders_rel in
+ let ty = it_mkProd_or_LetIn top_arity binders_rel in
+ let univs = Evd.universe_context !evdref in
+ (*FIXME poly? *)
+ let ce = definition_entry ~types:ty ~univs (Evarutil.nf_evar !evdref body) in
+ (** FIXME: include locality *)
+ let c = Declare.declare_constant recname (DefinitionEntry ce, IsDefinition Definition) in
+ let gr = ConstRef c in
+ if Impargs.is_implicit_args () || not (List.is_empty impls) then
+ Impargs.declare_manual_implicits false gr [impls]
+ in
+ let typ = it_mkProd_or_LetIn top_arity binders in
+ hook, name, typ
+ else
+ let typ = it_mkProd_or_LetIn top_arity binders_rel in
+ let hook l gr =
+ if Impargs.is_implicit_args () || not (List.is_empty impls) then
+ Impargs.declare_manual_implicits false gr [impls]
+ in hook, recname, typ
+ in
+ let hook = Lemmas.mk_hook hook in
+ let fullcoqc = Evarutil.nf_evar !evdref def in
+ let fullctyp = Evarutil.nf_evar !evdref typ in
+ Obligations.check_evars env !evdref;
+ let evars, _, evars_def, evars_typ =
+ Obligations.eterm_obligations env recname !evdref 0 fullcoqc fullctyp
+ in
+ let ctx = Evd.evar_universe_context !evdref in
+ ignore(Obligations.add_definition recname ~term:evars_def
+ evars_typ ctx evars ~hook)
let interp_recursive isfix fixl notations =
let env = Global.env() in
let fixnames = List.map (fun fix -> fix.fix_name) fixl in
(* Interp arities allowing for unresolved types *)
- let evdref = ref Evd.empty in
+ let evdref = ref (Evd.from_env env) in
let fixctxs, fiximppairs, fixannots =
- list_split3 (List.map (interp_fix_context evdref env isfix) fixl) in
+ List.split3 (List.map (interp_fix_context env evdref isfix) fixl) in
let fixctximpenvs, fixctximps = List.split fiximppairs in
- let fixccls,fixcclimps = List.split (list_map3 (interp_fix_ccl evdref) fixctximpenvs fixctxs fixl) in
+ let fixccls,fixcclimps = List.split (List.map3 (interp_fix_ccl evdref) fixctximpenvs fixctxs fixl) in
let fixtypes = List.map2 build_fix_type fixctxs fixccls in
let fixtypes = List.map (nf_evar !evdref) fixtypes in
- let fiximps = list_map3
+ let fiximps = List.map3
(fun ctximps cclimps (_,ctx) -> ctximps@(Impargs.lift_implicits (List.length ctx) cclimps))
fixctximps fixcclimps fixctxs in
- let env_rec = push_named_types env fixnames fixtypes in
+ let rec_sign =
+ List.fold_left2
+ (fun env' id t ->
+ if Flags.is_program_mode () then
+ let sort = Evarutil.evd_comb1 (Typing.e_type_of ~refresh:true env) evdref t in
+ 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
+ in
+ (id,None,fixprot) :: env'
+ else (id,None,t) :: env')
+ [] fixnames fixtypes
+ in
+ let env_rec = push_named_context rec_sign env in
(* Get interpretation metadatas *)
let impls = compute_internalization_env env Recursive fixnames fixtypes fiximps in
@@ -533,72 +1006,99 @@ let interp_recursive isfix fixl notations =
let fixdefs =
Metasyntax.with_syntax_protection (fun () ->
List.iter (Metasyntax.set_notation_for_interpretation impls) notations;
- list_map4
- (fun fixctximpenv -> interp_fix_body evdref env_rec (Idmap.fold Idmap.add fixctximpenv impls))
+ List.map4
+ (fun fixctximpenv -> interp_fix_body env_rec evdref (Id.Map.fold Id.Map.add fixctximpenv impls))
fixctximpenvs fixctxs fixl fixccls)
() in
(* Instantiate evars and check all are resolved *)
let evd = consider_remaining_unif_problems env_rec !evdref in
- let fixdefs = List.map (Option.map (nf_evar evd)) fixdefs in
- let fixtypes = List.map (nf_evar evd) fixtypes 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 evd = Typeclasses.resolve_typeclasses ~filter:Typeclasses.no_goals ~fail:true env evd in
- List.iter (Option.iter (check_evars env_rec Evd.empty evd)) fixdefs;
- List.iter (check_evars env Evd.empty evd) fixtypes;
- if not (List.mem None fixdefs) then begin
- let fixdefs = List.map Option.get fixdefs in
- check_mutuality env isfix (List.combine fixnames fixdefs)
- end;
(* Build the fix declaration block *)
- (fixnames,fixdefs,fixtypes), list_combine3 fixctxnames fiximps fixannots
+ (env,rec_sign,evd), (fixnames,fixdefs,fixtypes), List.combine3 fixctxnames fiximps fixannots
-let interp_fixpoint = interp_recursive true
-let interp_cofixpoint = interp_recursive false
-
-let declare_fixpoint ((fixnames,fixdefs,fixtypes),fiximps) indexes ntns =
- if List.mem None fixdefs then
+let check_recursive isfix env evd (fixnames,fixdefs,_) =
+ check_evars_are_solved env evd (Evd.empty,evd);
+ if List.for_all Option.has_some fixdefs then begin
+ let fixdefs = List.map Option.get fixdefs in
+ check_mutuality env isfix (List.combine fixnames fixdefs)
+ end
+
+let interp_fixpoint l ntns =
+ let (env,_,evd),fix,info = interp_recursive true l ntns in
+ check_recursive true env evd fix;
+ (fix,Evd.evar_universe_context evd,info)
+
+let interp_cofixpoint l ntns =
+ let (env,_,evd),fix,info = interp_recursive false l ntns in
+ check_recursive false env evd fix;
+ fix,Evd.evar_universe_context evd,info
+
+let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),ctx,fiximps) indexes ntns =
+ if List.exists Option.is_empty fixdefs then
(* Some bodies to define by proof *)
let thms =
- list_map3 (fun id t (len,imps,_) -> (id,(t,(len,imps)))) fixnames fixtypes fiximps in
+ List.map3 (fun id t (len,imps,_) -> (id,(t,(len,imps)))) fixnames fixtypes fiximps in
let init_tac =
Some (List.map (Option.cata Tacmach.refine_no_check Tacticals.tclIDTAC)
fixdefs) in
- Lemmas.start_proof_with_initialization (Global,DefinitionBody Fixpoint)
- (Some(false,indexes,init_tac)) thms None (fun _ _ -> ())
+ let init_tac =
+ Option.map (List.map Proofview.V82.tactic) init_tac
+ in
+ let evd = Evd.from_env ~ctx Environ.empty_env in
+ Lemmas.start_proof_with_initialization (Global,poly,DefinitionBody Fixpoint)
+ evd (Some(false,indexes,init_tac)) thms None (Lemmas.mk_hook (fun _ _ -> ()))
else begin
(* We shortcut the proof process *)
let fixdefs = List.map Option.get fixdefs in
let fixdecls = prepare_recursive_declaration fixnames fixtypes fixdefs in
- let indexes = search_guard dummy_loc (Global.env()) indexes fixdecls in
+ let env = Global.env() in
+ let indexes = search_guard Loc.ghost env indexes fixdecls in
let fiximps = List.map (fun (n,r,p) -> r) fiximps in
+ let vars = Universes.universes_of_constr (mkFix ((indexes,0),fixdecls)) in
let fixdecls =
- list_map_i (fun i _ -> mkFix ((indexes,i),fixdecls)) 0 fixnames in
- ignore (list_map4 (declare_fix Fixpoint) fixnames fixdecls fixtypes fiximps);
+ List.map_i (fun i _ -> mkFix ((indexes,i),fixdecls)) 0 fixnames in
+ let ctx = Evd.evar_universe_context_set ctx in
+ let ctx = Universes.restrict_universe_context ctx vars in
+ let fixdecls = List.map Term_typing.mk_pure_proof fixdecls in
+ let ctx = Univ.ContextSet.to_context ctx in
+ ignore (List.map4 (declare_fix (local, poly, Fixpoint) ctx)
+ fixnames fixdecls fixtypes fiximps);
(* Declare the recursive definitions *)
fixpoint_message (Some indexes) fixnames;
end;
(* Declare notations *)
List.iter Metasyntax.add_notation_interpretation ntns
-let declare_cofixpoint ((fixnames,fixdefs,fixtypes),fiximps) ntns =
- if List.mem None fixdefs then
+let declare_cofixpoint local poly ((fixnames,fixdefs,fixtypes),ctx,fiximps) ntns =
+ if List.exists Option.is_empty fixdefs then
(* Some bodies to define by proof *)
let thms =
- list_map3 (fun id t (len,imps,_) -> (id,(t,(len,imps)))) fixnames fixtypes fiximps in
+ List.map3 (fun id t (len,imps,_) -> (id,(t,(len,imps)))) fixnames fixtypes fiximps in
let init_tac =
Some (List.map (Option.cata Tacmach.refine_no_check Tacticals.tclIDTAC)
fixdefs) in
- Lemmas.start_proof_with_initialization (Global,DefinitionBody CoFixpoint)
- (Some(true,[],init_tac)) thms None (fun _ _ -> ())
+ let init_tac =
+ Option.map (List.map Proofview.V82.tactic) init_tac
+ in
+ let evd = Evd.from_env ~ctx Environ.empty_env in
+ Lemmas.start_proof_with_initialization (Global,poly, DefinitionBody CoFixpoint)
+ evd (Some(true,[],init_tac)) thms None (Lemmas.mk_hook (fun _ _ -> ()))
else begin
(* We shortcut the proof process *)
let fixdefs = List.map Option.get fixdefs in
let fixdecls = prepare_recursive_declaration fixnames fixtypes fixdefs in
- let fixdecls = list_map_i (fun i _ -> mkCoFix (i,fixdecls)) 0 fixnames in
+ let fixdecls = List.map_i (fun i _ -> mkCoFix (i,fixdecls)) 0 fixnames in
+ let fixdecls = List.map Term_typing.mk_pure_proof fixdecls in
let fiximps = List.map (fun (len,imps,idx) -> imps) fiximps in
- ignore (list_map4 (declare_fix CoFixpoint) fixnames fixdecls fixtypes fiximps);
+ let ctx = Evd.evar_universe_context_set ctx in
+ let ctx = Univ.ContextSet.to_context ctx in
+ ignore (List.map4 (declare_fix (local, poly, CoFixpoint) ctx)
+ fixnames fixdecls fixtypes fiximps);
(* Declare the recursive definitions *)
cofixpoint_message fixnames
end;
@@ -624,13 +1124,91 @@ let extract_cofixpoint_components l =
{fix_name = id; fix_annot = None; fix_binders = bl; fix_body = def; fix_type = typ}) fixl,
List.flatten ntnl
-let do_fixpoint l =
- let fixl,ntns = extract_fixpoint_components true l in
- let fix = interp_fixpoint fixl ntns in
- let possible_indexes =
- List.map compute_possible_guardness_evidences (snd fix) in
- declare_fixpoint fix possible_indexes ntns
+let out_def = function
+ | Some def -> def
+ | None -> error "Program Fixpoint needs defined bodies."
-let do_cofixpoint l =
+let do_program_recursive local p fixkind fixl ntns =
+ let isfix = fixkind != Obligations.IsCoFixpoint in
+ let (env, rec_sign, evd), fix, info =
+ interp_recursive isfix fixl ntns
+ in
+ (* Program-specific code *)
+ (* Get the interesting evars, those that were not instanciated *)
+ let evd = Typeclasses.resolve_typeclasses ~filter:Typeclasses.no_goals ~fail:true env evd in
+ (* Solve remaining evars *)
+ let evd = nf_evar_map_undefined evd in
+ let collect_evars id def typ imps =
+ (* Generalize by the recursive prototypes *)
+ let def =
+ nf_evar evd (Termops.it_mkNamedLambda_or_LetIn def rec_sign)
+ and typ =
+ nf_evar evd (Termops.it_mkNamedProd_or_LetIn typ rec_sign)
+ in
+ let evars, _, def, typ =
+ Obligations.eterm_obligations env id evd
+ (List.length rec_sign) def typ
+ in (id, def, typ, imps, evars)
+ in
+ let (fixnames,fixdefs,fixtypes) = fix in
+ let fiximps = List.map pi2 info in
+ let fixdefs = List.map out_def fixdefs in
+ let defs = List.map4 collect_evars fixnames fixdefs fixtypes fiximps in
+ let () = if isfix then begin
+ let possible_indexes = List.map compute_possible_guardness_evidences info in
+ let fixdecls = Array.of_list (List.map (fun x -> Name x) fixnames),
+ Array.of_list fixtypes,
+ Array.of_list (List.map (subst_vars (List.rev fixnames)) fixdefs)
+ in
+ let indexes =
+ Pretyping.search_guard Loc.ghost (Global.env ()) possible_indexes fixdecls in
+ List.iteri (fun i _ -> Inductive.check_fix env ((indexes,i),fixdecls)) fixl
+ end in
+ let ctx = Evd.evar_universe_context evd in
+ let kind = match fixkind with
+ | Obligations.IsFixpoint _ -> (local, p, Fixpoint)
+ | Obligations.IsCoFixpoint -> (local, p, CoFixpoint)
+ in
+ Obligations.add_mutual_definitions defs ~kind 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)] ->
+ 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
+
+ | [(n, CMeasureRec (m, r))], [(((_,id),_,bl,typ,def),ntn)] ->
+ build_wellfounded (id, n, bl, typ, out_def def)
+ (Option.default (CRef (lt_ref,None)) r) m ntn
+
+ | _, _ when List.for_all (fun (n, ro) -> ro == CStructRec) g ->
+ let fixl,ntns = extract_fixpoint_components true l in
+ let fixkind = Obligations.IsFixpoint g in
+ do_program_recursive local poly fixkind fixl ntns
+
+ | _, _ ->
+ errorlabstrm "do_program_fixpoint"
+ (str "Well-founded fixpoints not allowed in mutually recursive blocks")
+
+let do_fixpoint local poly l =
+ if Flags.is_program_mode () then do_program_fixpoint local poly l
+ else
+ let fixl, ntns = extract_fixpoint_components true l in
+ let fix = interp_fixpoint fixl ntns in
+ let possible_indexes =
+ List.map compute_possible_guardness_evidences (pi3 fix) in
+ declare_fixpoint local poly fix possible_indexes ntns
+
+let do_cofixpoint local poly l =
let fixl,ntns = extract_cofixpoint_components l in
- declare_cofixpoint (interp_cofixpoint fixl ntns) ntns
+ if Flags.is_program_mode () then
+ do_program_recursive local poly Obligations.IsCoFixpoint fixl ntns
+ else
+ let cofix = interp_cofixpoint fixl ntns in
+ declare_cofixpoint local poly cofix ntns
diff --git a/toplevel/command.mli b/toplevel/command.mli
index 456026bf..894333ad 100644
--- a/toplevel/command.mli
+++ b/toplevel/command.mli
@@ -1,54 +1,68 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Util
open Names
open Term
open Entries
open Libnames
+open Globnames
open Tacexpr
open Vernacexpr
-open Topconstr
+open Constrexpr
open Decl_kinds
open Redexpr
-open Constrintern
open Pfedit
(** This file is about the interpretation of raw commands into typed
ones and top-level declaration of the main Gallina objects *)
+val do_universe : Id.t Loc.located list -> unit
+val do_constraint : (Id.t Loc.located * Univ.constraint_type * Id.t Loc.located) list -> unit
+
(** {6 Hooks for Pcoq} *)
val set_declare_definition_hook : (definition_entry -> unit) -> unit
val get_declare_definition_hook : unit -> (definition_entry -> unit)
-val set_declare_assumptions_hook : (types -> unit) -> unit
(** {6 Definitions/Let} *)
val interp_definition :
- local_binder list -> red_expr option -> constr_expr ->
- constr_expr option -> definition_entry * Impargs.manual_implicits
+ local_binder list -> polymorphic -> red_expr option -> constr_expr ->
+ constr_expr option -> definition_entry * Evd.evar_map * Impargs.manual_implicits
-val declare_definition : identifier -> locality * definition_object_kind ->
- definition_entry -> Impargs.manual_implicits -> declaration_hook -> unit
+val declare_definition : Id.t -> definition_kind ->
+ definition_entry -> Impargs.manual_implicits ->
+ Globnames.global_reference Lemmas.declaration_hook -> Globnames.global_reference
+
+val do_definition : Id.t -> definition_kind ->
+ local_binder list -> red_expr option -> constr_expr ->
+ constr_expr option -> unit Lemmas.declaration_hook -> unit
(** {6 Parameters/Assumptions} *)
-val interp_assumption :
- local_binder list -> constr_expr -> types * Impargs.manual_implicits
+(* val interp_assumption : env -> evar_map ref -> *)
+(* local_binder list -> constr_expr -> *)
+(* types Univ.in_universe_context_set * Impargs.manual_implicits *)
-val declare_assumption : coercion_flag -> assumption_kind -> types ->
+(** returns [false] if the assumption is neither local to a section,
+ nor in a module type and meant to be instantiated. *)
+val declare_assumption : coercion_flag -> assumption_kind ->
+ types Univ.in_universe_context_set ->
Impargs.manual_implicits ->
- bool (** implicit *) -> Entries.inline -> variable located -> unit
+ bool (** implicit *) -> Vernacexpr.inline -> variable Loc.located ->
+ global_reference * Univ.Instance.t * bool
+
+val do_assumptions : locality * polymorphic * assumption_object_kind ->
+ Vernacexpr.inline -> simple_binder with_coercion list -> bool
-val declare_assumptions : variable located list ->
- coercion_flag -> assumption_kind -> types -> Impargs.manual_implicits ->
- bool -> Entries.inline -> unit
+(* val declare_assumptions : variable Loc.located list -> *)
+(* coercion_flag -> assumption_kind -> types Univ.in_universe_context_set -> *)
+(* Impargs.manual_implicits -> bool -> Vernacexpr.inline -> bool *)
(** {6 Inductive and coinductive types} *)
@@ -56,9 +70,9 @@ val declare_assumptions : variable located list ->
inductive declarations *)
type structured_one_inductive_expr = {
- ind_name : identifier;
+ ind_name : Id.t;
ind_arity : constr_expr;
- ind_lc : (identifier * constr_expr) list
+ ind_lc : (Id.t * constr_expr) list
}
type structured_inductive_expr =
@@ -75,26 +89,28 @@ type one_inductive_impls =
Impargs.manual_implicits list (** for constrs *)
val interp_mutual_inductive :
- structured_inductive_expr -> decl_notation list -> bool ->
+ structured_inductive_expr -> decl_notation list -> polymorphic ->
+ private_flag -> Decl_kinds.recursivity_kind ->
mutual_inductive_entry * one_inductive_impls list
(** Registering a mutual inductive definition together with its
associated schemes *)
val declare_mutual_inductive_with_eliminations :
- Declare.internal_flag -> mutual_inductive_entry -> one_inductive_impls list ->
+ mutual_inductive_entry -> one_inductive_impls list ->
mutual_inductive
(** Entry points for the vernacular commands Inductive and CoInductive *)
val do_mutual_inductive :
- (one_inductive_expr * decl_notation list) list -> bool -> unit
+ (one_inductive_expr * decl_notation list) list -> polymorphic ->
+ private_flag -> Decl_kinds.recursivity_kind -> unit
(** {6 Fixpoints and cofixpoints} *)
type structured_fixpoint_expr = {
- fix_name : identifier;
- fix_annot : identifier located option;
+ fix_name : Id.t;
+ fix_annot : Id.t Loc.located option;
fix_binders : local_binder list;
fix_body : constr_expr option;
fix_type : constr_expr
@@ -114,37 +130,42 @@ val extract_cofixpoint_components :
(** Typing global fixpoints and cofixpoint_expr *)
type recursive_preentry =
- identifier list * constr option list * types list
+ Id.t list * constr option list * types list
val interp_fixpoint :
structured_fixpoint_expr list -> decl_notation list ->
- recursive_preentry * (name list * Impargs.manual_implicits * int option) list
+ recursive_preentry * Evd.evar_universe_context *
+ (Name.t list * Impargs.manual_implicits * int option) list
val interp_cofixpoint :
structured_fixpoint_expr list -> decl_notation list ->
- recursive_preentry * (name list * Impargs.manual_implicits * int option) list
+ recursive_preentry * Evd.evar_universe_context *
+ (Name.t list * Impargs.manual_implicits * int option) list
(** Registering fixpoints and cofixpoints in the environment *)
val declare_fixpoint :
- recursive_preentry * (name list * Impargs.manual_implicits * int option) list ->
+ locality -> polymorphic ->
+ recursive_preentry * Evd.evar_universe_context *
+ (Name.t list * Impargs.manual_implicits * int option) list ->
lemma_possible_guards -> decl_notation list -> unit
-val declare_cofixpoint :
- recursive_preentry * (name list * Impargs.manual_implicits * int option) list ->
- decl_notation list -> unit
+val declare_cofixpoint : locality -> polymorphic ->
+ recursive_preentry * Evd.evar_universe_context *
+ (Name.t list * Impargs.manual_implicits * int option) list ->
+ decl_notation list -> unit
(** Entry points for the vernacular commands Fixpoint and CoFixpoint *)
val do_fixpoint :
- (fixpoint_expr * decl_notation list) list -> unit
+ locality -> polymorphic -> (fixpoint_expr * decl_notation list) list -> unit
val do_cofixpoint :
- (cofixpoint_expr * decl_notation list) list -> unit
+ locality -> polymorphic -> (cofixpoint_expr * decl_notation list) list -> unit
(** Utils *)
-val check_mutuality : Environ.env -> bool -> (identifier * types) list -> unit
+val check_mutuality : Environ.env -> bool -> (Id.t * types) list -> unit
-val declare_fix : definition_object_kind -> identifier ->
- constr -> types -> Impargs.manual_implicits -> global_reference
+val declare_fix : definition_kind -> Univ.universe_context -> Id.t ->
+ Entries.proof_output -> types -> Impargs.manual_implicits -> global_reference
diff --git a/toplevel/coqinit.ml b/toplevel/coqinit.ml
index 0b83bbb8..03074ced 100644
--- a/toplevel/coqinit.ml
+++ b/toplevel/coqinit.ml
@@ -1,18 +1,19 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Util
open Pp
-open System
-open Toplevel
-let (/) s1 s2 = s1 ^ "/" ^ s2
+let ( / ) s1 s2 = s1 ^ "/" ^ s2
-let set_debug () = Flags.debug := true
+let set_debug () =
+ let () = Backtrace.record_backtrace true in
+ Flags.debug := true
(* Loading of the ressource file.
rcfile is either $XDG_CONFIG_HOME/.coqrc.VERSION, or $XDG_CONFIG_HOME/.coqrc if the first one
@@ -30,16 +31,19 @@ let load_rcfile() =
if !load_rc then
try
if !rcfile_specified then
- if file_readable_p !rcfile then
+ if CUnix.file_readable_p !rcfile then
Vernac.load_vernac false !rcfile
else raise (Sys_error ("Cannot read rcfile: "^ !rcfile))
- else try let inferedrc = List.find file_readable_p [
- Envars.xdg_config_home/rcdefaultname^"."^Coq_config.version;
- Envars.xdg_config_home/rcdefaultname;
- System.home/"."^rcdefaultname^"."^Coq_config.version;
- System.home/"."^rcdefaultname;
- ] in
- Vernac.load_vernac false inferedrc
+ else
+ try
+ let warn x = 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;
+ Envars.home ~warn / "."^rcdefaultname^"."^Coq_config.version;
+ Envars.home ~warn / "."^rcdefaultname
+ ] in
+ Vernac.load_vernac false inferedrc
with Not_found -> ()
(*
Flags.if_verbose
@@ -47,79 +51,74 @@ let load_rcfile() =
" found. Skipping rcfile loading."))
*)
with reraise ->
- (msgnl (str"Load of rcfile failed.");
- raise reraise)
+ let reraise = Errors.push reraise in
+ let () = msg_info (str"Load of rcfile failed.") in
+ iraise reraise
else
- Flags.if_verbose msgnl (str"Skipping rcfile loading.")
+ Flags.if_verbose msg_info (str"Skipping rcfile loading.")
(* Puts dir in the path of ML and in the LoadPath *)
let coq_add_path unix_path s =
- Mltop.add_path ~unix_path ~coq_root:(Names.make_dirpath [Nameops.coq_root;Names.id_of_string s])
-let coq_add_rec_path unix_path = Mltop.add_rec_path ~unix_path ~coq_root:(Names.make_dirpath [Nameops.coq_root])
+ Mltop.add_path ~unix_path ~coq_root:(Names.DirPath.make [Nameops.coq_root;Names.Id.of_string s]) ~implicit:true;
+ Mltop.add_ml_dir unix_path
-(* By the option -include -I or -R of the command line *)
-let includes = ref []
-let push_include (s, alias) = includes := (s,alias,false) :: !includes
-let push_rec_include (s, alias) = includes := (s,alias,true) :: !includes
+(* Recursively puts dir in the LoadPath if -nois was not passed *)
+let add_stdlib_path ~unix_path ~coq_root ~with_ml =
+ if !Flags.load_init then
+ Mltop.add_rec_path ~unix_path ~coq_root ~implicit:true
+ else
+ Mltop.add_path ~unix_path ~coq_root ~implicit:false;
+ if with_ml then
+ Mltop.add_rec_ml_dir unix_path
-(* The list of all theories in the standard library /!\ order does matter *)
-let theories_dirs_map = [
- "theories/Unicode", "Unicode" ;
- "theories/Classes", "Classes" ;
- "theories/Program", "Program" ;
- "theories/MSets", "MSets" ;
- "theories/FSets", "FSets" ;
- "theories/Reals", "Reals" ;
- "theories/Strings", "Strings" ;
- "theories/Sorting", "Sorting" ;
- "theories/Setoids", "Setoids" ;
- "theories/Sets", "Sets" ;
- "theories/Structures", "Structures" ;
- "theories/Lists", "Lists" ;
- "theories/Vectors", "Vectors" ;
- "theories/Wellfounded", "Wellfounded" ;
- "theories/Relations", "Relations" ;
- "theories/Numbers", "Numbers" ;
- "theories/QArith", "QArith" ;
- "theories/PArith", "PArith" ;
- "theories/NArith", "NArith" ;
- "theories/ZArith", "ZArith" ;
- "theories/Arith", "Arith" ;
- "theories/Bool", "Bool" ;
- "theories/Logic", "Logic" ;
- "theories/Init", "Init"
- ]
+let add_userlib_path ~unix_path =
+ Mltop.add_path ~unix_path ~coq_root:Nameops.default_root_prefix ~implicit:false;
+ Mltop.add_rec_ml_dir unix_path
+
+(* Options -I, -I-as, and -R of the command line *)
+let includes = ref []
+let push_include s alias recursive implicit =
+ includes := (s,alias,recursive,implicit) :: !includes
+let ml_includes = ref []
+let push_ml_include s = ml_includes := s :: !ml_includes
(* Initializes the LoadPath *)
let init_load_path () =
let coqlib = Envars.coqlib () in
let user_contrib = coqlib/"user-contrib" in
- let xdg_dirs = Envars.xdg_dirs in
+ let xdg_dirs = Envars.xdg_dirs ~warn:(fun x -> msg_warning (str x)) in
let coqpath = Envars.coqpath in
- let dirs = ["states";"plugins"] in
+ let coq_root = Names.DirPath.make [Nameops.coq_root] in
(* NOTE: These directories are searched from last to first *)
(* first, developer specific directory to open *)
if Coq_config.local then coq_add_path (coqlib/"dev") "dev";
+ (* main loops *)
+ if Coq_config.local || !Flags.boot then begin
+ let () = Mltop.add_ml_dir (coqlib/"stm") in
+ Mltop.add_ml_dir (coqlib/"ide")
+ end;
+ Mltop.add_ml_dir (coqlib/"toploop");
(* then standard library *)
- List.iter
- (fun (s,alias) -> Mltop.add_rec_path ~unix_path:(coqlib/s) ~coq_root:(Names.make_dirpath [Names.id_of_string alias; Nameops.coq_root]))
- theories_dirs_map;
- (* then states and plugins *)
- List.iter (fun s -> coq_add_rec_path (coqlib/s)) dirs;
+ add_stdlib_path ~unix_path:(coqlib/"theories") ~coq_root ~with_ml:false;
+ (* then plugins *)
+ add_stdlib_path ~unix_path:(coqlib/"plugins") ~coq_root ~with_ml:true;
(* then user-contrib *)
if Sys.file_exists user_contrib then
- Mltop.add_rec_path ~unix_path:user_contrib ~coq_root:Nameops.default_root_prefix;
+ add_userlib_path ~unix_path:user_contrib;
(* then directories in XDG_DATA_DIRS and XDG_DATA_HOME *)
- List.iter (fun s -> Mltop.add_rec_path ~unix_path:s ~coq_root:Nameops.default_root_prefix) xdg_dirs;
+ List.iter (fun s -> add_userlib_path ~unix_path:s) xdg_dirs;
(* then directories in COQPATH *)
- List.iter (fun s -> Mltop.add_rec_path ~unix_path:s ~coq_root:Nameops.default_root_prefix) coqpath;
+ List.iter (fun s -> add_userlib_path ~unix_path:s) coqpath;
(* then current directory *)
- Mltop.add_path ~unix_path:"." ~coq_root:Nameops.default_root_prefix;
- (* additional loadpath, given with -I -include -R options *)
+ Mltop.add_path ~unix_path:"." ~coq_root:Nameops.default_root_prefix ~implicit:false;
+ (* additional loadpath, given with options -I-as, -Q, and -R *)
List.iter
- (fun (unix_path, coq_root, reci) ->
- if reci then Mltop.add_rec_path ~unix_path ~coq_root else Mltop.add_path ~unix_path ~coq_root)
- (List.rev !includes)
+ (fun (unix_path, coq_root, reci, implicit) ->
+ (if reci then Mltop.add_rec_path else Mltop.add_path)
+ ~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)
let init_library_roots () =
includes := []
@@ -134,12 +133,14 @@ let init_ocaml_path () =
List.iter add_subdir
[ [ "config" ]; [ "dev" ]; [ "lib" ]; [ "kernel" ]; [ "library" ];
[ "pretyping" ]; [ "interp" ]; [ "parsing" ]; [ "proofs" ];
- [ "tactics" ]; [ "toplevel" ]; [ "translate" ]; [ "ide" ] ]
+ [ "tactics" ]; [ "toplevel" ]; [ "printing" ]; [ "intf" ];
+ [ "grammar" ]; [ "ide" ] ]
let get_compat_version = function
+ | "8.4" -> Flags.V8_4
| "8.3" -> Flags.V8_3
| "8.2" -> Flags.V8_2
| ("8.1" | "8.0") as s ->
- msg_warn ("Compatibility with version "^s^" not supported.");
+ msg_warning (strbrk ("Compatibility with version "^s^" not supported."));
Flags.V8_2
- | s -> Util.error ("Unknown compatibility version \""^s^"\".")
+ | s -> Errors.error ("Unknown compatibility version \""^s^"\".")
diff --git a/toplevel/coqinit.mli b/toplevel/coqinit.mli
index d9484443..5f7133c3 100644
--- a/toplevel/coqinit.mli
+++ b/toplevel/coqinit.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -15,8 +15,10 @@ val set_rcfile : string -> unit
val no_load_rc : unit -> unit
val load_rcfile : unit -> unit
-val push_include : string * Names.dir_path -> unit
-val push_rec_include : string * Names.dir_path -> unit
+val push_include : string -> Names.DirPath.t -> bool -> bool -> unit
+(** [push_include phys_path log_path recursive implicit] *)
+
+val push_ml_include : string -> unit
val init_load_path : unit -> unit
val init_library_roots : unit -> unit
diff --git a/toplevel/toplevel.ml b/toplevel/coqloop.ml
index d38b1f78..52fa9e01 100644
--- a/toplevel/toplevel.ml
+++ b/toplevel/coqloop.ml
@@ -1,19 +1,17 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Pp
+open Errors
open Util
open Flags
-open Cerrors
open Vernac
-open Vernacexpr
open Pcoq
-open Compat
(* A buffer for the character read from a channel. We store the command
* entered to be able to report errors without pretty-printing. *)
@@ -22,7 +20,7 @@ type input_buffer = {
mutable prompt : unit -> string;
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 bols : int list; (* offsets in str of beginning of lines *)
mutable tokens : Gram.parsable; (* stream of tokens *)
mutable start : int } (* stream count of the first char of the buffer *)
@@ -33,7 +31,7 @@ let resize_buffer ibuf =
String.blit ibuf.str 0 nstr 0 (String.length ibuf.str);
ibuf.str <- nstr
-(* Delete all irrelevent lines of the input buffer. Keep the last line
+(* 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. *)
let resynch_buffer ibuf =
@@ -58,8 +56,8 @@ let emacs_prompt_endstring() = Printer.emacs_str "</prompt>"
beginning of line. *)
let prompt_char ic ibuf count =
let bol = match ibuf.bols with
- | ll::_ -> ibuf.len == ll
- | [] -> ibuf.len == 0
+ | ll::_ -> Int.equal ibuf.len ll
+ | [] -> Int.equal ibuf.len 0
in
if bol && not !print_emacs then msgerr (str (ibuf.prompt()));
try
@@ -88,7 +86,7 @@ let reset_input_buffer ic ibuf =
let get_bols_of_loc ibuf (bp,ep) =
let add_line (b,e) lines =
- if b < 0 or e < b then anomaly "Bad location";
+ if b < 0 || e < b then anomaly (Pp.str "Bad location");
match lines with
| ([],None) -> ([], Some (b,e))
| (fl,oe) -> ((b,e)::fl, oe)
@@ -118,13 +116,13 @@ let blanch_utf8_string s bp ep =
fixed-size char and therefore contract all utf-8 code into one
space; in any case, preserve tabulation so
that its effective interpretation in terms of spacing is preserved *)
- if s.[i] = '\t' then s'.[!j] <- '\t';
+ if s.[i] == '\t' then s'.[!j] <- '\t';
if n < 0x80 || 0xC0 <= n then incr j
done;
String.sub s' 0 !j
let print_highlight_location ib loc =
- let (bp,ep) = unloc loc in
+ let (bp,ep) = Loc.unloc loc in
let bp = bp - ib.start
and ep = ep - ib.start in
let highlight_lines =
@@ -146,68 +144,55 @@ let print_highlight_location ib loc =
str sn ++ str dn) in
(l1 ++ li ++ ln)
in
- let loc = make_loc (bp,ep) in
+ let loc = Loc.make_loc (bp,ep) in
(str"Toplevel input, characters " ++ Cerrors.print_loc loc ++ str":" ++ fnl () ++
highlight_lines ++ fnl ())
(* Functions to report located errors in a file. *)
-let print_location_in_file s inlibrary fname loc =
+let print_location_in_file {outer=s;inner=fname} loc =
let errstrm = str"Error while reading " ++ str s in
- if loc = dummy_loc then
+ if Loc.is_ghost loc then
hov 1 (errstrm ++ spc() ++ str" (unknown location):") ++ fnl ()
else
let errstrm =
- if s = fname then mt() else errstrm ++ str":" ++ fnl() in
- if inlibrary then
- hov 0 (errstrm ++ str"Module " ++ str ("\""^fname^"\"") ++ spc() ++
- str"characters " ++ Cerrors.print_loc loc) ++ fnl ()
- else
- let (bp,ep) = unloc loc in
- 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
+ if String.equal s fname then mt() else errstrm ++ str":" ++ fnl()
+ in
+ let (bp,ep) = Loc.unloc loc in
+ let line_of_pos lin bol cnt =
try
- let (line, bol) = line_of_pos 1 0 0 in
+ 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;
- hov 0 (* No line break so as to follow emacs error message format *)
- (errstrm ++ str"File " ++ str ("\""^fname^"\"") ++
+ 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 ("\""^fname^"\"") ++
str", line " ++ int line ++ str", characters " ++
- Cerrors.print_loc (make_loc (bp-bol,ep-bol))) ++ str":" ++
+ Cerrors.print_loc (Loc.make_loc (bp-bol,ep-bol))) ++ str":" ++
fnl ()
- with e when Errors.noncritical e ->
- (close_in ic;
- hov 1 (errstrm ++ spc() ++ str"(invalid location):") ++ fnl ())
-
-let print_command_location ib dloc =
- match dloc with
- | Some (bp,ep) ->
- (str"Error during interpretation of command:" ++ fnl () ++
- str(String.sub ib.str (bp-ib.start) (ep-bp)) ++ fnl ())
- | None -> (mt ())
-
-let valid_loc dloc loc =
- loc <> dummy_loc
- & match dloc with
- | Some dloc ->
- let (bd,ed) = unloc dloc in let (b,e) = unloc loc in bd<=b & e<=ed
- | _ -> true
-
-let valid_buffer_loc ib dloc loc =
- valid_loc dloc loc &
- let (b,e) = unloc loc in b-ib.start >= 0 & e-ib.start < ib.len & b<=e
+ with e when Errors.noncritical e ->
+ hov 1 (errstrm ++ spc() ++ str"(invalid location):") ++ fnl ()
+
+let valid_buffer_loc ib loc =
+ not (Loc.is_ghost loc) &&
+ let (b,e) = Loc.unloc loc in b-ib.start >= 0 && e-ib.start < ib.len && b<=e
(*s The Coq prompt is the name of the focused proof, if any, and "Coq"
otherwise. We trap all exceptions to prevent the error message printing
from cycling. *)
let make_prompt () =
try
- (Names.string_of_id (Pfedit.get_current_proof_name ())) ^ " < "
+ (Names.Id.to_string (Pfedit.get_current_proof_name ())) ^ " < "
with Proof_global.NoCurrentProof ->
"Coq < "
@@ -216,7 +201,7 @@ let make_prompt () =
let l' = ref l in
let res =
while List.length !l' > 1 do
- pl := !pl ^ "|" Names.string_of_id x;
+ pl := !pl ^ "|" Names.Id.to_string x;
l':=List.tl !l'
done in
let last = try List.hd !l' with _ -> in
@@ -232,12 +217,12 @@ let make_prompt () =
"n |lem1|lem2|lem3| p < "
*)
let make_emacs_prompt() =
- let statnum = string_of_int (Lib.current_command_label ()) in
- let dpth = Pfedit.current_proof_depth() in
- let pending = Pfedit.get_all_proof_names() in
+ let statnum = Stateid.to_string (Stm.get_current_state ()) in
+ let dpth = Stm.current_proof_depth() in
+ let pending = Stm.get_all_proof_names() in
let pendingprompt =
List.fold_left
- (fun acc x -> acc ^ (if acc <> "" then "|" else "") ^ Names.string_of_id x)
+ (fun acc x -> acc ^ (if String.is_empty acc then "" else "|") ^ Names.Id.to_string x)
"" pending in
let proof_info = if dpth >= 0 then string_of_int dpth else "0" in
if !Flags.print_emacs then statnum ^ " |" ^ pendingprompt ^ "| " ^ proof_info ^ " < "
@@ -267,113 +252,109 @@ let set_prompt prompt =
^ prompt ()
^ emacs_prompt_endstring())
-(* Removes and prints the location of the error. The following exceptions
- need not be located. *)
-let rec is_pervasive_exn = function
- | Out_of_memory | Stack_overflow | Sys.Break -> true
- | Error_in_file (_,_,e) -> is_pervasive_exn e
- | Loc.Exc_located (_,e) -> is_pervasive_exn e
- | DuringCommandInterp (_,e) -> is_pervasive_exn e
- | _ -> false
-
-(* Toplevel error explanation, dealing with locations, Drop, Ctrl-D
- May raise only the following exceptions: Drop and End_of_input,
- meaning we get out of the Coq loop *)
-let print_toplevel_error exc =
- let (dloc,exc) =
- match exc with
- | DuringCommandInterp (loc,ie) ->
- if loc = dummy_loc then (None,ie) else (Some loc, ie)
- | _ -> (None, exc)
- in
- let (locstrm,exc) =
- match exc with
- | Loc.Exc_located (loc, ie) ->
- if valid_buffer_loc top_buffer dloc loc then
- (print_highlight_location top_buffer loc, ie)
- else
- ((mt ()) (* print_command_location top_buffer dloc *), ie)
- | Error_in_file (s, (inlibrary, fname, loc), ie) ->
- (print_location_in_file s inlibrary fname loc, ie)
- | _ ->
- ((mt ()) (* print_command_location top_buffer dloc *), exc)
+(* The following exceptions need not be located. *)
+
+let locate_exn = function
+ | Out_of_memory | Stack_overflow | Sys.Break -> false
+ | _ -> true
+
+(* Toplevel error explanation. *)
+
+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 ->
+ if locate_exn e && valid_buffer_loc top_buffer loc then
+ print_highlight_location top_buffer loc
+ else mt ()
in
- match exc with
- | End_of_input ->
- msgerrnl (mt ()); pp_flush(); exit 0
- | Vernacexpr.Drop -> (* Last chance *)
- if Mltop.is_ocaml_top() then raise Vernacexpr.Drop;
- (str"Error: There is no ML toplevel." ++ fnl ())
- | Vernacexpr.Quit ->
- raise Vernacexpr.Quit
- | _ ->
- (if is_pervasive_exn exc then (mt ()) else locstrm) ++
- Errors.print exc
+ locmsg ++ Errors.iprint (e, info)
(* Read the input stream until a dot is encountered *)
let parse_to_dot =
- let rec dot st = match get_tok (Stream.next st) with
- | Tok.KEYWORD "." -> ()
+ let rec dot st = match Compat.get_tok (Stream.next st) with
+ | Tok.KEYWORD ("."|"...") -> ()
| Tok.EOI -> raise End_of_input
| _ -> dot st
in
Gram.Entry.of_parser "Coqtoplevel.dot" dot
-(* We assume that when a lexer error occurs, at least one char was eaten *)
+(* If an error occurred while parsing, we try to read the input until a dot
+ token is encountered.
+ We assume that when a lexer error occurs, at least one char was eaten *)
+
let rec discard_to_dot () =
try
Gram.entry_parse parse_to_dot top_buffer.tokens
- with Loc.Exc_located(_,(Token.Error _|Lexer.Error.E _)) ->
- discard_to_dot()
-
-
-(* If the error occured while parsing, we read the input until a dot token
- * in encountered. *)
-
-let process_error = function
- | DuringCommandInterp _ as e -> e
- | e ->
- if is_pervasive_exn e then
- e
- else
- try
- discard_to_dot (); e
- with
- | End_of_input -> End_of_input
- | any -> if is_pervasive_exn any then any else e
-
-(* do_vernac reads and executes a toplevel phrase, and print error
- messages when an exception is raised, except for the following:
- Drop: kill the Coq toplevel, going down to the Caml toplevel if it exists.
+ with
+ | Compat.Token.Error _ | Lexer.Error.E _ -> discard_to_dot ()
+ | End_of_input -> raise End_of_input
+ | e when Errors.noncritical e -> ()
+
+let read_sentence () =
+ try Vernac.parse_sentence (top_buffer.tokens, None)
+ with reraise ->
+ let reraise = Errors.push reraise in
+ discard_to_dot ();
+ iraise reraise
+
+(** [do_vernac] reads and executes a toplevel phrase, and print error
+ messages when an exception is raised, except for the following:
+ - Drop: kill the Coq toplevel, going down to the Caml toplevel if it exists.
Otherwise, exit.
- End_of_input: Ctrl-D was typed in, we will quit *)
+ - End_of_input: Ctrl-D was typed in, we will quit.
+
+ In particular, this is normally the only place where a Sys.Break
+ is caught and handled (i.e. not re-raised).
+*)
+
let do_vernac () =
msgerrnl (mt ());
if !print_emacs then msgerr (str (top_buffer.prompt()));
resynch_buffer top_buffer;
- begin
- try
- raw_do_vernac top_buffer.tokens
- with any ->
- msgnl (print_toplevel_error (process_error any))
- end;
- flush_all()
-
-(* coq and go read vernacular expressions until Drop is entered.
- * Ctrl-C will raise the exception Break instead of aborting Coq.
- * Here we catch the exceptions terminating the Coq loop, and decide
- * if we really must quit.
- *)
+ try
+ Vernac.eval_expr (read_sentence ())
+ 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 ())
+ | any ->
+ let any = Errors.push any in
+ Format.set_formatter_out_channel stdout;
+ let msg = print_toplevel_error any ++ fnl () in
+ pp_with ~pp_tag:Ppstyle.pp_tag !Pp_control.std_ft msg;
+ pp_flush ()
+
+(** Main coq loop : read vernacular expressions until Drop is entered.
+ Ctrl-C is handled internally as Sys.Break instead of aborting Coq.
+ Normally, the only exceptions that can come out of [do_vernac] and
+ exit the loop are Drop and Quit. Any other exception there indicates
+ an issue with [print_toplevel_error] above. *)
+
+(*
+let feed_emacs = function
+ | { Interface.id = Interface.State id;
+ Interface.content = Interface.GlobRef (_,a,_,c,_) } ->
+ prerr_endline ("<info>" ^"<id>"^Stateid.to_string id ^"</id>"
+ ^a^" "^c^ "</info>")
+ | _ -> ()
+*)
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() done
+ while true do do_vernac(); flush_all() done
with
- | Vernacexpr.Drop -> ()
- | End_of_input -> msgerrnl (mt ()); pp_flush(); exit 0
- | Vernacexpr.Quit -> exit 0
+ | Errors.Drop -> ()
+ | Errors.Quit -> exit 0
| any ->
- msgerrnl (str"Anomaly. Please report.");
- loop ()
+ msgerrnl (str"Anomaly: main loop exited with exception: " ++
+ str (Printexc.to_string any) ++
+ fnl() ++ str"Please report.");
+ loop ()
diff --git a/toplevel/toplevel.mli b/toplevel/coqloop.mli
index 3be3115f..8ed661e6 100644
--- a/toplevel/toplevel.mli
+++ b/toplevel/coqloop.mli
@@ -1,13 +1,12 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Pp
-open Pcoq
(** The Coq toplevel loop. *)
@@ -31,9 +30,9 @@ val set_prompt : (unit -> string) -> unit
May raise only the following exceptions: [Drop] and [End_of_input],
meaning we get out of the Coq loop. *)
-val print_toplevel_error : exn -> std_ppcmds
+val print_toplevel_error : Exninfo.iexn -> std_ppcmds
-(** Parse and execute a vernac command. *)
+(** Parse and execute one vernac command. *)
val do_vernac : unit -> unit
diff --git a/toplevel/coqtop.ml b/toplevel/coqtop.ml
index d2eb7d9c..142f3386 100644
--- a/toplevel/coqtop.ml
+++ b/toplevel/coqtop.ml
@@ -1,36 +1,96 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Pp
+open Errors
open Util
open System
open Flags
open Names
open Libnames
-open Nameops
open States
-open Toplevel
open Coqinit
+let () = at_exit flush_all
+
+let ( / ) = Filename.concat
+
+let fatal_error info anomaly =
+ let msg = info ++ fnl () in
+ pp_with ~pp_tag:Ppstyle.pp_tag !Pp_control.err_ft msg;
+ flush_all ();
+ exit (if anomaly then 129 else 1)
+
let get_version_date () =
try
- let coqlib = Envars.coqlib () in
- let ch = open_in (Filename.concat coqlib "revision") in
+ let ch = open_in (Envars.coqlib () / "revision") in
let ver = input_line ch in
let rev = input_line ch in
- (ver,rev)
+ let () = close_in ch in
+ (ver,rev)
with e when Errors.noncritical e ->
(Coq_config.version,Coq_config.date)
let print_header () =
- let (ver,rev) = (get_version_date ()) in
- Printf.printf "Welcome to Coq %s (%s)\n" ver rev;
- flush stdout
+ let (ver,rev) = get_version_date () in
+ ppnl (str ("Welcome to Coq "^ver^" ("^rev^")"));
+ pp_flush ()
+
+let warning s = msg_warning (strbrk s)
+
+let toploop = ref None
+
+let color : [`ON | `AUTO | `OFF] ref = ref `AUTO
+let set_color = function
+| "on" -> color := `ON
+| "off" -> color := `OFF
+| "auto" -> color := `AUTO
+| _ -> prerr_endline ("Error: on/off/auto expected after option color"); exit 1
+
+let init_color () =
+ let has_color = match !color with
+ | `OFF -> false
+ | `ON -> true
+ | `AUTO ->
+ Terminal.has_style Unix.stdout &&
+ Terminal.has_style Unix.stderr
+ 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 ()
+ | Some "" ->
+ (** No color output *)
+ ()
+ | Some s ->
+ (** Overwrite all colors *)
+ Ppstyle.clear_styles ();
+ Ppstyle.parse_config s;
+ Ppstyle.init_color_output ()
+ end
+
+let toploop_init = ref begin fun x ->
+ let () = init_color () in
+ let () = CoqworkmgrApi.(init !Flags.async_proofs_worker_priority) in
+ x
+ end
+
+let toploop_run = ref (fun () ->
+ if Dumpglob.dump () then begin
+ if_verbose warning "Dumpglob cannot be used in interactive mode.";
+ Dumpglob.noglob ()
+ end;
+ Coqloop.loop();
+ (* Initialise and launch the Ocaml toplevel *)
+ Coqinit.init_ocaml_path();
+ Mltop.ocaml_toploop())
let output_context = ref false
@@ -38,7 +98,8 @@ let memory_stat = ref false
let print_memory_stat () =
if !memory_stat then
- Format.printf "total heap size = %d kbytes\n" (heap_size_kb ())
+ ppnl
+ (str "total heap size = " ++ int (CObj.heap_size_kb ()) ++ str " kbytes")
let _ = at_exit print_memory_stat
@@ -47,40 +108,40 @@ let set_engagement c = engagement := Some c
let engage () =
match !engagement with Some c -> Global.set_engagement c | None -> ()
+let type_in_type = ref false
+let set_type_in_type () = type_in_type := true
+let set_hierarchy () = if !type_in_type then Global.set_type_in_type ()
+
let set_batch_mode () = batch_mode := true
-let toplevel_default_name = make_dirpath [id_of_string "Top"]
+let toplevel_default_name = DirPath.make [Id.of_string "Top"]
let toplevel_name = ref (Some toplevel_default_name)
let set_toplevel_name dir =
- if dir = empty_dirpath then error "Need a non empty toplevel module name";
+ if DirPath.is_empty dir then error "Need a non empty toplevel module name";
toplevel_name := Some dir
let unset_toplevel_name () = toplevel_name := None
let remove_top_ml () = Mltop.remove ()
-let inputstate = ref None
-let set_inputstate s = inputstate:= Some s
-let inputstate () =
- match !inputstate with
- | Some "" -> ()
- | Some s -> intern_state s
- | None -> intern_state "initial.coq"
+let inputstate = ref ""
+let set_inputstate s =
+ let () = msg_warning (str "The inputstate option is deprecated and discouraged.") in
+ inputstate:=s
+let inputstate () = if not (String.is_empty !inputstate) then intern_state !inputstate
let outputstate = ref ""
-let set_outputstate s = outputstate:=s
-let outputstate () = if !outputstate <> "" then extern_state !outputstate
+let set_outputstate s =
+ let () = msg_warning (str "The outputstate option is deprecated and discouraged.") in
+ outputstate:=s
+let outputstate () = if not (String.is_empty !outputstate) then extern_state !outputstate
-let set_default_include d = push_include (d,Nameops.default_root_prefix)
-let set_include d p =
+let set_include d p recursive implicit =
let p = dirpath_of_string p in
- push_include (d,p)
-let set_rec_include d p =
- let p = dirpath_of_string p in
- push_rec_include(d,p)
+ push_include d p recursive implicit
let load_vernacular_list = ref ([] : (string * bool) list)
let add_load_vernacular verb s =
- load_vernacular_list := ((make_suffix s ".v"),verb) :: !load_vernacular_list
+ load_vernacular_list := ((CUnix.make_suffix s ".v"),verb) :: !load_vernacular_list
let load_vernacular () =
List.iter
(fun (s,b) ->
@@ -96,29 +157,57 @@ let load_vernac_obj () =
List.iter (fun f -> Library.require_library_from_file None f None)
(List.rev !load_vernacular_obj)
+let require_prelude () =
+ let vo = Envars.coqlib () / "theories/Init/Prelude.vo" in
+ let vio = Envars.coqlib () / "theories/Init/Prelude.vio" in
+ let m =
+ if Sys.file_exists vo then vo else
+ if Sys.file_exists vio then vio else vo in
+ Library.require_library_from_dirpath [Coqlib.prelude_module,m] (Some true)
+
let require_list = ref ([] : string list)
let add_require s = require_list := s :: !require_list
let require () =
+ if !load_init then silently require_prelude ();
List.iter (fun s -> Library.require_library_from_file None s (Some false))
(List.rev !require_list)
let compile_list = ref ([] : (bool * string) list)
+
+let glob_opt = ref false
+
let add_compile verbose s =
set_batch_mode ();
Flags.make_silent true;
+ if not !glob_opt then Dumpglob.dump_to_dotglob ();
+ (** make the file name explicit; needed not to break up Coq loadpath stuff. *)
+ let s =
+ let open Filename in
+ if is_implicit s
+ then concat current_dir_name s
+ else s
+ in
compile_list := (verbose,s) :: !compile_list
+
+let compile_file (v,f) =
+ if Flags.do_beautify () then
+ with_option beautify_file (Vernac.compile v) f
+ else
+ Vernac.compile v f
+
let compile_files () =
- let init_state = States.freeze() in
- let coqdoc_init_state = Dumpglob.coqdoc_freeze () in
- List.iter
- (fun (v,f) ->
- States.unfreeze init_state;
- Dumpglob.coqdoc_unfreeze coqdoc_init_state;
- if Flags.do_beautify () then
- with_option beautify_file (Vernac.compile v) f
- else
- Vernac.compile v f)
- (List.rev !compile_list)
+ 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)
(*s options for the virtual machine *)
@@ -129,248 +218,402 @@ let set_vm_opt () =
Vm.set_transp_values (not !boxed_val);
Vconv.set_use_vm !use_vm
+(** Options for proof general *)
+
+let set_emacs () =
+ Flags.print_emacs := true;
+ Pp.make_pp_emacs ();
+ Vernacentries.qed_display_script := false;
+ color := `OFF
+
+(** GC tweaking *)
+
+(** Coq is a heavy user of persistent data structures and symbolic ASTs, so the
+ minor heap is heavily sollicited. Unfortunately, the default size is far too
+ small, so we enlarge it a lot (128 times larger).
+
+ To better handle huge memory consumers, we also augment the default major
+ heap increment and the GC pressure coefficient.
+*)
+
+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
+
(*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 () =
- if !batch_mode then
- Usage.print_usage_coqc ()
- else
- Usage.print_usage_coqtop () ;
- flush stderr ;
- exit 1
+ Envars.set_coqlib Errors.error;
+ init_load_path ();
+ if !batch_mode then Usage.print_usage_coqc ()
+ else begin
+ Mltop.load_ml_objects_raw_rex
+ (Str.regexp (if Mltop.is_native then "^.*top.cmxs$" else "^.*top.cma$"));
+ Usage.print_usage_coqtop ()
+ end
+
+let print_style_tags () =
+ let () = init_color () in
+ let tags = Ppstyle.dump () in
+ let iter (t, st) =
+ let st = match st with Some st -> st | None -> Terminal.make () in
+ let opt =
+ Terminal.eval st ^
+ String.concat "." (Ppstyle.repr t) ^
+ Terminal.reset ^ "\n"
+ in
+ print_string opt
+ in
+ List.iter iter tags;
+ flush_all ()
-let warning s = msg_warning (str s)
+let error_missing_arg s =
+ prerr_endline ("Error: extra argument expected after option "^s);
+ prerr_endline "See -help for the syntax of supported options";
+ exit 1
-let ide_slave = ref false
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 parse_args arglist =
- let glob_opt = ref false in
- let rec parse = function
- | [] -> []
- | "-with-geoproof" :: s :: rem ->
- if s = "yes" then Coq_config.with_geoproof := true
- else if s = "no" then Coq_config.with_geoproof := false
- else usage ();
- parse rem
- | "-impredicative-set" :: rem ->
- set_engagement Declarations.ImpredicativeSet; parse rem
-
- | ("-I"|"-include") :: d :: "-as" :: p :: rem -> set_include d p; parse rem
- | ("-I"|"-include") :: d :: "-as" :: [] -> usage ()
- | ("-I"|"-include") :: d :: rem -> set_default_include d; parse rem
- | ("-I"|"-include") :: [] -> usage ()
-
- | "-R" :: d :: "-as" :: p :: rem -> set_rec_include d p;parse rem
- | "-R" :: d :: "-as" :: [] -> usage ()
- | "-R" :: d :: p :: rem -> set_rec_include d p;parse rem
- | "-R" :: ([] | [_]) -> usage ()
-
- | "-top" :: d :: rem -> set_toplevel_name (dirpath_of_string d); parse rem
- | "-top" :: [] -> usage ()
-
- | "-exclude-dir" :: f :: rem -> exclude_search_in_dirname f; parse rem
- | "-exclude-dir" :: [] -> usage ()
-
- | "-notop" :: rem -> unset_toplevel_name (); parse rem
- | "-q" :: rem -> no_load_rc (); parse rem
-
- | "-opt" :: rem -> warning "option -opt deprecated, call with .opt suffix\n"; parse rem
- | "-byte" :: rem -> warning "option -byte deprecated, call with .byte suffix\n"; parse rem
- | "-full" :: rem -> warning "option -full deprecated\n"; parse rem
-
- | "-batch" :: rem -> set_batch_mode (); parse rem
- | "-boot" :: rem -> boot := true; no_load_rc (); parse rem
- | "-quality" :: rem -> term_quality := true; no_load_rc (); parse rem
- | "-outputstate" :: s :: rem ->
- Flags.load_proofs := Flags.Force; set_outputstate s; parse rem
- | "-outputstate" :: [] -> usage ()
-
- | "-nois" :: rem -> set_inputstate ""; parse rem
-
- | ("-inputstate"|"-is") :: s :: rem -> set_inputstate s; parse rem
- | ("-inputstate"|"-is") :: [] -> usage ()
+let print_where = ref false
+let print_config = ref false
+let print_tags = ref false
- | "-load-ml-object" :: f :: rem -> Mltop.dir_ml_load f; parse rem
- | "-load-ml-object" :: [] -> usage ()
+let get_priority opt s =
+ try Flags.priority_of_string s
+ with Invalid_argument _ ->
+ prerr_endline ("Error: low/high expected after "^opt); exit 1
- | "-load-ml-source" :: f :: rem -> Mltop.dir_ml_use f; parse rem
- | "-load-ml-source" :: [] -> usage ()
+let get_async_proofs_mode opt = function
+ | "off" -> Flags.APoff
+ | "on" -> Flags.APon
+ | "lazy" -> Flags.APonLazy
+ | _ -> prerr_endline ("Error: on/off/lazy expected after "^opt); exit 1
- | ("-load-vernac-source"|"-l") :: f :: rem ->
- add_load_vernacular false f; parse rem
- | ("-load-vernac-source"|"-l") :: [] -> usage ()
+let get_cache opt = function
+ | "force" -> Some Flags.Force
+ | _ -> prerr_endline ("Error: force expected after "^opt); exit 1
- | ("-load-vernac-source-verbose"|"-lv") :: f :: rem ->
- add_load_vernacular true f; parse rem
- | ("-load-vernac-source-verbose"|"-lv") :: [] -> usage ()
- | "-load-vernac-object" :: f :: rem -> add_vernac_obj f; parse rem
- | "-load-vernac-object" :: [] -> usage ()
+let set_worker_id opt s =
+ assert (s <> "master");
+ Flags.async_proofs_worker_id := s
- | "-dump-glob" :: "stdout" :: rem -> Dumpglob.dump_to_stdout (); glob_opt := true; parse rem
- (* À ne pas documenter : l'option 'stdout' n'étant
- éventuellement utile que pour le debugging... *)
- | "-dump-glob" :: f :: rem -> Dumpglob.dump_into_file f; glob_opt := true; parse rem
- | "-dump-glob" :: [] -> usage ()
- | ("-no-glob" | "-noglob") :: rem -> Dumpglob.noglob (); glob_opt := true; parse rem
+let get_bool opt = function
+ | "yes" -> true
+ | "no" -> false
+ | _ -> prerr_endline ("Error: yes/no expected after option "^opt); exit 1
- | "-require" :: f :: rem -> add_require f; parse rem
- | "-require" :: [] -> usage ()
+let get_int opt n =
+ try int_of_string n
+ with Failure _ ->
+ prerr_endline ("Error: integer expected after option "^opt); exit 1
- | "-compile" :: f :: rem -> add_compile false f; if not !glob_opt then Dumpglob.dump_to_dotglob (); parse rem
- | "-compile" :: [] -> usage ()
+let get_host_port opt s =
+ match CString.split ':' s with
+ | [host; port] -> Some (Spawned.Socket(host, int_of_string port))
+ | ["stdfds"] -> Some Spawned.AnonPipe
+ | _ ->
+ prerr_endline ("Error: host:port or stdfds expected after option "^opt);
+ exit 1
- | "-compile-verbose" :: f :: rem -> add_compile true f; if not !glob_opt then Dumpglob.dump_to_dotglob (); parse rem
- | "-compile-verbose" :: [] -> usage ()
+let get_task_list s = List.map int_of_string (Str.split (Str.regexp ",") s)
- | "-force-load-proofs" :: rem -> Flags.load_proofs := Flags.Force; parse rem
- | "-lazy-load-proofs" :: rem -> Flags.load_proofs := Flags.Lazy; parse rem
- | "-dont-load-proofs" :: rem -> Flags.load_proofs := Flags.Dont; parse rem
+let vio_tasks = ref []
- | "-beautify" :: rem -> make_beautify true; parse rem
-
- | "-unsafe" :: f :: rem -> add_unsafe f; parse rem
- | "-unsafe" :: [] -> usage ()
-
- | "-debug" :: rem -> set_debug (); parse rem
-
- | "-compat" :: v :: rem ->
- Flags.compat_version := get_compat_version v; parse rem
- | "-compat" :: [] -> usage ()
-
- | "-verbose-compat-notations" :: rem -> verb_compat_ntn := true; parse rem
- | "-no-compat-notations" :: rem -> no_compat_ntn := true; parse rem
-
- | "-vm" :: rem -> use_vm := true; parse rem
- | "-emacs" :: rem ->
- Flags.print_emacs := true; Pp.make_pp_emacs();
- Vernacentries.qed_display_script := false;
- parse rem
- | "-emacs-U" :: rem ->
- warning "Obsolete option \"-emacs-U\", use -emacs instead.";
- parse ("-emacs" :: rem)
-
- | "-unicode" :: rem -> add_require "Utf8_core"; parse rem
-
- | "-coqlib" :: d :: rem -> Flags.coqlib_spec:=true; Flags.coqlib:=d; parse rem
- | "-coqlib" :: [] -> usage ()
-
- | "-where" :: _ -> print_endline (Envars.coqlib ()); exit (if !filter_opts then 2 else 0)
-
- | ("-config"|"--config") :: _ -> Usage.print_config (); exit (if !filter_opts then 2 else 0)
-
- | ("-quiet"|"-silent") :: rem -> Flags.make_silent true; parse rem
-
- | ("-?"|"-h"|"-H"|"-help"|"--help") :: _ -> usage ()
-
- | ("-v"|"--version") :: _ -> Usage.version (if !filter_opts then 2 else 0)
-
- | "-init-file" :: f :: rem -> set_rcfile f; parse rem
- | "-init-file" :: [] -> usage ()
-
- | "-notactics" :: rem ->
- warning "Obsolete option \"-notactics\".";
- remove_top_ml (); parse rem
-
- | "-just-parsing" :: rem -> Vernac.just_parsing := true; parse rem
-
- | ("-m" | "--memory") :: rem -> memory_stat := true; parse rem
-
- | "-xml" :: rem -> Flags.xml_export := true; parse rem
-
- | "-output-context" :: rem -> output_context := true; parse rem
-
- (* Scanned in Flags. *)
- | "-v7" :: rem -> error "This version of Coq does not support v7 syntax"
- | "-v8" :: rem -> parse rem
-
- | "-no-hash-consing" :: rem -> Flags.hash_cons_proofs := false; parse rem
-
- | "-ideslave" :: rem -> ide_slave := true; parse rem
+let add_vio_task f =
+ set_batch_mode ();
+ Flags.make_silent true;
+ vio_tasks := f :: !vio_tasks
+
+let check_vio_tasks () =
+ let rc =
+ List.fold_left (fun acc t -> Vio_checking.check_vio t && acc)
+ true (List.rev !vio_tasks) in
+ if not rc then exit 1
+
+let vio_files = ref []
+let vio_files_j = ref 0
+let vio_checking = ref false
+let add_vio_file f =
+ set_batch_mode ();
+ Flags.make_silent true;
+ vio_files := f :: !vio_files
+
+let set_vio_checking_j opt j =
+ try vio_files_j := int_of_string j
+ with Failure _ ->
+ prerr_endline ("The first argument of " ^ opt ^ " must the number");
+ prerr_endline "of concurrent workers to be used (a positive integer).";
+ prerr_endline "Makefiles generated by coq_makefile should be called";
+ prerr_endline "setting the J variable like in 'make vio2vo J=3'";
+ exit 1
+
+let is_not_dash_option = function
+ | Some f when String.length f > 0 && f.[0] <> '-' -> true
+ | _ -> false
+
+let schedule_vio_checking () =
+ if !vio_files <> [] && !vio_checking then
+ Vio_checking.schedule_vio_checking !vio_files_j !vio_files
+let schedule_vio_compilation () =
+ if !vio_files <> [] && not !vio_checking then
+ Vio_checking.schedule_vio_compilation !vio_files_j !vio_files
+
+let get_native_name s =
+ (* We ignore even critical errors because this mode has to be super silent *)
+ try
+ String.concat Filename.dir_sep [Filename.dirname s;
+ Nativelib.output_dir; Library.native_name_from_filename s]
+ with _ -> ""
- | "-filteropts" :: rem -> filter_opts := true; parse rem
+let parse_args arglist =
+ let args = ref arglist in
+ let extras = ref [] in
+ let rec parse () = match !args with
+ | [] -> List.rev !extras
+ | opt :: rem ->
+ args := rem;
+ let next () = match !args with
+ | x::rem -> args := rem; x
+ | [] -> error_missing_arg opt
+ in
+ let peek_next () = match !args with
+ | x::_ -> Some x
+ | [] -> None
+ in
+ begin match opt with
+
+ (* Complex options with many args *)
+ |"-I"|"-include" ->
+ begin match rem with
+ | d :: "-as" :: p :: rem -> set_include d p false true; args := rem
+ | d :: "-as" :: [] -> error_missing_arg "-as"
+ | d :: rem -> push_ml_include d; args := rem
+ | [] -> error_missing_arg opt
+ end
+ |"-Q" ->
+ begin match rem with
+ | d :: p :: rem -> set_include d p true false; args := rem
+ | _ -> error_missing_arg opt
+ end
+ |"-R" ->
+ begin match rem with
+ | d :: "-as" :: [] -> error_missing_arg "-as"
+ | d :: "-as" :: p :: rem
+ | d :: p :: rem -> set_include d p true true; args := rem
+ | _ -> error_missing_arg opt
+ end
- | s :: rem ->
- if !filter_opts then
- s :: (parse rem)
- else
- (prerr_endline ("Don't know what to do with " ^ s); usage ())
+ (* Options with two arg *)
+ |"-check-vio-tasks" ->
+ let tno = get_task_list (next ()) in
+ let tfile = next () in
+ add_vio_task (tno,tfile)
+ |"-schedule-vio-checking" ->
+ vio_checking := true;
+ set_vio_checking_j opt (next ());
+ add_vio_file (next ());
+ while is_not_dash_option (peek_next ()) do add_vio_file (next ()); done
+ |"-schedule-vio2vo" ->
+ set_vio_checking_j opt (next ());
+ add_vio_file (next ());
+ while is_not_dash_option (peek_next ()) do add_vio_file (next ()); done
+
+ (* Options with one arg *)
+ |"-coqlib" -> Flags.coqlib_spec:=true; Flags.coqlib:=(next ())
+ |"-async-proofs" ->
+ Flags.async_proofs_mode := get_async_proofs_mode opt (next())
+ |"-async-proofs-j" ->
+ Flags.async_proofs_n_workers := (get_int opt (next ()))
+ |"-async-proofs-cache" ->
+ Flags.async_proofs_cache := get_cache opt (next ())
+ |"-async-proofs-tac-j" ->
+ Flags.async_proofs_n_tacworkers := (get_int opt (next ()))
+ |"-async-proofs-worker-priority" ->
+ Flags.async_proofs_worker_priority := get_priority opt (next ())
+ |"-async-proofs-private-flags" ->
+ Flags.async_proofs_private_flags := Some (next ());
+ |"-worker-id" -> set_worker_id opt (next ())
+ |"-compat" -> Flags.compat_version := get_compat_version (next ())
+ |"-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 ())
+ |"-init-file" -> set_rcfile (next ())
+ |"-inputstate"|"-is" -> set_inputstate (next ())
+ |"-load-ml-object" -> Mltop.dir_ml_load (next ())
+ |"-load-ml-source" -> Mltop.dir_ml_use (next ())
+ |"-load-vernac-object" -> add_vernac_obj (next ())
+ |"-load-vernac-source"|"-l" -> add_load_vernacular false (next ())
+ |"-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
+ |"-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 ())
+
+ (* Options with zero arg *)
+ |"-async-queries-always-delegate"
+ |"-async-proofs-always-delegate"
+ |"-async-proofs-full" ->
+ Flags.async_proofs_full := true;
+ |"-async-proofs-never-reopen-branch" ->
+ Flags.async_proofs_never_reopen_branch := true;
+ |"-batch" -> set_batch_mode ()
+ |"-beautify" -> make_beautify true
+ |"-boot" -> boot := true; no_load_rc ()
+ |"-bt" -> Backtrace.record_backtrace true
+ |"-color" -> set_color (next ())
+ |"-config"|"--config" -> print_config := true
+ |"-debug" -> set_debug ()
+ |"-emacs" -> set_emacs ()
+ |"-filteropts" -> filter_opts := true
+ |"-h"|"-H"|"-?"|"-help"|"--help" -> usage ()
+ |"-ideslave" -> toploop := Some "coqidetop"; Flags.ide_slave := true
+ |"-impredicative-set" -> set_engagement Declarations.ImpredicativeSet
+ |"-indices-matter" -> Indtypes.enforce_indices_matter ()
+ |"-just-parsing" -> Vernac.just_parsing := true
+ |"-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
+ |"-no-native-compiler" -> no_native_compiler := true
+ |"-notop" -> unset_toplevel_name ()
+ |"-output-context" -> output_context := true
+ |"-q" -> no_load_rc ()
+ |"-quiet"|"-silent" -> Flags.make_silent true
+ |"-quick" -> Flags.compilation_mode := BuildVio
+ |"-list-tags" -> print_tags := true
+ |"-time" -> Flags.time := true
+ |"-type-in-type" -> set_type_in_type ()
+ |"-unicode" -> add_require "Utf8_core"
+ |"-v"|"--version" -> Usage.version (exitcode ())
+ |"-verbose-compat-notations" -> verb_compat_ntn := true
+ |"-vm" -> use_vm := true
+ |"-where" -> print_where := true
+
+ (* Deprecated options *)
+ |"-byte" -> warning "option -byte deprecated, call with .byte suffix"
+ |"-opt" -> warning "option -opt deprecated, call with .opt suffix"
+ |"-full" -> warning "option -full deprecated"
+ |"-notactics" -> warning "Obsolete option \"-notactics\"."; remove_top_ml ()
+ |"-emacs-U" ->
+ warning "Obsolete option \"-emacs-U\", use -emacs instead."; set_emacs ()
+ |"-v7" -> error "This version of Coq does not support v7 syntax"
+ |"-v8" -> warning "Obsolete option \"-v8\"."
+ |"-lazy-load-proofs" -> warning "Obsolete option \"-lazy-load-proofs\"."
+ |"-dont-load-proofs" -> warning "Obsolete option \"-dont-load-proofs\"."
+ |"-force-load-proofs" -> warning "Obsolete option \"-force-load-proofs\"."
+ |"-unsafe" -> warning "Obsolete option \"-unsafe\"."; ignore (next ())
+ |"-quality" -> warning "Obsolete option \"-quality\"."
+ |"-xml" -> warning "Obsolete option \"-xml\"."
+
+ (* Unknown option *)
+ | s -> extras := s :: !extras
+ end;
+ parse ()
in
try
- parse arglist
+ parse ()
with
- | UserError(_,s) as e -> begin
- try
- Stream.empty s; exit 1
- with Stream.Failure ->
- msgnl (Errors.print e); exit 1
- end
- | any -> begin msgnl (Errors.print any); exit 1 end
+ | 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)
let init arglist =
+ init_gc ();
Sys.catch_break false; (* Ctrl-C is fatal during the initialisation *)
Lib.init();
(* Default Proofb Mode starts with an alternative default. *)
Goptions.set_string_option_value ["Default";"Proof";"Mode"] "Classic";
begin
try
- let foreign_args = parse_args arglist in
- if !filter_opts then
- (print_string (String.concat "\n" foreign_args); exit 0);
- if !ide_slave then begin
- Flags.make_silent true;
- Ide_slave.init_stdout ()
+ let extras = parse_args arglist in
+ (* 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;
+ 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 ()));
+ if !filter_opts then (print_string (String.concat "\n" extras); exit 0);
+ init_load_path ();
+ Option.iter Mltop.load_ml_object_raw !toploop;
+ let extras = !toploop_init extras in
+ if not (List.is_empty extras) then begin
+ prerr_endline ("Don't know what to do with "^String.concat " " extras);
+ prerr_endline "See -help for the list of supported options";
+ exit 1
end;
if_verbose print_header ();
- init_load_path ();
inputstate ();
Mltop.init_known_plugins ();
set_vm_opt ();
engage ();
+ set_hierarchy ();
(* 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|| !compile_list=[]) && Global.env_is_empty() then
- Option.iter Declaremods.start_library !toplevel_name;
+ if (not !batch_mode || List.is_empty !compile_list)
+ && Global.env_is_initial ()
+ then Option.iter Declaremods.start_library !toplevel_name;
init_library_roots ();
load_vernac_obj ();
require ();
+ Stm.init ();
load_rcfile();
load_vernacular ();
compile_files ();
+ schedule_vio_checking ();
+ schedule_vio_compilation ();
+ check_vio_tasks ();
outputstate ()
with any ->
+ let any = Errors.push any in
flush_all();
- if not !batch_mode then message "Error during initialization:";
- msgnl (Toplevel.print_toplevel_error any);
- exit 1
+ let msg =
+ if !batch_mode then mt ()
+ else str "Error during initialization:" ++ fnl ()
+ in
+ fatal_error (msg ++ Coqloop.print_toplevel_error any) (Errors.is_anomaly (fst any))
end;
- if !batch_mode then
- (flush_all();
- if !output_context then
- Pp.ppnl (with_option raw_print Prettyp.print_full_pure_context ());
- Profile.print_profile ();
- exit 0);
- (* We initialize the command history stack with a first entry *)
- Backtrack.mark_command Vernacexpr.VernacNop
+ if !batch_mode then begin
+ flush_all();
+ if !output_context then
+ Pp.ppnl (with_option raw_print Prettyp.print_full_pure_context ());
+ Profile.print_profile ();
+ exit 0
+ end
let init_toplevel = init
let start () =
- init_toplevel (List.tl (Array.to_list Sys.argv));
- if !ide_slave then
- Ide_slave.loop ()
- else
- Toplevel.loop();
- (* Initialise and launch the Ocaml toplevel *)
- Coqinit.init_ocaml_path();
- Mltop.ocaml_toploop();
+ 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 ... *)
+ !toploop_run ();
exit 1
(* [Coqtop.start] will be called by the code produced by coqmktop *)
diff --git a/toplevel/coqtop.mli b/toplevel/coqtop.mli
index 842ec4ef..356ccdcc 100644
--- a/toplevel/coqtop.mli
+++ b/toplevel/coqtop.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -9,8 +9,14 @@
(** The Coq main module. The following function [start] will parse the
command line, print the banner, initialize the load path, load the input
state, load the files given on the command line, load the ressource file,
- produce the output state if any, and finally will launch [Toplevel.loop]. *)
+ produce the output state if any, and finally will launch [Coqloop.loop]. *)
val init_toplevel : string list -> unit
val start : unit -> unit
+
+
+(* For other toploops *)
+val toploop_init : (string list -> string list) ref
+val toploop_run : (unit -> unit) ref
+
diff --git a/toplevel/discharge.ml b/toplevel/discharge.ml
index af0bbeaf..7d5d61fb 100644
--- a/toplevel/discharge.ml
+++ b/toplevel/discharge.ml
@@ -1,15 +1,17 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Names
+open Errors
open Util
-open Sign
+open Context
open Term
+open Vars
open Entries
open Declarations
open Cooking
@@ -20,7 +22,7 @@ open Cooking
let detype_param = function
| (Name id,None,p) -> id, Entries.LocalAssum p
| (Name id,Some p,_) -> id, Entries.LocalDef p
- | (Anonymous,_,_) -> anomaly"Unnamed inductive local variable"
+ | (Anonymous,_,_) -> anomaly (Pp.str "Unnamed inductive local variable")
(* Replace
@@ -37,29 +39,31 @@ 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 subs = list_tabulate (fun k -> lift nhyp (mkApp(mkRel (k+1),args))) ntyp in
+ let args = Array.of_list args in
+ let subs = List.init ntyp (fun k -> lift nhyp (mkApp(mkRel (k+1),args))) in
let inds' =
List.map
- (function (tname,arity,cnames,lc) ->
+ (function (tname,arity,template,cnames,lc) ->
let lc' = List.map (substl subs) lc in
let lc'' = List.map (fun b -> Termops.it_mkNamedProd_wo_LetIn b hyps) lc' in
let arity' = Termops.it_mkNamedProd_wo_LetIn arity hyps in
- (tname,arity',cnames,lc''))
+ (tname,arity',template,cnames,lc''))
inds in
let nparams' = nparams + Array.length args in
(* To be sure to be the same as before, should probably be moved to process_inductive *)
- let params' = let (_,arity,_,_) = List.hd inds' in
+ let params' = let (_,arity,_,_,_) = List.hd inds' in
let (params,_) = decompose_prod_n_assum nparams' arity in
List.map detype_param params
in
let ind'' =
List.map
- (fun (a,arity,c,lc) ->
+ (fun (a,arity,template,c,lc) ->
let _, short_arity = decompose_prod_n_assum nparams' arity in
let shortlc =
List.map (fun c -> snd (decompose_prod_n_assum nparams' c)) lc in
{ mind_entry_typename = a;
mind_entry_arity = short_arity;
+ mind_entry_template = template;
mind_entry_consnames = c;
mind_entry_lc = shortlc })
inds'
@@ -67,27 +71,49 @@ let abstract_inductive hyps nparams inds =
let refresh_polymorphic_type_of_inductive (_,mip) =
match mip.mind_arity with
- | Monomorphic s ->
- s.mind_user_arity
- | Polymorphic ar ->
- let ctx = List.rev mip.mind_arity_ctxt in
- mkArity (List.rev ctx,Termops.new_Type_sort())
+ | RegularArity s -> s.mind_user_arity, false
+ | TemplateArity ar ->
+ let ctx = List.rev mip.mind_arity_ctxt in
+ mkArity (List.rev ctx, Type ar.template_level), true
-let process_inductive sechyps modlist mib =
+let process_inductive (sechyps,abs_ctx) modlist mib =
let nparams = mib.mind_nparams in
+ let subst, univs =
+ if mib.mind_polymorphic then
+ let inst = Univ.UContext.instance mib.mind_universes in
+ let cstrs = Univ.UContext.constraints mib.mind_universes in
+ inst, Univ.UContext.make (inst, Univ.subst_instance_constraints inst cstrs)
+ else Univ.Instance.empty, mib.mind_universes
+ in
let inds =
- array_map_to_list
+ Array.map_to_list
(fun mip ->
- let arity = expmod_constr modlist (refresh_polymorphic_type_of_inductive (mib,mip)) in
- let lc = Array.map (expmod_constr modlist) mip.mind_user_lc in
- (mip.mind_typename,
- arity,
- Array.to_list mip.mind_consnames,
- Array.to_list lc))
+ let ty, template = refresh_polymorphic_type_of_inductive (mib,mip) in
+ let arity = expmod_constr modlist ty in
+ let arity = Vars.subst_instance_constr subst arity in
+ let lc = Array.map
+ (fun c -> Vars.subst_instance_constr subst (expmod_constr modlist c))
+ mip.mind_user_lc
+ in
+ (mip.mind_typename,
+ arity, template,
+ Array.to_list mip.mind_consnames,
+ Array.to_list lc))
mib.mind_packets in
let sechyps' = map_named_context (expmod_constr modlist) sechyps in
let (params',inds') = abstract_inductive sechyps' nparams inds in
- { mind_entry_record = mib.mind_record;
+ let abs_ctx = Univ.instantiate_univ_context abs_ctx in
+ let univs = Univ.UContext.union abs_ctx univs in
+ let record = match mib.mind_record with
+ | Some (Some (id, _, _)) -> Some (Some id)
+ | Some None -> Some None
+ | None -> None
+ in
+ { mind_entry_record = record;
mind_entry_finite = mib.mind_finite;
mind_entry_params = params';
- mind_entry_inds = inds' }
+ mind_entry_inds = inds';
+ mind_entry_polymorphic = mib.mind_polymorphic;
+ mind_entry_private = mib.mind_private;
+ mind_entry_universes = univs
+ }
diff --git a/toplevel/discharge.mli b/toplevel/discharge.mli
index 297a0731..386e4e3e 100644
--- a/toplevel/discharge.mli
+++ b/toplevel/discharge.mli
@@ -1,15 +1,15 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Sign
-open Cooking
+open Context
open Declarations
open Entries
+open Opaqueproof
val process_inductive :
- named_context -> work_list -> mutual_inductive_body -> mutual_inductive_entry
+ named_context Univ.in_universe_context -> work_list -> mutual_inductive_body -> mutual_inductive_entry
diff --git a/toplevel/g_obligations.ml4 b/toplevel/g_obligations.ml4
new file mode 100644
index 00000000..24661e12
--- /dev/null
+++ b/toplevel/g_obligations.ml4
@@ -0,0 +1,135 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \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 8dd18163..9341f2f7 100644
--- a/toplevel/himsg.ml
+++ b/toplevel/himsg.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,32 +8,131 @@
open Pp
open Util
-open Flags
open Names
open Nameops
open Namegen
open Term
open Termops
-open Inductive
open Indtypes
-open Sign
open Environ
open Pretype_errors
open Type_errors
open Typeclasses_errors
open Indrec
-open Reduction
open Cases
open Logic
open Printer
-open Glob_term
open Evd
-open Libnames
-open Declarations
+
+(* 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' ->
+ 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
+ l := (mkRel 1) :: List.map (Vars.lift 1) !l;
+ push_rel (na',c',t') env in
+ let env = process_rel_context contract_context env in
+ (env, List.map (Vars.substl !l) lc)
+
+let contract2 env a b = match contract env [a;b] with
+ | env, [a;b] -> env,a,b | _ -> assert false
+
+let contract3 env a b c = match contract env [a;b;c] with
+ | env, [a;b;c] -> env,a,b,c | _ -> assert false
+
+let contract4 env a b c d = match contract env [a;b;c;d] with
+ | env, [a;b;c;d] -> (env,a,b,c),d | _ -> assert false
+
+let contract1_vect env a v =
+ match contract env (a :: Array.to_list v) with
+ | env, a::l -> env,a,Array.of_list l
+ | _ -> assert false
+
+let rec contract3' env a b c = function
+ | OccurCheck (evk,d) -> let x,d = contract4 env a b c d in x,OccurCheck(evk,d)
+ | NotClean ((evk,args),env',d) ->
+ let env',d,args = contract1_vect env' d args in
+ contract3 env a b c,NotClean((evk,args),env',d)
+ | ConversionFailed (env',t1,t2) ->
+ let (env',t1,t2) = contract2 env' t1 t2 in
+ contract3 env a b c, ConversionFailed (env',t1,t2)
+ | NotSameArgSize | NotSameHead | NoCanonicalStructure
+ | MetaOccurInBody _ | InstanceNotSameType _
+ | UnifUnivInconsistency _ as x -> contract3 env a b c, x
+ | CannotSolveConstraint ((pb,env,t,u),x) ->
+ let env,t,u = contract2 env t u in
+ let y,x = contract3' env a b c x in
+ y,CannotSolveConstraint ((pb,env,t,u),x)
+
+(** Printers *)
let pr_lconstr c = quote (pr_lconstr c)
-let pr_lconstr_env e c = quote (pr_lconstr_env e c)
-let pr_ljudge_env e c = let v,t = pr_ljudge_env e c in (quote v,quote t)
+let pr_lconstr_env e s c = quote (pr_lconstr_env e s c)
+let pr_ljudge_env e s c = let v,t = pr_ljudge_env e s c in (quote v,quote t)
+
+(** A canonisation procedure for constr such that comparing there
+ externalisation catches more equalities *)
+let canonize_constr c =
+ (* replaces all the names in binders by [dn] ("default name"),
+ ensures that [alpha]-equivalent terms will have the same
+ externalisation. *)
+ let dn = Name.Anonymous in
+ let rec canonize_binders c =
+ match Term.kind_of_term c with
+ | Prod (_,t,b) -> mkProd(dn,t,b)
+ | Lambda (_,t,b) -> mkLambda(dn,t,b)
+ | LetIn (_,u,t,b) -> mkLetIn(dn,u,t,b)
+ | _ -> Term.map_constr canonize_binders c
+ in
+ canonize_binders c
+
+(** Tries to realize when the two terms, albeit different are printed the same. *)
+let display_eq ~flags env sigma t1 t2 =
+ (* terms are canonized, then their externalisation is compared syntactically *)
+ let open Constrextern in
+ let t1 = canonize_constr t1 in
+ let t2 = canonize_constr t2 in
+ let ct1 = Flags.with_options flags (fun () -> extern_constr false env sigma t1) () in
+ let ct2 = Flags.with_options flags (fun () -> extern_constr false env sigma t2) () in
+ Constrexpr_ops.constr_expr_eq ct1 ct2
+
+(** This function adds some explicit printing flags if the two arguments are
+ printed alike. *)
+let rec pr_explicit_aux env sigma t1 t2 = function
+| [] ->
+ (** no specified flags: default. *)
+ (quote (Printer.pr_lconstr_env env sigma t1), quote (Printer.pr_lconstr_env env sigma t2))
+| flags :: rem ->
+ let equal = display_eq ~flags env sigma t1 t2 in
+ if equal then
+ (** The two terms are the same from the user point of view *)
+ pr_explicit_aux env sigma t1 t2 rem
+ else
+ let open Constrextern in
+ let ct1 = Flags.with_options flags (fun () -> extern_constr false env sigma t1) ()
+ in
+ let ct2 = Flags.with_options flags (fun () -> extern_constr false env sigma t2) ()
+ in
+ quote (Ppconstr.pr_lconstr_expr ct1), quote (Ppconstr.pr_lconstr_expr ct2)
+
+let explicit_flags =
+ let open Constrextern in
+ [ []; (** First, try with the current flags *)
+ [print_implicits]; (** Then with implicit *)
+ [print_universes]; (** Then with universes *)
+ [print_universes; print_implicits]; (** With universes AND implicits *)
+ [print_implicits; print_coercions; print_no_symbol]; (** Then more! *)
+ [print_universes; print_implicits; print_coercions; print_no_symbol] (** and more! *) ]
+
+let pr_explicit env sigma t1 t2 = pr_explicit_aux env sigma t1 t2 explicit_flags
let pr_db env i =
try
@@ -42,8 +141,8 @@ let pr_db env i =
| Anonymous, _, _ -> str "<>"
with Not_found -> str "UNBOUND_REL_" ++ int i
-let explain_unbound_rel env n =
- let pe = pr_ne_context_of (str "In environment") env in
+let explain_unbound_rel env sigma n =
+ let pe = pr_ne_context_of (str "In environment") env sigma in
str "Unbound reference: " ++ pe ++
str "The reference " ++ int n ++ str " is free."
@@ -52,24 +151,25 @@ let explain_unbound_var env v =
str "No such section variable or assumption: " ++ var ++ str "."
let explain_not_type env sigma j =
- let j = j_nf_evar sigma j in
- let pe = pr_ne_context_of (str "In environment") env in
- let pc,pt = pr_ljudge_env env j in
+ let j = Evarutil.j_nf_evar sigma j in
+ let pe = pr_ne_context_of (str "In environment") env sigma in
+ let pc,pt = pr_ljudge_env env sigma j in
pe ++ str "The term" ++ brk(1,1) ++ pc ++ spc () ++
str "has type" ++ spc () ++ pt ++ spc () ++
str "which should be Set, Prop or Type."
-let explain_bad_assumption env j =
- let pe = pr_ne_context_of (str "In environment") env in
- let pc,pt = pr_ljudge_env env j in
+let explain_bad_assumption env sigma j =
+ let pe = pr_ne_context_of (str "In environment") env sigma in
+ let pc,pt = pr_ljudge_env env sigma j in
pe ++ str "Cannot declare a variable or hypothesis over the term" ++
brk(1,1) ++ pc ++ spc () ++ str "of type" ++ spc () ++ pt ++ spc () ++
str "because this term is not a type."
-let explain_reference_variables c =
- let pc = pr_lconstr c in
- str "The constant" ++ spc () ++ pc ++ spc () ++
- str "refers to variables which are not in the context."
+let explain_reference_variables id c =
+ (* c is intended to be a global reference *)
+ let pc = pr_global (Globnames.global_of_constr c) in
+ pc ++ strbrk " depends on the variable " ++ pr_id id ++
+ strbrk " which is not declared in the context."
let rec pr_disjunction pr = function
| [a] -> pr a
@@ -77,10 +177,16 @@ let rec pr_disjunction pr = function
| a::l -> pr a ++ str "," ++ spc () ++ pr_disjunction pr l
| [] -> assert false
-let explain_elim_arity env ind sorts c pj okinds =
+let pr_puniverses f env (c,u) =
+ f env c ++
+ (if Flags.is_universe_polymorphism () && not (Univ.Instance.is_empty u) then
+ str"(*" ++ Univ.Instance.pr Universes.pr_with_global_universes u ++ str"*)"
+ else mt())
+
+let explain_elim_arity env sigma ind sorts c pj okinds =
let env = make_all_name_different env in
- let pi = pr_inductive env ind in
- let pc = pr_lconstr_env env c in
+ let pi = pr_inductive env (fst ind) in
+ let pc = pr_lconstr_env env sigma c in
let msg = match okinds with
| Some(kp,ki,explanation) ->
let pki = pr_sort_family ki in
@@ -93,7 +199,7 @@ let explain_elim_arity env ind sorts c pj okinds =
| WrongArity ->
"wrong arity" in
let ppar = pr_disjunction (fun s -> quote (pr_sort_family s)) sorts in
- let ppt = pr_lconstr_env env ((strip_prod_assum pj.uj_type)) in
+ let ppt = pr_lconstr_env env sigma ((strip_prod_assum pj.uj_type)) in
hov 0
(str "the return type has sort" ++ spc () ++ ppt ++ spc () ++
str "while it" ++ spc () ++ str "should be " ++ ppar ++ str ".") ++
@@ -112,10 +218,10 @@ let explain_elim_arity env ind sorts c pj okinds =
fnl () ++ msg
let explain_case_not_inductive env sigma cj =
- let cj = j_nf_evar sigma cj in
+ let cj = Evarutil.j_nf_evar sigma cj in
let env = make_all_name_different env in
- let pc = pr_lconstr_env env cj.uj_val in
- let pct = pr_lconstr_env env cj.uj_type in
+ let pc = pr_lconstr_env env sigma cj.uj_val in
+ let pct = pr_lconstr_env env sigma cj.uj_type in
match kind_of_term cj.uj_type with
| Evar _ ->
str "Cannot infer a type for this expression."
@@ -125,116 +231,177 @@ let explain_case_not_inductive env sigma cj =
str "which is not a (co-)inductive type."
let explain_number_branches env sigma cj expn =
- let cj = j_nf_evar sigma cj in
+ let cj = Evarutil.j_nf_evar sigma cj in
let env = make_all_name_different env in
- let pc = pr_lconstr_env env cj.uj_val in
- let pct = pr_lconstr_env env cj.uj_type in
+ let pc = pr_lconstr_env env sigma cj.uj_val in
+ let pct = pr_lconstr_env env sigma cj.uj_type in
str "Matching on term" ++ brk(1,1) ++ pc ++ spc () ++
str "of type" ++ brk(1,1) ++ pct ++ spc () ++
str "expects " ++ int expn ++ str " branches."
let explain_ill_formed_branch env sigma c ci actty expty =
- let simp t = Reduction.nf_betaiota (nf_evar sigma t) in
- let c = nf_evar sigma c in
+ let simp t = Reduction.nf_betaiota env (Evarutil.nf_evar sigma t) in
+ let c = Evarutil.nf_evar sigma c in
let env = make_all_name_different env in
- let pc = pr_lconstr_env env c in
- let pa = pr_lconstr_env env (simp actty) in
- let pe = pr_lconstr_env env (simp expty) in
+ let pc = pr_lconstr_env env sigma c in
+ let pa, pe = pr_explicit env sigma (simp actty) (simp expty) in
strbrk "In pattern-matching on term" ++ brk(1,1) ++ pc ++
spc () ++ strbrk "the branch for constructor" ++ spc () ++
- quote (pr_constructor env ci) ++
+ quote (pr_puniverses pr_constructor env ci) ++
spc () ++ str "has type" ++ brk(1,1) ++ pa ++ spc () ++
str "which should be" ++ brk(1,1) ++ pe ++ str "."
-let explain_generalization env (name,var) j =
- let pe = pr_ne_context_of (str "In environment") env in
- let pv = pr_ltype_env env var in
- let (pc,pt) = pr_ljudge_env (push_rel_assum (name,var) env) j in
+let explain_generalization env sigma (name,var) j =
+ let pe = pr_ne_context_of (str "In environment") env sigma in
+ let pv = pr_ltype_env env sigma var in
+ let (pc,pt) = pr_ljudge_env (push_rel_assum (name,var) env) sigma j in
pe ++ str "Cannot generalize" ++ brk(1,1) ++ pv ++ spc () ++
str "over" ++ brk(1,1) ++ pc ++ str "," ++ spc () ++
str "it has type" ++ spc () ++ pt ++
spc () ++ str "which should be Set, Prop or Type."
-let explain_actual_type env sigma j pt =
- let j = j_nf_betaiotaevar sigma j in
- let pt = Reductionops.nf_betaiota sigma pt in
- let pe = pr_ne_context_of (str "In environment") env in
- let (pc,pct) = pr_ljudge_env env j in
- let pt = pr_lconstr_env env pt in
+let rec explain_unification_error env sigma p1 p2 = function
+ | None -> mt()
+ | Some e ->
+ let rec aux p1 p2 = function
+ | OccurCheck (evk,rhs) ->
+ let rhs = Evarutil.nf_evar sigma rhs in
+ [str "cannot define " ++ quote (pr_existential_key sigma evk) ++
+ strbrk " with term " ++ pr_lconstr_env env sigma rhs ++
+ strbrk " that would depend on itself"]
+ | NotClean ((evk,args),env,c) ->
+ let c = Evarutil.nf_evar sigma c in
+ let args = Array.map (Evarutil.nf_evar sigma) args in
+ [str "cannot instantiate " ++ quote (pr_existential_key sigma evk)
+ ++ strbrk " because " ++ pr_lconstr_env env sigma c ++
+ strbrk " is not in its scope" ++
+ (if Array.is_empty args then mt() else
+ strbrk ": available arguments are " ++
+ pr_sequence (pr_lconstr_env env sigma) (List.rev (Array.to_list args)))]
+ | NotSameArgSize | NotSameHead | NoCanonicalStructure ->
+ (* Error speaks from itself *) []
+ | ConversionFailed (env,t1,t2) ->
+ if Term.eq_constr t1 p1 && Term.eq_constr t2 p2 then [] else
+ let env = make_all_name_different env in
+ let t1 = Evarutil.nf_evar sigma t1 in
+ let t2 = Evarutil.nf_evar sigma t2 in
+ if not (Term.eq_constr t1 p1) || not (Term.eq_constr t2 p2) then
+ let t1, t2 = pr_explicit env sigma t1 t2 in
+ [str "cannot unify " ++ t1 ++ strbrk " and " ++ t2]
+ else []
+ | MetaOccurInBody evk ->
+ [str "instance for " ++ quote (pr_existential_key sigma evk) ++
+ strbrk " refers to a metavariable - please report your example"]
+ | InstanceNotSameType (evk,env,t,u) ->
+ let t, u = pr_explicit env sigma t u in
+ [str "unable to find a well-typed instantiation for " ++
+ quote (pr_existential_key sigma evk) ++
+ strbrk ": cannot ensure that " ++
+ t ++ strbrk " is a subtype of " ++ u]
+ | UnifUnivInconsistency p ->
+ if !Constrextern.print_universes then
+ [str "universe inconsistency: " ++
+ Univ.explain_universe_inconsistency Universes.pr_with_global_universes p]
+ else
+ [str "universe inconsistency"]
+ | CannotSolveConstraint ((pb,env,t,u),e) ->
+ let t = Evarutil.nf_evar sigma t in
+ let u = Evarutil.nf_evar sigma u in
+ (strbrk "cannot satisfy constraint " ++ pr_lconstr_env env sigma t ++
+ str " == " ++ pr_lconstr_env env sigma u)
+ :: aux t u e
+ in
+ match aux p1 p2 e with
+ | [] -> mt ()
+ | l -> spc () ++ str "(" ++
+ prlist_with_sep pr_semicolon (fun x -> x) l ++ str ")"
+
+let explain_actual_type env sigma j t reason =
+ let env = make_all_name_different env in
+ let j = Evarutil.j_nf_betaiotaevar sigma j in
+ let t = Reductionops.nf_betaiota sigma t in
+ (** Actually print *)
+ let pe = pr_ne_context_of (str "In environment") env sigma in
+ let pc = pr_lconstr_env env sigma (Environ.j_val j) in
+ let (pt, pct) = pr_explicit env sigma t (Environ.j_type j) in
+ let ppreason = explain_unification_error env sigma j.uj_type t reason in
pe ++
+ hov 0 (
str "The term" ++ brk(1,1) ++ pc ++ spc () ++
- str "has type" ++ brk(1,1) ++ pct ++ brk(1,1) ++
- str "while it is expected to have type" ++ brk(1,1) ++ pt ++ str "."
+ str "has type" ++ brk(1,1) ++ pct ++ spc () ++
+ str "while it is expected to have type" ++ brk(1,1) ++ pt ++
+ ppreason ++ str ".")
let explain_cant_apply_bad_type env sigma (n,exptyp,actualtyp) rator randl =
- let randl = jv_nf_betaiotaevar sigma randl in
- let exptyp = nf_evar sigma exptyp in
+ let randl = Evarutil.jv_nf_betaiotaevar sigma randl in
+ let exptyp = Evarutil.nf_evar sigma exptyp in
let actualtyp = Reductionops.nf_betaiota sigma actualtyp in
- let rator = j_nf_evar sigma rator in
+ let rator = Evarutil.j_nf_evar sigma rator in
let env = make_all_name_different env in
+ let actualtyp, exptyp = pr_explicit env sigma actualtyp exptyp in
let nargs = Array.length randl in
-(* let pe = pr_ne_context_of (str "in environment") env in*)
- let pr,prt = pr_ljudge_env env rator in
- let term_string1 = str (plural nargs "term") in
+(* let pe = pr_ne_context_of (str "in environment") env sigma in*)
+ let pr,prt = pr_ljudge_env env sigma rator in
+ let term_string1 = str (String.plural nargs "term") in
let term_string2 =
- if nargs>1 then str "The " ++ nth n ++ str " term" else str "This term" in
- let appl = prvect_with_sep pr_fnl
+ if nargs>1 then str "The " ++ pr_nth n ++ str " term" else str "This term"
+ in
+ let appl = prvect_with_sep fnl
(fun c ->
- let pc,pct = pr_ljudge_env env c in
+ let pc,pct = pr_ljudge_env env sigma c in
hov 2 (pc ++ spc () ++ str ": " ++ pct)) randl
in
- str "Illegal application (Type Error): " ++ (* pe ++ *) fnl () ++
+ str "Illegal application: " ++ (* pe ++ *) fnl () ++
str "The term" ++ brk(1,1) ++ pr ++ spc () ++
str "of type" ++ brk(1,1) ++ prt ++ spc () ++
str "cannot be applied to the " ++ term_string1 ++ fnl () ++
str " " ++ v 0 appl ++ fnl () ++ term_string2 ++ str " has type" ++
- brk(1,1) ++ pr_lconstr_env env actualtyp ++ spc () ++
+ brk(1,1) ++ actualtyp ++ spc () ++
str "which should be coercible to" ++ brk(1,1) ++
- pr_lconstr_env env exptyp ++ str "."
+ exptyp ++ str "."
let explain_cant_apply_not_functional env sigma rator randl =
- let randl = jv_nf_evar sigma randl in
- let rator = j_nf_evar sigma rator in
+ let randl = Evarutil.jv_nf_evar sigma randl in
+ let rator = Evarutil.j_nf_evar sigma rator in
let env = make_all_name_different env in
let nargs = Array.length randl in
-(* let pe = pr_ne_context_of (str "in environment") env in*)
- let pr = pr_lconstr_env env rator.uj_val in
- let prt = pr_lconstr_env env rator.uj_type in
- let appl = prvect_with_sep pr_fnl
+(* let pe = pr_ne_context_of (str "in environment") env sigma in*)
+ let pr = pr_lconstr_env env sigma rator.uj_val in
+ let prt = pr_lconstr_env env sigma rator.uj_type in
+ let appl = prvect_with_sep fnl
(fun c ->
- let pc = pr_lconstr_env env c.uj_val in
- let pct = pr_lconstr_env env c.uj_type in
+ let pc = pr_lconstr_env env sigma c.uj_val in
+ let pct = pr_lconstr_env env sigma c.uj_type in
hov 2 (pc ++ spc () ++ str ": " ++ pct)) randl
in
str "Illegal application (Non-functional construction): " ++
(* pe ++ *) fnl () ++
str "The expression" ++ brk(1,1) ++ pr ++ spc () ++
str "of type" ++ brk(1,1) ++ prt ++ spc () ++
- str "cannot be applied to the " ++ str (plural nargs "term") ++ fnl () ++
- str " " ++ v 0 appl
+ str "cannot be applied to the " ++ str (String.plural nargs "term") ++
+ fnl () ++ str " " ++ v 0 appl
let explain_unexpected_type env sigma actual_type expected_type =
- let actual_type = nf_evar sigma actual_type in
- let expected_type = nf_evar sigma expected_type in
- let pract = pr_lconstr_env env actual_type in
- let prexp = pr_lconstr_env env expected_type in
+ let actual_type = Evarutil.nf_evar sigma actual_type in
+ let expected_type = Evarutil.nf_evar sigma expected_type in
+ let pract, prexp = pr_explicit env sigma actual_type expected_type in
str "Found type" ++ spc () ++ pract ++ spc () ++
str "where" ++ spc () ++ prexp ++ str " was expected."
let explain_not_product env sigma c =
- let c = nf_evar sigma c in
- let pr = pr_lconstr_env env c in
+ let c = Evarutil.nf_evar sigma c in
+ let pr = pr_lconstr_env env sigma c in
str "The type of this term is a product" ++ spc () ++
str "while it is expected to be" ++
(if is_Type c then str " a sort" else (brk(1,1) ++ pr)) ++ str "."
(* TODO: use the names *)
(* (co)fixpoints *)
-let explain_ill_formed_rec_body env err names i fixenv vdefj =
+let explain_ill_formed_rec_body env sigma err names i fixenv vdefj =
let prt_name i =
match names.(i) with
Name id -> str "Recursive definition of " ++ pr_id id
- | Anonymous -> str "The " ++ nth i ++ str " definition" in
+ | Anonymous -> str "The " ++ pr_nth i ++ str " definition" in
let st = match err with
@@ -242,175 +409,188 @@ let explain_ill_formed_rec_body env err names i fixenv vdefj =
| NotEnoughAbstractionInFixBody ->
str "Not enough abstractions in the definition"
| RecursionNotOnInductiveType c ->
- str "Recursive definition on" ++ spc () ++ pr_lconstr_env env c ++ spc () ++
- str "which should be an inductive type"
+ str "Recursive definition on" ++ spc () ++ pr_lconstr_env env sigma c ++
+ spc () ++ str "which should be an inductive type"
| RecursionOnIllegalTerm(j,(arg_env, arg),le,lt) ->
let arg_env = make_all_name_different arg_env in
let called =
match names.(j) with
Name id -> pr_id id
- | Anonymous -> str "the " ++ nth i ++ str " definition" in
+ | Anonymous -> str "the " ++ pr_nth i ++ str " definition" in
let pr_db x = quote (pr_db env x) in
let vars =
match (lt,le) with
([],[]) -> assert false
| ([],[x]) -> str "a subterm of " ++ pr_db x
| ([],_) -> str "a subterm of the following variables: " ++
- prlist_with_sep pr_spc pr_db le
+ pr_sequence pr_db le
| ([x],_) -> pr_db x
| _ ->
str "one of the following variables: " ++
- prlist_with_sep pr_spc pr_db lt in
+ pr_sequence pr_db lt in
str "Recursive call to " ++ called ++ spc () ++
strbrk "has principal argument equal to" ++ spc () ++
- pr_lconstr_env arg_env arg ++ strbrk " instead of " ++ vars
+ pr_lconstr_env arg_env sigma arg ++ strbrk " instead of " ++ vars
| NotEnoughArgumentsForFixCall j ->
let called =
match names.(j) with
Name id -> pr_id id
- | Anonymous -> str "the " ++ nth i ++ str " definition" in
+ | Anonymous -> str "the " ++ pr_nth i ++ str " definition" in
str "Recursive call to " ++ called ++ str " has not enough arguments"
(* CoFixpoint guard errors *)
| CodomainNotInductiveType c ->
- str "The codomain is" ++ spc () ++ pr_lconstr_env env c ++ spc () ++
+ str "The codomain is" ++ spc () ++ pr_lconstr_env env sigma c ++ spc () ++
str "which should be a coinductive type"
| NestedRecursiveOccurrences ->
str "Nested recursive occurrences"
| UnguardedRecursiveCall c ->
- str "Unguarded recursive call in" ++ spc () ++ pr_lconstr_env env c
+ str "Unguarded recursive call in" ++ spc () ++ pr_lconstr_env env sigma c
| RecCallInTypeOfAbstraction c ->
str "Recursive call forbidden in the domain of an abstraction:" ++
- spc () ++ pr_lconstr_env env c
+ spc () ++ pr_lconstr_env env sigma c
| RecCallInNonRecArgOfConstructor c ->
str "Recursive call on a non-recursive argument of constructor" ++
- spc () ++ pr_lconstr_env env c
+ spc () ++ pr_lconstr_env env sigma c
| RecCallInTypeOfDef c ->
str "Recursive call forbidden in the type of a recursive definition" ++
- spc () ++ pr_lconstr_env env c
+ spc () ++ pr_lconstr_env env sigma c
| RecCallInCaseFun c ->
- str "Invalid recursive call in a branch of" ++ spc () ++ pr_lconstr_env env c
+ str "Invalid recursive call in a branch of" ++
+ spc () ++ pr_lconstr_env env sigma c
| RecCallInCaseArg c ->
str "Invalid recursive call in the argument of \"match\" in" ++ spc () ++
- pr_lconstr_env env c
+ pr_lconstr_env env sigma c
| RecCallInCasePred c ->
- str "Invalid recursive call in the \"return\" clause of \"match\" in" ++ spc () ++
- pr_lconstr_env env c
+ str "Invalid recursive call in the \"return\" clause of \"match\" in" ++
+ spc () ++ pr_lconstr_env env sigma c
| NotGuardedForm c ->
- str "Sub-expression " ++ pr_lconstr_env env c ++
+ str "Sub-expression " ++ pr_lconstr_env env sigma c ++
strbrk " not in guarded form (should be a constructor," ++
strbrk " an abstraction, a match, a cofix or a recursive call)"
+ | ReturnPredicateNotCoInductive c ->
+ str "The return clause of the following pattern matching should be" ++
+ strbrk " a coinductive type:" ++
+ spc () ++ pr_lconstr_env env sigma c
in
prt_name i ++ str " is ill-formed." ++ fnl () ++
- pr_ne_context_of (str "In environment") env ++
+ pr_ne_context_of (str "In environment") env sigma ++
st ++ str "." ++ fnl () ++
(try (* May fail with unresolved globals. *)
let fixenv = make_all_name_different fixenv in
- let pvd = pr_lconstr_env fixenv vdefj.(i).uj_val 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 ())
let explain_ill_typed_rec_body env sigma i names vdefj vargs =
- let vdefj = jv_nf_evar sigma vdefj in
- let vargs = Array.map (nf_evar sigma) vargs in
+ let vdefj = Evarutil.jv_nf_evar sigma vdefj in
+ let vargs = Array.map (Evarutil.nf_evar sigma) vargs in
let env = make_all_name_different env in
- let pvd,pvdt = pr_ljudge_env env (vdefj.(i)) in
- let pv = pr_lconstr_env env vargs.(i) in
+ let pvd = pr_lconstr_env env sigma vdefj.(i).uj_val in
+ let pvdt, pv = pr_explicit env sigma vdefj.(i).uj_type vargs.(i) in
str "The " ++
- (if Array.length vdefj = 1 then mt () else nth (i+1) ++ spc ()) ++
+ (match vdefj with [|_|] -> mt () | _ -> pr_nth (i+1) ++ spc ()) ++
str "recursive definition" ++ spc () ++ pvd ++ spc () ++
str "has type" ++ spc () ++ pvdt ++ spc () ++
str "while it should be" ++ spc () ++ pv ++ str "."
let explain_cant_find_case_type env sigma c =
- let c = nf_evar sigma c in
+ let c = Evarutil.nf_evar sigma c in
let env = make_all_name_different env in
- let pe = pr_lconstr_env env c in
+ let pe = pr_lconstr_env env sigma c in
str "Cannot infer type of pattern-matching on" ++ ws 1 ++ pe ++ str "."
let explain_occur_check env sigma ev rhs =
- let rhs = nf_evar sigma rhs in
+ let rhs = Evarutil.nf_evar sigma rhs in
let env = make_all_name_different env in
- let id = Evd.string_of_existential ev in
- let pt = pr_lconstr_env env rhs in
- str "Cannot define " ++ str id ++ str " with term" ++ brk(1,1) ++
- pt ++ spc () ++ str "that would depend on itself."
-
-let pr_ne_context_of header footer env =
- if Environ.rel_context env = empty_rel_context &
- Environ.named_context env = empty_named_context
- then footer
- else pr_ne_context_of header env
-
-let explain_hole_kind env evi = function
- | QuestionMark _ -> str "this placeholder"
- | CasesType ->
- str "the type of this pattern-matching problem"
- | BinderType (Name id) ->
- str "the type of " ++ Nameops.pr_id id
- | BinderType Anonymous ->
- str "the type of this anonymous binder"
- | ImplicitArg (c,(n,ido),b) ->
+ let pt = pr_lconstr_env env sigma rhs in
+ str "Cannot define " ++ pr_existential_key sigma ev ++ str " with term" ++
+ brk(1,1) ++ pt ++ spc () ++ str "that would depend on itself."
+
+let pr_trailing_ne_context_of env sigma =
+ if List.is_empty (Environ.rel_context env) &&
+ List.is_empty (Environ.named_context env)
+ then str "."
+ else (str " in environment:"++ pr_context_unlimited env sigma)
+
+let rec explain_evar_kind env sigma evk ty = function
+ | Evar_kinds.QuestionMark _ ->
+ strbrk "this placeholder of type " ++ ty
+ | Evar_kinds.CasesType false ->
+ strbrk "the type of this pattern-matching problem"
+ | Evar_kinds.CasesType true ->
+ strbrk "a subterm of type " ++ ty ++
+ strbrk " in the type of this pattern-matching problem"
+ | Evar_kinds.BinderType (Name id) ->
+ strbrk "the type of " ++ Nameops.pr_id id
+ | Evar_kinds.BinderType Anonymous ->
+ strbrk "the type of this anonymous binder"
+ | Evar_kinds.ImplicitArg (c,(n,ido),b) ->
let id = Option.get ido in
- str "the implicit parameter " ++
- pr_id id ++ spc () ++ str "of" ++
- spc () ++ Nametab.pr_global_env Idset.empty c
- | InternalHole ->
- str "an internal placeholder" ++
- Option.cata (fun evi ->
- let env = Evd.evar_env evi in
- str " of type " ++ pr_lconstr_env env evi.evar_concl ++
- pr_ne_context_of (str " in environment:"++ fnl ()) (mt ()) env)
- (mt ()) evi
- | TomatchTypeParameter (tyi,n) ->
- str "the " ++ nth n ++
- str " argument of the inductive type (" ++ pr_inductive env tyi ++
- str ") of this term"
- | GoalEvar ->
- str "an existential variable"
- | ImpossibleCase ->
- str "the type of an impossible pattern-matching clause"
- | MatchingVar _ ->
+ strbrk "the implicit parameter " ++ pr_id id ++ spc () ++ str "of" ++
+ spc () ++ Nametab.pr_global_env Id.Set.empty c ++
+ strbrk " whose type is " ++ ty
+ | Evar_kinds.InternalHole -> strbrk "an internal placeholder of type " ++ ty
+ | Evar_kinds.TomatchTypeParameter (tyi,n) ->
+ strbrk "the " ++ pr_nth n ++
+ strbrk " argument of the inductive type (" ++ pr_inductive env tyi ++
+ strbrk ") of this term"
+ | Evar_kinds.GoalEvar ->
+ strbrk "an existential variable of type " ++ ty
+ | Evar_kinds.ImpossibleCase ->
+ strbrk "the type of an impossible pattern-matching clause"
+ | Evar_kinds.MatchingVar _ ->
assert false
-
-let explain_not_clean env sigma ev t k =
- let t = nf_evar sigma t in
- let env = make_all_name_different env in
- let id = Evd.string_of_existential ev in
- let var = pr_lconstr_env env t in
- str "Tried to instantiate " ++ explain_hole_kind env None k ++
- str " (" ++ str id ++ str ")" ++ spc () ++
- str "with a term using variable " ++ var ++ spc () ++
- str "which is not in its scope."
-
-let explain_unsolvability = function
- | None -> mt()
- | Some (SeveralInstancesFound n) ->
- strbrk " (several distinct possible instances found)"
-
-let explain_typeclass_resolution env evi k =
+ | Evar_kinds.VarInstance id ->
+ strbrk "an instance of type " ++ ty ++
+ str " for the variable " ++ pr_id id
+ | Evar_kinds.SubEvar evk' ->
+ let evi = Evd.find sigma evk' in
+ let pc = match evi.evar_body with
+ | Evar_defined c -> pr_lconstr_env env sigma (Evarutil.nf_evar sigma c)
+ | Evar_empty -> assert false in
+ let ty' = Evarutil.nf_evar sigma evi.evar_concl in
+ pr_existential_key sigma evk ++ str " of type " ++ ty ++
+ str " in the partial instance " ++ pc ++
+ str " found for " ++ explain_evar_kind env sigma evk'
+ (pr_lconstr_env env sigma ty') (snd evi.evar_source)
+
+let explain_typeclass_resolution env sigma evi k =
match Typeclasses.class_of_constr evi.evar_concl with
- | Some c ->
- let env = Evd.evar_env evi in
+ | Some _ ->
+ let env = Evd.evar_filtered_env evi in
fnl () ++ str "Could not find an instance for " ++
- pr_lconstr_env env evi.evar_concl ++
- pr_ne_context_of (str " in environment:"++ fnl ()) (str ".") env
+ pr_lconstr_env env sigma evi.evar_concl ++
+ pr_trailing_ne_context_of env sigma
| _ -> mt()
-let explain_unsolvable_implicit env evi k explain =
- str "Cannot infer " ++ explain_hole_kind env (Some evi) k ++
- explain_unsolvability explain ++ str "." ++
- explain_typeclass_resolution env evi k
+let explain_placeholder_kind env sigma c e =
+ match e with
+ | Some (SeveralInstancesFound n) ->
+ strbrk " (several distinct possible type class instances found)"
+ | None ->
+ match Typeclasses.class_of_constr c with
+ | Some _ -> strbrk " (no type class instance found)"
+ | _ -> mt ()
+
+let explain_unsolvable_implicit env sigma evk explain =
+ let evi = Evarutil.nf_evar_info sigma (Evd.find_undefined sigma evk) in
+ let env = Evd.evar_filtered_env evi in
+ let type_of_hole = pr_lconstr_env env sigma evi.evar_concl in
+ let pe = pr_trailing_ne_context_of env sigma in
+ strbrk "Cannot infer " ++
+ explain_evar_kind env sigma evk type_of_hole (snd evi.evar_source) ++
+ explain_placeholder_kind env sigma evi.evar_concl explain ++ pe
let explain_var_not_found env id =
str "The variable" ++ spc () ++ pr_id id ++
spc () ++ str "was not found" ++
spc () ++ str "in the current" ++ spc () ++ str "environment" ++ str "."
-let explain_wrong_case_info env ind ci =
+let explain_wrong_case_info env (ind,u) ci =
let pi = pr_inductive (Global.env()) ind in
- if ci.ci_ind = ind then
+ if eq_ind ci.ci_ind ind then
str "Pattern-matching expression on an object of inductive type" ++
spc () ++ pi ++ spc () ++ str "has invalid information."
else
@@ -419,71 +599,86 @@ let explain_wrong_case_info env ind ci =
str "was given to a pattern-matching expression on the inductive type" ++
spc () ++ pc ++ str "."
-let explain_cannot_unify env sigma m n =
- let m = nf_evar sigma m in
- let n = nf_evar sigma n in
- let pm = pr_lconstr_env env m in
- let pn = pr_lconstr_env env n in
- str "Impossible to unify" ++ brk(1,1) ++ pm ++ spc () ++
- str "with" ++ brk(1,1) ++ pn ++ str "."
+let explain_cannot_unify env sigma m n e =
+ let env = make_all_name_different env in
+ let m = Evarutil.nf_evar sigma m in
+ let n = Evarutil.nf_evar sigma n in
+ let pm, pn = pr_explicit env sigma m n in
+ let ppreason = explain_unification_error env sigma m n e in
+ let pe = pr_ne_context_of (str "In environment") env sigma in
+ pe ++ str "Unable to unify" ++ brk(1,1) ++ pm ++ spc () ++
+ str "with" ++ brk(1,1) ++ pn ++ ppreason ++ str "."
let explain_cannot_unify_local env sigma m n subn =
- let pm = pr_lconstr_env env m in
- let pn = pr_lconstr_env env n in
- let psubn = pr_lconstr_env env subn in
- str "Impossible to unify" ++ brk(1,1) ++ pm ++ spc () ++
+ let pm = pr_lconstr_env env sigma m in
+ let pn = pr_lconstr_env env sigma n in
+ let psubn = pr_lconstr_env env sigma subn in
+ str "Unable to unify" ++ brk(1,1) ++ pm ++ spc () ++
str "with" ++ brk(1,1) ++ pn ++ spc () ++ str "as" ++ brk(1,1) ++
psubn ++ str " contains local variables."
-let explain_refiner_cannot_generalize env ty =
+let explain_refiner_cannot_generalize env sigma ty =
str "Cannot find a well-typed generalisation of the goal with type: " ++
- pr_lconstr_env env ty ++ str "."
+ pr_lconstr_env env sigma ty ++ str "."
-let explain_no_occurrence_found env c id =
- str "Found no subterm matching " ++ pr_lconstr_env env c ++
+let explain_no_occurrence_found env sigma c id =
+ str "Found no subterm matching " ++ pr_lconstr_env env sigma c ++
str " in " ++
(match id with
| Some id -> pr_id id
| None -> str"the current goal") ++ str "."
-let explain_cannot_unify_binding_type env m n =
- let pm = pr_lconstr_env env m in
- let pn = pr_lconstr_env env n in
+let explain_cannot_unify_binding_type env sigma m n =
+ let pm = pr_lconstr_env env sigma m in
+ let pn = pr_lconstr_env env sigma n in
str "This binding has type" ++ brk(1,1) ++ pm ++ spc () ++
str "which should be unifiable with" ++ brk(1,1) ++ pn ++ str "."
-let explain_cannot_find_well_typed_abstraction env p l =
+let explain_cannot_find_well_typed_abstraction env sigma p l e =
str "Abstracting over the " ++
- str (plural (List.length l) "term") ++ spc () ++
- hov 0 (pr_enum (pr_lconstr_env env) l) ++ spc () ++
- str "leads to a term" ++ spc () ++ pr_lconstr_env env p ++ spc () ++
- str "which is ill-typed."
+ str (String.plural (List.length l) "term") ++ spc () ++
+ hov 0 (pr_enum (pr_lconstr_env env sigma) l) ++ spc () ++
+ str "leads to a term" ++ spc () ++ pr_lconstr_goal_style_env env sigma p ++
+ spc () ++ str "which is ill-typed." ++
+ (match e with None -> mt () | Some e -> fnl () ++ str "Reason is: " ++ e)
+
+let explain_wrong_abstraction_type env sigma na abs expected result =
+ let ppname = match na with Name id -> pr_id id ++ spc () | _ -> mt () in
+ str "Cannot instantiate metavariable " ++ ppname ++ strbrk "of type " ++
+ pr_lconstr_env env sigma expected ++ strbrk " with abstraction " ++
+ pr_lconstr_env env sigma abs ++ strbrk " of incompatible type " ++
+ pr_lconstr_env env sigma result ++ str "."
let explain_abstraction_over_meta _ m n =
strbrk "Too complex unification problem: cannot find a solution for both " ++
pr_name m ++ spc () ++ str "and " ++ pr_name n ++ str "."
-let explain_non_linear_unification env m t =
+let explain_non_linear_unification env sigma m t =
strbrk "Cannot unambiguously instantiate " ++
pr_name m ++ str ":" ++
strbrk " which would require to abstract twice on " ++
- pr_lconstr_env env t ++ str "."
+ pr_lconstr_env env sigma t ++ str "."
+
+let explain_unsatisfied_constraints env sigma cst =
+ strbrk "Unsatisfied constraints: " ++
+ Univ.pr_constraints (Evd.pr_evd_level sigma) cst ++
+ spc () ++ str "(maybe a bugged tactic)."
let explain_type_error env sigma err =
let env = make_all_name_different env in
match err with
| UnboundRel n ->
- explain_unbound_rel env n
+ explain_unbound_rel env sigma n
| UnboundVar v ->
explain_unbound_var env v
| NotAType j ->
explain_not_type env sigma j
| BadAssumption c ->
- explain_bad_assumption env c
- | ReferenceVariables id ->
- explain_reference_variables id
+ explain_bad_assumption env sigma c
+ | ReferenceVariables (id,c) ->
+ explain_reference_variables id c
| ElimArity (ind, aritylst, c, pj, okinds) ->
- explain_elim_arity env ind aritylst c pj okinds
+ explain_elim_arity env sigma ind aritylst c pj okinds
| CaseNotInductive cj ->
explain_case_not_inductive env sigma cj
| NumberBranches (cj, n) ->
@@ -491,42 +686,128 @@ let explain_type_error env sigma err =
| IllFormedBranch (c, i, actty, expty) ->
explain_ill_formed_branch env sigma c i actty expty
| Generalization (nvar, c) ->
- explain_generalization env nvar c
+ explain_generalization env sigma nvar c
| ActualType (j, pt) ->
- explain_actual_type env sigma j pt
+ explain_actual_type env sigma j pt None
| CantApplyBadType (t, rator, randl) ->
explain_cant_apply_bad_type env sigma t rator randl
| CantApplyNonFunctional (rator, randl) ->
explain_cant_apply_not_functional env sigma rator randl
| IllFormedRecBody (err, lna, i, fixenv, vdefj) ->
- explain_ill_formed_rec_body env err lna i fixenv vdefj
+ explain_ill_formed_rec_body env sigma err lna i fixenv vdefj
| IllTypedRecBody (i, lna, vdefj, vargs) ->
explain_ill_typed_rec_body env sigma i lna vdefj vargs
| WrongCaseInfo (ind,ci) ->
explain_wrong_case_info env ind ci
+ | UnsatisfiedConstraints cst ->
+ explain_unsatisfied_constraints env sigma cst
+
+let pr_position (cl,pos) =
+ let clpos = match cl with
+ | None -> str " of the goal"
+ | Some (id,Locus.InHyp) -> str " of hypothesis " ++ pr_id id
+ | Some (id,Locus.InHypTypeOnly) -> str " of the type of hypothesis " ++ pr_id id
+ | Some (id,Locus.InHypValueOnly) -> str " of the body of hypothesis " ++ pr_id id in
+ int pos ++ clpos
+
+let explain_cannot_unify_occurrences env sigma nested ((cl2,pos2),t2) ((cl1,pos1),t1) e =
+ if nested then
+ str "Found nested occurrences of the pattern at positions " ++
+ int pos1 ++ strbrk " and " ++ pr_position (cl2,pos2) ++ str "."
+ else
+ let ppreason = match e with None -> mt() | Some (c1,c2,e) -> explain_unification_error env sigma c1 c2 (Some e) in
+ str "Found incompatible occurrences of the pattern" ++ str ":" ++
+ spc () ++ str "Matched term " ++ pr_lconstr_env env sigma t2 ++
+ strbrk " at position " ++ pr_position (cl2,pos2) ++
+ strbrk " is not compatible with matched term " ++
+ pr_lconstr_env env sigma t1 ++ strbrk " at position " ++
+ pr_position (cl1,pos1) ++ ppreason ++ str "."
+
+let pr_constraints printenv env sigma evars cstrs =
+ let (ev, evi) = Evar.Map.choose evars in
+ if Evar.Map.for_all (fun ev' evi' ->
+ eq_named_context_val evi.evar_hyps evi'.evar_hyps) evars
+ then
+ let l = Evar.Map.bindings evars in
+ let env' = reset_with_named_context evi.evar_hyps env in
+ let pe =
+ if printenv then
+ pr_ne_context_of (str "In environment:") env' sigma
+ else mt ()
+ in
+ let evs =
+ prlist
+ (fun (ev, evi) -> fnl () ++ pr_existential_key sigma ev ++
+ str " : " ++ pr_lconstr_env env' sigma evi.evar_concl) l
+ in
+ h 0 (pe ++ evs ++ pr_evar_constraints cstrs)
+ else
+ let filter evk _ = Evar.Map.mem evk evars in
+ pr_evar_map_filter ~with_univs:false filter sigma
+
+let explain_unsatisfiable_constraints env sigma constr comp =
+ let (_, constraints) = Evd.extract_all_conv_pbs sigma in
+ let undef = Evd.undefined_map (Evarutil.nf_evar_map_undefined sigma) in
+ (** Only keep evars that are subject to resolution and members of the given
+ component. *)
+ let is_kept evk evi = match comp with
+ | None -> Typeclasses.is_resolvable evi
+ | Some comp -> Typeclasses.is_resolvable evi && Evar.Set.mem evk comp
+ in
+ let undef =
+ let m = Evar.Map.filter is_kept undef in
+ if Evar.Map.is_empty m then undef
+ else m
+ in
+ match constr with
+ | None ->
+ str "Unable to satisfy the following constraints:" ++ fnl () ++
+ pr_constraints true env sigma undef constraints
+ | Some (ev, k) ->
+ let cstr =
+ let remaining = Evar.Map.remove ev undef in
+ if not (Evar.Map.is_empty remaining) then
+ str "With the following constraints:" ++ fnl () ++
+ pr_constraints false env sigma remaining constraints
+ else mt ()
+ in
+ let info = Evar.Map.find ev undef in
+ explain_typeclass_resolution env sigma info k ++ fnl () ++ cstr
let explain_pretype_error env sigma err =
- let env = env_nf_betaiotaevar sigma env in
+ let env = Evarutil.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
- | OccurCheck (n,c) -> explain_occur_check env sigma n c
- | NotClean (n,c,k) -> explain_not_clean env sigma n c k
- | UnsolvableImplicit (evi,k,exp) -> explain_unsolvable_implicit env evi k exp
+ | ActualTypeNotCoercible (j,t,e) ->
+ let {uj_val = c; uj_type = actty} = j in
+ let (env, c, actty, expty), e = contract3' env c actty t e in
+ let j = {uj_val = c; uj_type = actty} in
+ explain_actual_type env sigma j t (Some e)
+ | UnifOccurCheck (ev,rhs) -> explain_occur_check env sigma ev rhs
+ | UnsolvableImplicit (evk,exp) -> explain_unsolvable_implicit env sigma evk exp
| VarNotFound id -> explain_var_not_found env id
- | UnexpectedType (actual,expect) -> explain_unexpected_type env sigma actual expect
+ | UnexpectedType (actual,expect) ->
+ let env, actual, expect = contract2 env actual expect in
+ explain_unexpected_type env sigma actual expect
| NotProduct c -> explain_not_product env sigma c
- | CannotUnify (m,n) -> explain_cannot_unify env sigma m n
+ | CannotUnify (m,n,e) ->
+ let env, m, n = contract2 env m n in
+ explain_cannot_unify env sigma m n e
| CannotUnifyLocal (m,n,sn) -> explain_cannot_unify_local env sigma m n sn
- | CannotGeneralize ty -> explain_refiner_cannot_generalize env ty
- | NoOccurrenceFound (c, id) -> explain_no_occurrence_found env c id
- | CannotUnifyBindingType (m,n) -> explain_cannot_unify_binding_type env m n
- | CannotFindWellTypedAbstraction (p,l) ->
- explain_cannot_find_well_typed_abstraction env p l
+ | CannotGeneralize ty -> explain_refiner_cannot_generalize env sigma ty
+ | NoOccurrenceFound (c, id) -> explain_no_occurrence_found env sigma c id
+ | CannotUnifyBindingType (m,n) -> explain_cannot_unify_binding_type env sigma m n
+ | CannotFindWellTypedAbstraction (p,l,e) ->
+ explain_cannot_find_well_typed_abstraction env sigma p l
+ (Option.map (fun (env',e) -> explain_type_error env' sigma e) e)
+ | WrongAbstractionType (n,a,t,u) ->
+ explain_wrong_abstraction_type env sigma n a t u
| AbstractionOverMeta (m,n) -> explain_abstraction_over_meta env m n
- | NonLinearUnification (m,c) -> explain_non_linear_unification env m c
+ | NonLinearUnification (m,c) -> explain_non_linear_unification env sigma m c
| TypingError t -> explain_type_error env sigma t
-
+ | CannotUnifyOccurrences (b,c1,c2,e) -> explain_cannot_unify_occurrences env sigma b c1 c2 e
+ | UnsatisfiableConstraints (c,comp) -> explain_unsatisfiable_constraints env sigma c comp
(* Module errors *)
open Modops
@@ -541,14 +822,14 @@ let explain_not_match_error = function
| ModuleTypeFieldExpected ->
strbrk "a module type is expected"
| NotConvertibleInductiveField id | NotConvertibleConstructorField id ->
- str "types given to " ++ str (string_of_id id) ++ str " differ"
+ str "types given to " ++ str (Id.to_string id) ++ str " differ"
| NotConvertibleBodyField ->
str "the body of definitions differs"
| NotConvertibleTypeField (env, typ1, typ2) ->
str "expected type" ++ spc () ++
- quote (Printer.safe_pr_lconstr_env env typ2) ++ spc () ++
+ quote (Printer.safe_pr_lconstr_env env Evd.empty typ2) ++ spc () ++
str "but found type" ++ spc () ++
- quote (Printer.safe_pr_lconstr_env env typ1)
+ quote (Printer.safe_pr_lconstr_env env Evd.empty typ1)
| NotSameConstructorNamesField ->
str "constructor names differ"
| NotSameInductiveNameInBlockField ->
@@ -566,24 +847,45 @@ 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 (string_of_id id) | _ -> str "_") nal
+ pr_enum (function Name id -> str (Id.to_string id) | _ -> str "_") nal
| NotEqualInductiveAliases ->
str "Aliases to inductive types do not match"
| NoTypeConstraintExpected ->
- strbrk "a definition whose type is constrained can only be subtype of a definition whose type is itself constrained"
+ strbrk "a definition whose type is constrained can only be subtype " ++
+ strbrk "of a definition whose type is itself constrained"
+ | PolymorphicStatusExpected b ->
+ let status b = if b then str"polymorphic" else str"monomorphic" in
+ str "a " ++ status b ++ str" declaration was expected, but a " ++
+ status (not b) ++ str" declaration was found"
+ | IncompatibleInstances ->
+ str"polymorphic universe instances do not match"
+ | IncompatibleUniverses incon ->
+ str"the universe constraints are inconsistent: " ++
+ Univ.explain_universe_inconsistency Universes.pr_with_global_universes incon
+ | IncompatiblePolymorphism (env, t1, t2) ->
+ str "conversion of polymorphic values generates additional constraints: " ++
+ quote (Printer.safe_pr_lconstr_env env Evd.empty t1) ++ spc () ++
+ str "compared to " ++ spc () ++
+ quote (Printer.safe_pr_lconstr_env env Evd.empty t2)
+ | IncompatibleConstraints cst ->
+ str " the expected (polymorphic) constraints do not imply " ++
+ quote (Univ.pr_constraints (Evd.pr_evd_level Evd.empty) cst)
let explain_signature_mismatch l spec why =
- str "Signature components for label " ++ str (string_of_label l) ++
+ str "Signature components for label " ++ str (Label.to_string l) ++
str " do not match:" ++ spc () ++ explain_not_match_error why ++ str "."
let explain_label_already_declared l =
- str ("The label "^string_of_label l^" is already declared.")
+ str ("The label "^Label.to_string l^" is already declared.")
let explain_application_to_not_path _ =
str "Application of modules is restricted to paths."
-let explain_not_a_functor mtb =
- str "Application of not a functor."
+let explain_not_a_functor () =
+ str "Application of a non-functor."
+
+let explain_is_a_functor () =
+ str "Illegal use of a functor."
let explain_incompatible_module_types mexpr1 mexpr2 =
str "Incompatible module types."
@@ -592,20 +894,11 @@ let explain_not_equal_module_paths mp1 mp2 =
str "Non equal modules."
let explain_no_such_label l =
- str "No such label " ++ str (string_of_label l) ++ str "."
+ str "No such label " ++ str (Label.to_string l) ++ str "."
let explain_incompatible_labels l l' =
str "Opening and closing labels are not the same: " ++
- str (string_of_label l) ++ str " <> " ++ str (string_of_label l') ++ str "!"
-
-let explain_signature_expected mtb =
- str "Signature expected."
-
-let explain_no_module_to_end () =
- str "No open module to end."
-
-let explain_no_module_type_to_end () =
- str "No open module type to end."
+ str (Label.to_string l) ++ str " <> " ++ str (Label.to_string l') ++ str "!"
let explain_not_a_module s =
quote (str s) ++ str " is not a module."
@@ -614,45 +907,41 @@ let explain_not_a_module_type s =
quote (str s) ++ str " is not a module type."
let explain_not_a_constant l =
- quote (pr_label l) ++ str " is not a constant."
+ quote (Label.print l) ++ str " is not a constant."
let explain_incorrect_label_constraint l =
str "Incorrect constraint for label " ++
- quote (pr_label l) ++ str "."
+ quote (Label.print l) ++ str "."
let explain_generative_module_expected l =
- str "The module " ++ str (string_of_label l) ++
- strbrk " is not generative. Only components of generative modules can be changed using the \"with\" construct."
-
-let explain_non_empty_local_context = function
- | None -> str "The local context is not empty."
- | Some l ->
- str "The local context of the component " ++
- str (string_of_label l) ++ str " is not empty."
+ str "The module " ++ str (Label.to_string 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 (string_of_label l) ++ str " is missing in "
+ str "The field " ++ str (Label.to_string l) ++ str " is missing in "
++ str s ++ str "."
+let explain_higher_order_include () =
+ str "You cannot Include a higher-order structure."
+
let explain_module_error = function
| SignatureMismatch (l,spec,err) -> explain_signature_mismatch l spec err
| LabelAlreadyDeclared l -> explain_label_already_declared l
| ApplicationToNotPath mexpr -> explain_application_to_not_path mexpr
- | NotAFunctor mtb -> explain_not_a_functor mtb
+ | NotAFunctor -> explain_not_a_functor ()
+ | IsAFunctor -> explain_is_a_functor ()
| IncompatibleModuleTypes (m1,m2) -> explain_incompatible_module_types m1 m2
| NotEqualModulePaths (mp1,mp2) -> explain_not_equal_module_paths mp1 mp2
| NoSuchLabel l -> explain_no_such_label l
| IncompatibleLabels (l1,l2) -> explain_incompatible_labels l1 l2
- | SignatureExpected mtb -> explain_signature_expected mtb
- | NoModuleToEnd -> explain_no_module_to_end ()
- | NoModuleTypeToEnd -> explain_no_module_type_to_end ()
| NotAModule s -> explain_not_a_module s
| NotAModuleType s -> explain_not_a_module_type s
| NotAConstant l -> explain_not_a_constant l
| IncorrectWithConstraint l -> explain_incorrect_label_constraint l
| GenerativeModuleExpected l -> explain_generative_module_expected l
- | NonEmptyLocalContect lopt -> explain_non_empty_local_context lopt
| LabelMissing (l,s) -> explain_label_missing l s
+ | HigherOrderInclude -> explain_higher_order_include ()
(* Module internalization errors *)
@@ -681,72 +970,27 @@ let explain_module_internalization_error = function
(* Typeclass errors *)
let explain_not_a_class env c =
- pr_constr_env env c ++ str" is not a declared type class."
+ pr_constr_env env Evd.empty c ++ str" is not a declared type class."
let explain_unbound_method env cid id =
- str "Unbound method name " ++ Nameops.pr_id (snd id) ++ spc () ++ str"of class" ++ spc () ++
- pr_global cid ++ str "."
+ str "Unbound method name " ++ Nameops.pr_id (snd id) ++ spc () ++
+ str"of class" ++ spc () ++ pr_global cid ++ str "."
let pr_constr_exprs exprs =
hv 0 (List.fold_right
(fun d pps -> ws 2 ++ Ppconstr.pr_constr_expr d ++ pps)
exprs (mt ()))
-let explain_no_instance env (_,id) l =
- str "No instance found for class " ++ Nameops.pr_id id ++ spc () ++
- str "applied to arguments" ++ spc () ++
- prlist_with_sep pr_spc (pr_lconstr_env env) l
-
-let is_goal_evar evi = match evi.evar_source with (_, GoalEvar) -> true | _ -> false
-
-let pr_constraints printenv env evm =
- let l = Evd.to_list evm in
- assert(l <> []);
- let (ev, evi) = List.hd l in
- if List.for_all (fun (ev', evi') ->
- eq_named_context_val evi.evar_hyps evi'.evar_hyps) l
- then
- let pe = pr_ne_context_of (str "In environment:") (mt ())
- (reset_with_named_context evi.evar_hyps env) in
- (if printenv then pe ++ fnl () else mt ()) ++
- prlist_with_sep (fun () -> fnl ())
- (fun (ev, evi) -> str(string_of_existential ev) ++
- str " : " ++ pr_lconstr evi.evar_concl) l ++ fnl() ++
- pr_evar_map_constraints evm
- else
- pr_evar_map None evm
-
-let explain_unsatisfiable_constraints env evd constr =
- let evm = Evd.undefined_evars (Evarutil.nf_evar_map_undefined evd) in
- (* Remove goal evars *)
- let undef = fold_undefined
- (fun ev evi evm' ->
- if is_goal_evar evi then Evd.remove evm' ev else evm') evm evm
- in
- match constr with
- | None ->
- str"Unable to satisfy the following constraints:" ++ fnl() ++
- pr_constraints true env undef
- | Some (ev, k) ->
- explain_typeclass_resolution env (Evd.find evm ev) k ++ fnl () ++
- (let remaining = Evd.remove undef ev in
- if Evd.has_undefined remaining then
- str"With the following constraints:" ++ fnl() ++
- pr_constraints false env remaining
- else mt ())
-
let explain_mismatched_contexts env c i j =
str"Mismatched contexts while declaring instance: " ++ brk (1,1) ++
- hov 1 (str"Expected:" ++ brk (1, 1) ++ pr_rel_context env j) ++ fnl () ++ brk (1,1) ++
+ hov 1 (str"Expected:" ++ brk (1, 1) ++ pr_rel_context env Evd.empty j) ++
+ fnl () ++ brk (1,1) ++
hov 1 (str"Found:" ++ brk (1, 1) ++ pr_constr_exprs i)
-let explain_typeclass_error env err =
- match err with
- | NotAClass c -> explain_not_a_class env c
- | UnboundMethod (cid, id) -> explain_unbound_method env cid id
- | NoInstance (id, l) -> explain_no_instance env id l
- | UnsatisfiableConstraints (evd, c) -> explain_unsatisfiable_constraints env evd c
- | MismatchedContextInstance (c, i, j) -> explain_mismatched_contexts env c i j
+let explain_typeclass_error env = function
+ | NotAClass c -> explain_not_a_class env c
+ | UnboundMethod (cid, id) -> explain_unbound_method env cid id
+ | MismatchedContextInstance (c,i,j) -> explain_mismatched_contexts env c i j
(* Refiner errors *)
@@ -758,7 +1002,7 @@ let explain_refiner_bad_type arg ty conclty =
let explain_refiner_unresolved_bindings l =
str "Unable to find an instance for the " ++
- str (plural (List.length l) "variable") ++ spc () ++
+ str (String.plural (List.length l) "variable") ++ spc () ++
prlist_with_sep pr_comma pr_name l ++ str"."
let explain_refiner_cannot_apply t harg =
@@ -784,6 +1028,9 @@ let explain_meta_in_type c =
str "In refiner, a meta appears in the type " ++ brk(1,1) ++ pr_lconstr c ++
str " of another meta"
+let explain_no_such_hyp id =
+ str "No such hypothesis: " ++ pr_id id
+
let explain_refiner_error = function
| BadType (arg,ty,conclty) -> explain_refiner_bad_type arg ty conclty
| UnresolvedBindings t -> explain_refiner_unresolved_bindings t
@@ -793,50 +1040,51 @@ let explain_refiner_error = function
| DoesNotOccurIn (c,hyp) -> explain_does_not_occur_in c hyp
| NonLinearProof c -> explain_non_linear_proof c
| MetaInType c -> explain_meta_in_type c
+ | NoSuchHyp id -> explain_no_such_hyp id
(* Inductive errors *)
-let error_non_strictly_positive env c v =
- let pc = pr_lconstr_env env c in
- let pv = pr_lconstr_env env v in
+let error_non_strictly_positive env c v =
+ let pc = pr_lconstr_env env Evd.empty c in
+ let pv = pr_lconstr_env env Evd.empty v in
str "Non strictly positive occurrence of " ++ pv ++ str " in" ++
brk(1,1) ++ pc ++ str "."
let error_ill_formed_inductive env c v =
- let pc = pr_lconstr_env env c in
- let pv = pr_lconstr_env env v in
+ let pc = pr_lconstr_env env Evd.empty c in
+ let pv = pr_lconstr_env env Evd.empty v in
str "Not enough arguments applied to the " ++ pv ++
str " in" ++ brk(1,1) ++ pc ++ str "."
let error_ill_formed_constructor env id c v nparams nargs =
- let pv = pr_lconstr_env env v in
- let atomic = (nb_prod c = 0) in
+ let pv = pr_lconstr_env env Evd.empty v in
+ let atomic = Int.equal (nb_prod c) 0 in
str "The type of constructor" ++ brk(1,1) ++ pr_id id ++ brk(1,1) ++
str "is not valid;" ++ brk(1,1) ++
strbrk (if atomic then "it must be " else "its conclusion must be ") ++
pv ++
(* warning: because of implicit arguments it is difficult to say which
parameters must be explicitly given *)
- (if nparams<>0 then
- strbrk " applied to its " ++ str (plural nparams "parameter")
+ (if not (Int.equal nparams 0) then
+ strbrk " applied to its " ++ str (String.plural nparams "parameter")
else
mt()) ++
- (if nargs<>0 then
- str (if nparams<>0 then " and" else " applied") ++
- strbrk " to some " ++ str (plural nargs "argument")
+ (if not (Int.equal nargs 0) then
+ str (if not (Int.equal nparams 0) then " and" else " applied") ++
+ strbrk " to some " ++ str (String.plural nargs "argument")
else
mt()) ++ str "."
let pr_ltype_using_barendregt_convention_env env c =
(* Use goal_concl_style as an approximation of Barendregt's convention (?) *)
- quote (pr_goal_concl_style_env env c)
+ quote (pr_goal_concl_style_env env Evd.empty c)
let error_bad_ind_parameters env c n v1 v2 =
let pc = pr_ltype_using_barendregt_convention_env env c in
- let pv1 = pr_lconstr_env env v1 in
- let pv2 = pr_lconstr_env env v2 in
+ let pv1 = pr_lconstr_env env Evd.empty v1 in
+ let pv2 = pr_lconstr_env env Evd.empty v2 in
str "Last occurrence of " ++ pv2 ++ str " must have " ++ pv1 ++
- str " as " ++ nth n ++ str " argument in " ++ brk(1,1) ++ pc ++ str "."
+ str " as " ++ pr_nth n ++ str " argument in " ++ brk(1,1) ++ pc ++ str "."
let error_same_names_types id =
str "The name" ++ spc () ++ pr_id id ++ spc () ++
@@ -852,7 +1100,8 @@ let error_same_names_overlap idl =
prlist_with_sep pr_comma pr_id idl ++ str "."
let error_not_an_arity env c =
- str "The type" ++ spc () ++ pr_lconstr_env env c ++ spc () ++ str "is not an arity."
+ str "The type" ++ spc () ++ pr_lconstr_env env Evd.empty c ++ spc () ++
+ str "is not an arity."
let error_bad_entry () =
str "Bad inductive definition."
@@ -864,12 +1113,12 @@ let error_large_non_prop_inductive_not_in_type () =
let error_not_allowed_case_analysis isrec kind i =
str (if isrec then "Induction" else "Case analysis") ++
- strbrk " on sort " ++ pr_sort kind ++
+ strbrk " on sort " ++ pr_sort Evd.empty kind ++
strbrk " is not allowed for inductive definition " ++
- pr_inductive (Global.env()) i ++ str "."
+ pr_inductive (Global.env()) (fst i) ++ str "."
let error_not_mutual_in_scheme ind ind' =
- if ind = ind' then
+ if eq_ind ind ind' then
str "The inductive type " ++ pr_inductive (Global.env()) ind ++
str " occurs twice."
else
@@ -890,7 +1139,8 @@ let explain_inductive_error = function
| SameNamesOverlap idl -> error_same_names_overlap idl
| NotAnArity (env, c) -> error_not_an_arity env c
| BadEntry -> error_bad_entry ()
- | LargeNonPropInductiveNotInType -> error_large_non_prop_inductive_not_in_type ()
+ | LargeNonPropInductiveNotInType ->
+ error_large_non_prop_inductive_not_in_type ()
(* Recursion schemes errors *)
@@ -901,9 +1151,9 @@ let explain_recursion_scheme_error = function
(* Pattern-matching errors *)
-let explain_bad_pattern env cstr ty =
+let explain_bad_pattern env sigma cstr ty =
let env = make_all_name_different env in
- let pt = pr_lconstr_env env ty in
+ let pt = pr_lconstr_env env sigma ty in
let pc = pr_constructor env cstr in
str "Found the constructor " ++ pc ++ brk(1,1) ++
str "while matching a term of type " ++ pt ++ brk(1,1) ++
@@ -918,128 +1168,134 @@ let explain_bad_constructor env cstr ind =
str "is expected."
let decline_string n s =
- if n = 0 then "no "^s^"s"
- else if n = 1 then "1 "^s
+ if Int.equal n 0 then "no "^s^"s"
+ else if Int.equal n 1 then "1 "^s
else (string_of_int n^" "^s^"s")
let explain_wrong_numarg_constructor env cstr n =
str "The constructor " ++ pr_constructor env cstr ++
- str " expects " ++ str (decline_string n "argument") ++ str "."
+ str " (in type " ++ pr_inductive env (inductive_of_constructor cstr) ++
+ str ") expects " ++ str (decline_string n "argument") ++ str "."
let explain_wrong_numarg_inductive env ind n =
str "The inductive type " ++ pr_inductive env ind ++
str " expects " ++ str (decline_string n "argument") ++ str "."
-let explain_wrong_predicate_arity env pred nondep_arity dep_arity=
- let env = make_all_name_different env in
- let pp = pr_lconstr_env env pred in
- str "The elimination predicate " ++ spc () ++ pp ++ fnl () ++
- str "should be of arity" ++ spc () ++
- pr_lconstr_env env nondep_arity ++ spc () ++
- str "(for non dependent case) or" ++
- spc () ++ pr_lconstr_env env dep_arity ++ spc () ++ str "(for dependent case)."
-
-let explain_needs_inversion env x t =
- let env = make_all_name_different env in
- let px = pr_lconstr_env env x in
- let pt = pr_lconstr_env env t in
- str "Sorry, I need inversion to compile pattern matching on term " ++
- px ++ str " of type: " ++ pt ++ str "."
-
let explain_unused_clause env pats =
(* Without localisation
let s = if List.length pats > 1 then "s" else "" in
(str ("Unused clause with pattern"^s) ++ spc () ++
- hov 0 (prlist_with_sep pr_spc pr_cases_pattern pats) ++ str ")")
+ hov 0 (pr_sequence pr_cases_pattern pats) ++ str ")")
*)
str "This clause is redundant."
let explain_non_exhaustive env pats =
str "Non exhaustive pattern-matching: no clause found for " ++
- str (plural (List.length pats) "pattern") ++
- spc () ++ hov 0 (prlist_with_sep pr_spc pr_cases_pattern pats)
+ str (String.plural (List.length pats) "pattern") ++
+ spc () ++ hov 0 (pr_sequence pr_cases_pattern pats)
-let explain_cannot_infer_predicate env typs =
+let explain_cannot_infer_predicate env sigma typs =
let env = make_all_name_different env in
let pr_branch (cstr,typ) =
let cstr,_ = decompose_app cstr in
- str "For " ++ pr_lconstr_env env cstr ++ str ": " ++ pr_lconstr_env env typ
+ str "For " ++ pr_lconstr_env env sigma cstr ++ str ": " ++ pr_lconstr_env env sigma typ
in
str "Unable to unify the types found in the branches:" ++
- spc () ++ hov 0 (prlist_with_sep pr_fnl pr_branch (Array.to_list typs))
+ spc () ++ hov 0 (prlist_with_sep fnl pr_branch (Array.to_list typs))
-let explain_pattern_matching_error env = function
+let explain_pattern_matching_error env sigma = function
| BadPattern (c,t) ->
- explain_bad_pattern env c t
+ explain_bad_pattern env sigma c t
| BadConstructor (c,ind) ->
explain_bad_constructor env c ind
| WrongNumargConstructor (c,n) ->
explain_wrong_numarg_constructor env c n
| WrongNumargInductive (c,n) ->
explain_wrong_numarg_inductive env c n
- | WrongPredicateArity (pred,n,dep) ->
- explain_wrong_predicate_arity env pred n dep
- | NeedsInversion (x,t) ->
- explain_needs_inversion env x t
| UnusedClause tms ->
explain_unused_clause env tms
| NonExhaustive tms ->
explain_non_exhaustive env tms
| CannotInferPredicate typs ->
- explain_cannot_infer_predicate env typs
+ explain_cannot_infer_predicate env sigma typs
let explain_reduction_tactic_error = function
- | Tacred.InvalidAbstraction (env,c,(env',e)) ->
+ | Tacred.InvalidAbstraction (env,sigma,c,(env',e)) ->
str "The abstracted term" ++ spc () ++
- quote (pr_goal_concl_style_env env c) ++
+ quote (pr_goal_concl_style_env env sigma c) ++
spc () ++ str "is not well typed." ++ fnl () ++
explain_type_error env' Evd.empty e
-let explain_ltac_call_trace (nrep,last,trace,loc) =
- let calls =
- (nrep,last) :: List.rev (List.map(fun(n,_,ck)->(n,ck))trace) in
- let tacexpr_differ te te' =
- (* NB: The following comparison may raise an exception
- since a tacexpr may embed a functional part via a TacExtend *)
- try te <> te' with Invalid_argument _ -> false
+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
- let pr_call (n,ck) =
- (match ck with
- | Proof_type.LtacNotationCall s -> quote (str s)
- | Proof_type.LtacNameCall cst -> quote (Pptactic.pr_ltac_constant cst)
- | 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,otac) -> quote
- (Pptactic.pr_glob_tactic (Global.env())
- (Tacexpr.TacAtom (dummy_loc,te)))
- ++ (match !otac with
- | Some te' when tacexpr_differ (Obj.magic te') te ->
- strbrk " (expanded to " ++ quote
- (Pptactic.pr_tactic (Global.env())
- (Tacexpr.TacAtom (dummy_loc,te')))
- ++ str ")"
- | _ -> mt ())
- | Proof_type.LtacConstrInterp (c,(vars,unboundvars)) ->
- let filter =
- function (id,None) -> None | (id,Some id') -> Some(id,([],mkVar id')) in
- let unboundvars = list_map_filter filter unboundvars in
- quote (pr_glob_constr_env (Global.env()) c) ++
- (if unboundvars <> [] or vars <> [] then
- strbrk " (with " ++
- prlist_with_sep pr_comma
- (fun (id,c) ->
- pr_id id ++ str ":=" ++ Printer.pr_lconstr_under_binders c)
- (List.rev vars @ unboundvars) ++ str ")"
- else mt())) ++
- (if n=2 then str " (repeated twice)"
- else if n>2 then str " (repeated "++int n++str" times)"
- else mt()) in
- if calls <> [] then
- let kind_of_last_call = match list_last calls with
- | (_,Proof_type.LtacConstrInterp _) -> ", last term evaluation failed."
- | _ -> ", last call failed." 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
- mt ()
+ (* 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 b478497b..3d5442bb 100644
--- a/toplevel/himsg.mli
+++ b/toplevel/himsg.mli
@@ -1,13 +1,12 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Pp
-open Names
open Indtypes
open Environ
open Type_errors
@@ -32,14 +31,13 @@ val explain_recursion_scheme_error : recursion_scheme_error -> std_ppcmds
val explain_refiner_error : refiner_error -> std_ppcmds
val explain_pattern_matching_error :
- env -> pattern_matching_error -> std_ppcmds
+ env -> Evd.evar_map -> pattern_matching_error -> std_ppcmds
val explain_reduction_tactic_error :
Tacred.reduction_tactic_error -> std_ppcmds
-val explain_ltac_call_trace :
- int * Proof_type.ltac_call_kind * Proof_type.ltac_trace * Util.loc ->
- 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
diff --git a/toplevel/ide_intf.ml b/toplevel/ide_intf.ml
deleted file mode 100644
index cbb60b06..00000000
--- a/toplevel/ide_intf.ml
+++ /dev/null
@@ -1,713 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(** Protocol version of this file. This is the date of the last modification. *)
-
-(** WARNING: TO BE UPDATED WHEN MODIFIED! *)
-
-let protocol_version = "20130425~2"
-
-(** * Interface of calls to Coq by CoqIde *)
-
-open Xml_parser
-open Interface
-
-type xml = Xml_parser.xml
-
-(** We use phantom types and GADT to protect ourselves against wild casts *)
-
-type 'a call =
- | Interp of interp_sty
- | Rewind of rewind_sty
- | Goal of goals_sty
- | Evars of evars_sty
- | Hints of hints_sty
- | Status of status_sty
- | Search of search_sty
- | GetOptions of get_options_sty
- | SetOptions of set_options_sty
- | InLoadPath of inloadpath_sty
- | MkCases of mkcases_sty
- | Quit of quit_sty
- | About of about_sty
-
-type unknown
-
-(** The actual calls *)
-
-let interp x : interp_rty call = Interp x
-let rewind x : rewind_rty call = Rewind x
-let goals x : goals_rty call = Goal x
-let evars x : evars_rty call = Evars x
-let hints x : hints_rty call = Hints x
-let status x : status_rty call = Status x
-let get_options x : get_options_rty call = GetOptions x
-let set_options x : set_options_rty call = SetOptions x
-let inloadpath x : inloadpath_rty call = InLoadPath x
-let mkcases x : mkcases_rty call = MkCases x
-let search x : search_rty call = Search x
-let quit x : quit_rty call = Quit x
-
-(** * Coq answers to CoqIde *)
-
-let abstract_eval_call handler (c : 'a call) =
- let mkGood x : 'a value = Good (Obj.magic x) in
- try
- match c with
- | Interp x -> mkGood (handler.interp x)
- | Rewind x -> mkGood (handler.rewind x)
- | Goal x -> mkGood (handler.goals x)
- | Evars x -> mkGood (handler.evars x)
- | Hints x -> mkGood (handler.hints x)
- | Status x -> mkGood (handler.status x)
- | Search x -> mkGood (handler.search x)
- | GetOptions x -> mkGood (handler.get_options x)
- | SetOptions x -> mkGood (handler.set_options x)
- | InLoadPath x -> mkGood (handler.inloadpath x)
- | MkCases x -> mkGood (handler.mkcases x)
- | Quit x -> mkGood (handler.quit x)
- | About x -> mkGood (handler.about x)
- with any ->
- Fail (handler.handle_exn any)
-
-(* To read and typecheck the answers we give a description of the types,
- and a way to statically check that the reified version is in sync *)
-module ReifType : sig
-
- type 'a val_t
-
- val unit_t : unit val_t
- val string_t : string val_t
- val int_t : int val_t
- val bool_t : bool val_t
- val goals_t : goals val_t
- val evar_t : evar val_t
- val state_t : status val_t
- val coq_info_t : coq_info val_t
- val option_state_t : option_state val_t
- val option_t : 'a val_t -> 'a option val_t
- val list_t : 'a val_t -> 'a list val_t
- val coq_object_t : 'a val_t -> 'a coq_object val_t
- val pair_t : 'a val_t -> 'b val_t -> ('a * 'b) val_t
- val union_t : 'a val_t -> 'b val_t -> ('a ,'b) Util.union val_t
-
- type value_type = private
- | Unit | String | Int | Bool | Goals | Evar | State | Option_state | Coq_info
- | Option of value_type
- | List of value_type
- | Coq_object of value_type
- | Pair of value_type * value_type
- | Union of value_type * value_type
-
- val check : 'a val_t -> value_type
-
-end = struct
-
- type value_type =
- | Unit | String | Int | Bool | Goals | Evar | State | Option_state | Coq_info
- | Option of value_type
- | List of value_type
- | Coq_object of value_type
- | Pair of value_type * value_type
- | Union of value_type * value_type
-
- type 'a val_t = value_type
- let check x = x
-
- let unit_t = Unit
- let string_t = String
- let int_t = Int
- let bool_t = Bool
- let goals_t = Goals
- let evar_t = Evar
- let state_t = State
- let coq_info_t = Coq_info
- let option_state_t = Option_state
- let option_t x = Option x
- let list_t x = List x
- let coq_object_t x = Coq_object x
- let pair_t x y = Pair (x, y)
- let union_t x y = Union (x, y)
-
-end
-
-open ReifType
-
-(* For every (call : 'a call), we build the reification of 'a.
- * In OCaml 4 we could use GATDs to do that I guess *)
-let expected_answer_type call : value_type =
- let hint = list_t (pair_t string_t string_t) in
- let hints = pair_t (list_t hint) hint in
- let options = pair_t (list_t string_t) option_state_t in
- let objs = coq_object_t string_t in
- match call with
- | Interp _ -> check (string_t : interp_rty val_t)
- | Rewind _ -> check (int_t : rewind_rty val_t)
- | Goal _ -> check (option_t goals_t : goals_rty val_t)
- | Evars _ -> check (option_t (list_t evar_t) : evars_rty val_t)
- | Hints _ -> check (option_t hints : hints_rty val_t)
- | Status _ -> check (state_t : status_rty val_t)
- | Search _ -> check (list_t objs : search_rty val_t)
- | GetOptions _ -> check (list_t options : get_options_rty val_t)
- | SetOptions _ -> check (unit_t : set_options_rty val_t)
- | InLoadPath _ -> check (bool_t : inloadpath_rty val_t)
- | MkCases _ -> check (list_t (list_t string_t) : mkcases_rty val_t)
- | Quit _ -> check (unit_t : quit_rty val_t)
- | About _ -> check (coq_info_t : about_rty val_t)
-
-(** * XML data marshalling *)
-
-exception Marshal_error
-
-(** Utility functions *)
-
-let massoc x l =
- try List.assoc x l
- with Not_found -> raise Marshal_error
-
-let constructor t c args = Element (t, ["val", c], args)
-
-let do_match constr t mf = match constr with
-| Element (s, attrs, args) ->
- if s = t then
- let c = massoc "val" attrs in
- mf c args
- else raise Marshal_error
-| _ -> raise Marshal_error
-
-let singleton = function
-| [x] -> x
-| _ -> raise Marshal_error
-
-let raw_string = function
-| [] -> ""
-| [PCData s] -> s
-| _ -> raise Marshal_error
-
-let bool_arg tag b = if b then [tag, ""] else []
-
-(** Base types *)
-
-let of_unit () = Element ("unit", [], [])
-
-let to_unit : xml -> unit = function
- | Element ("unit", [], []) -> ()
- | _ -> raise Marshal_error
-
-let of_bool (b : bool) : xml =
- if b then constructor "bool" "true" []
- else constructor "bool" "false" []
-
-let to_bool xml : bool = do_match xml "bool"
- (fun s _ -> match s with
- | "true" -> true
- | "false" -> false
- | _ -> raise Marshal_error)
-
-let of_list (f : 'a -> xml) (l : 'a list) =
- Element ("list", [], List.map f l)
-
-let to_list (f : xml -> 'a) : xml -> 'a list = function
-| Element ("list", [], l) ->
- List.map f l
-| _ -> raise Marshal_error
-
-let of_option (f : 'a -> xml) : 'a option -> xml = function
-| None -> Element ("option", ["val", "none"], [])
-| Some x -> Element ("option", ["val", "some"], [f x])
-
-let to_option (f : xml -> 'a) : xml -> 'a option = function
-| Element ("option", ["val", "none"], []) -> None
-| Element ("option", ["val", "some"], [x]) -> Some (f x)
-| _ -> raise Marshal_error
-
-let of_string (s : string) : xml = Element ("string", [], [PCData s])
-
-let to_string : xml -> string = function
-| Element ("string", [], l) -> raw_string l
-| _ -> raise Marshal_error
-
-let of_int (i : int) : xml = Element ("int", [], [PCData (string_of_int i)])
-
-let to_int : xml -> int = function
-| Element ("int", [], [PCData s]) ->
- (try int_of_string s with Failure _ -> raise Marshal_error)
-| _ -> raise Marshal_error
-
-let of_pair (f : 'a -> xml) (g : 'b -> xml) : 'a * 'b -> xml =
- function (x,y) -> Element ("pair", [], [f x; g y])
-
-let to_pair (f : xml -> 'a) (g : xml -> 'b) : xml -> 'a * 'b = function
-| Element ("pair", [], [x; y]) -> (f x, g y)
-| _ -> raise Marshal_error
-
-let of_union (f : 'a -> xml) (g : 'b -> xml) : ('a,'b) Util.union -> xml =
-function
-| Util.Inl x -> Element ("union", ["val","in_l"], [f x])
-| Util.Inr x -> Element ("union", ["val","in_r"], [g x])
-
-let to_union (f : xml -> 'a) (g : xml -> 'b) : xml -> ('a,'b) Util.union=
-function
-| Element ("union", ["val","in_l"], [x]) -> Util.Inl (f x)
-| Element ("union", ["val","in_r"], [x]) -> Util.Inr (g x)
-| _ -> raise Marshal_error
-
-(** More elaborate types *)
-
-let of_option_value = function
-| IntValue i ->
- constructor "option_value" "intvalue" [of_option of_int i]
-| BoolValue b ->
- constructor "option_value" "boolvalue" [of_bool b]
-| StringValue s ->
- constructor "option_value" "stringvalue" [of_string s]
-
-let to_option_value xml = do_match xml "option_value"
- (fun s args -> match s with
- | "intvalue" -> IntValue (to_option to_int (singleton args))
- | "boolvalue" -> BoolValue (to_bool (singleton args))
- | "stringvalue" -> StringValue (to_string (singleton args))
- | _ -> raise Marshal_error
- )
-
-let of_option_state s =
- Element ("option_state", [], [
- of_bool s.opt_sync;
- of_bool s.opt_depr;
- of_string s.opt_name;
- of_option_value s.opt_value]
- )
-
-let to_option_state = function
-| Element ("option_state", [], [sync; depr; name; value]) ->
- {
- opt_sync = to_bool sync;
- opt_depr = to_bool depr;
- opt_name = to_string name;
- opt_value = to_option_value value;
- }
-| _ -> raise Marshal_error
-
-let of_search_constraint = function
-| Name_Pattern s ->
- constructor "search_constraint" "name_pattern" [of_string s]
-| Type_Pattern s ->
- constructor "search_constraint" "type_pattern" [of_string s]
-| SubType_Pattern s ->
- constructor "search_constraint" "subtype_pattern" [of_string s]
-| In_Module m ->
- constructor "search_constraint" "in_module" [of_list of_string m]
-| Include_Blacklist ->
- constructor "search_constraint" "include_blacklist" []
-
-let to_search_constraint xml = do_match xml "search_constraint"
- (fun s args -> match s with
- | "name_pattern" -> Name_Pattern (to_string (singleton args))
- | "type_pattern" -> Type_Pattern (to_string (singleton args))
- | "subtype_pattern" -> SubType_Pattern (to_string (singleton args))
- | "in_module" -> In_Module (to_list to_string (singleton args))
- | "include_blacklist" -> Include_Blacklist
- | _ -> raise Marshal_error)
-
-let of_coq_object f ans =
- let prefix = of_list of_string ans.coq_object_prefix in
- let qualid = of_list of_string ans.coq_object_qualid in
- let obj = f ans.coq_object_object in
- Element ("coq_object", [], [prefix; qualid; obj])
-
-let to_coq_object f = function
-| Element ("coq_object", [], [prefix; qualid; obj]) ->
- let prefix = to_list to_string prefix in
- let qualid = to_list to_string qualid in
- let obj = f obj in {
- coq_object_prefix = prefix;
- coq_object_qualid = qualid;
- coq_object_object = obj;
- }
-| _ -> raise Marshal_error
-
-let of_value f = function
-| Good x -> Element ("value", ["val", "good"], [f x])
-| Fail (loc, msg) ->
- let loc = match loc with
- | None -> []
- | Some (s, e) -> [("loc_s", string_of_int s); ("loc_e", string_of_int e)]
- in
- Element ("value", ["val", "fail"] @ loc, [PCData msg])
-
-let to_value f = function
-| Element ("value", attrs, l) ->
- let ans = massoc "val" attrs in
- if ans = "good" then Good (f (singleton l))
- else if ans = "fail" then
- let loc =
- try
- let loc_s = int_of_string (List.assoc "loc_s" attrs) in
- let loc_e = int_of_string (List.assoc "loc_e" attrs) in
- Some (loc_s, loc_e)
- with Not_found | Failure _ -> None
- in
- let msg = raw_string l in
- Fail (loc, msg)
- else raise Marshal_error
-| _ -> raise Marshal_error
-
-let of_call = function
-| Interp (id,raw, vrb, cmd) ->
- let flags = (bool_arg "raw" raw) @ (bool_arg "verbose" vrb) in
- Element ("call", ("val", "interp") :: ("id", string_of_int id) :: flags,
- [PCData cmd])
-| Rewind n ->
- Element ("call", ("val", "rewind") :: ["steps", string_of_int n], [])
-| Goal () ->
- Element ("call", ["val", "goal"], [])
-| Evars () ->
- Element ("call", ["val", "evars"], [])
-| Hints () ->
- Element ("call", ["val", "hints"], [])
-| Status () ->
- Element ("call", ["val", "status"], [])
-| Search flags ->
- let args = List.map (of_pair of_search_constraint of_bool) flags in
- Element ("call", ["val", "search"], args)
-| GetOptions () ->
- Element ("call", ["val", "getoptions"], [])
-| SetOptions opts ->
- let args = List.map (of_pair (of_list of_string) of_option_value) opts in
- Element ("call", ["val", "setoptions"], args)
-| InLoadPath file ->
- Element ("call", ["val", "inloadpath"], [PCData file])
-| MkCases ind ->
- Element ("call", ["val", "mkcases"], [PCData ind])
-| Quit () ->
- Element ("call", ["val", "quit"], [])
-| About () ->
- Element ("call", ["val", "about"], [])
-
-let to_call = function
-| Element ("call", attrs, l) ->
- let ans = massoc "val" attrs in
- begin match ans with
- | "interp" -> begin
- try
- let id = List.assoc "id" attrs in
- let raw = List.mem_assoc "raw" attrs in
- let vrb = List.mem_assoc "verbose" attrs in
- Interp (int_of_string id, raw, vrb, raw_string l)
- with Not_found -> raise Marshal_error end
- | "rewind" ->
- let steps = int_of_string (massoc "steps" attrs) in
- Rewind steps
- | "goal" -> Goal ()
- | "evars" -> Evars ()
- | "status" -> Status ()
- | "search" ->
- let args = List.map (to_pair to_search_constraint to_bool) l in
- Search args
- | "getoptions" -> GetOptions ()
- | "setoptions" ->
- let args = List.map (to_pair (to_list to_string) to_option_value) l in
- SetOptions args
- | "inloadpath" -> InLoadPath (raw_string l)
- | "mkcases" -> MkCases (raw_string l)
- | "hints" -> Hints ()
- | "quit" -> Quit ()
- | "about" -> About ()
- | _ -> raise Marshal_error
- end
-| _ -> raise Marshal_error
-
-let of_status s =
- let of_so = of_option of_string in
- let of_sl = of_list of_string in
- Element ("status", [],
- [
- of_sl s.status_path;
- of_so s.status_proofname;
- of_sl s.status_allproofs;
- of_int s.status_statenum;
- of_int s.status_proofnum;
- ]
- )
-
-let to_status = function
-| Element ("status", [], [path; name; prfs; snum; pnum]) ->
- {
- status_path = to_list to_string path;
- status_proofname = to_option to_string name;
- status_allproofs = to_list to_string prfs;
- status_statenum = to_int snum;
- status_proofnum = to_int pnum;
- }
-| _ -> raise Marshal_error
-
-let of_evar s =
- Element ("evar", [], [PCData s.evar_info])
-
-let to_evar = function
-| Element ("evar", [], data) -> { evar_info = raw_string data; }
-| _ -> raise Marshal_error
-
-let of_goal g =
- let hyp = of_list of_string g.goal_hyp in
- let ccl = of_string g.goal_ccl in
- let id = of_string g.goal_id in
- Element ("goal", [], [id; hyp; ccl])
-
-let to_goal = function
-| Element ("goal", [], [id; hyp; ccl]) ->
- let hyp = to_list to_string hyp in
- let ccl = to_string ccl in
- let id = to_string id in
- { goal_hyp = hyp; goal_ccl = ccl; goal_id = id; }
-| _ -> raise Marshal_error
-
-let of_goals g =
- let of_glist = of_list of_goal in
- let fg = of_list of_goal g.fg_goals in
- let bg = of_list (of_pair of_glist of_glist) g.bg_goals in
- Element ("goals", [], [fg; bg])
-
-let to_goals = function
-| Element ("goals", [], [fg; bg]) ->
- let to_glist = to_list to_goal in
- let fg = to_list to_goal fg in
- let bg = to_list (to_pair to_glist to_glist) bg in
- { fg_goals = fg; bg_goals = bg; }
-| _ -> raise Marshal_error
-
-let of_coq_info info =
- let version = of_string info.coqtop_version in
- let protocol = of_string info.protocol_version in
- let release = of_string info.release_date in
- let compile = of_string info.compile_date in
- Element ("coq_info", [], [version; protocol; release; compile])
-
-let to_coq_info = function
-| Element ("coq_info", [], [version; protocol; release; compile]) ->
- {
- coqtop_version = to_string version;
- protocol_version = to_string protocol;
- release_date = to_string release;
- compile_date = to_string compile;
- }
-| _ -> raise Marshal_error
-
-let of_message_level = function
-| Debug s -> constructor "message_level" "debug" [PCData s]
-| Info -> constructor "message_level" "info" []
-| Notice -> constructor "message_level" "notice" []
-| Warning -> constructor "message_level" "warning" []
-| Error -> constructor "message_level" "error" []
-
-let to_message_level xml = do_match xml "message_level"
- (fun s args -> match s with
- | "debug" -> Debug (raw_string args)
- | "info" -> Info
- | "notice" -> Notice
- | "warning" -> Warning
- | "error" -> Error
- | _ -> raise Marshal_error)
-
-let of_message msg =
- let lvl = of_message_level msg.message_level in
- let content = of_string msg.message_content in
- Element ("message", [], [lvl; content])
-
-let to_message xml = match xml with
-| Element ("message", [], [lvl; content]) ->
- { message_level = to_message_level lvl; message_content = to_string content }
-| _ -> raise Marshal_error
-
-let is_message = function
-| Element ("message", _, _) -> true
-| _ -> false
-
-let of_loc loc =
- let start, stop = loc in
- Element ("loc",[("start",string_of_int start);("stop",string_of_int stop)],[])
-
-let to_loc xml = match xml with
-| Element ("loc", l,[]) ->
- (try
- let start = List.assoc "start" l in
- let stop = List.assoc "stop" l in
- (int_of_string start, int_of_string stop)
- with Not_found | Invalid_argument _ -> raise Marshal_error)
-| _ -> raise Marshal_error
-
-let to_feedback_content xml = do_match xml "feedback_content"
- (fun s args -> match s with
- | "addedaxiom" -> AddedAxiom
- | "processed" -> Processed
- | "globref" ->
- (match args with
- | [loc; filepath; modpath; ident; ty] ->
- GlobRef(to_loc loc, to_string filepath, to_string modpath,
- to_string ident, to_string ty)
- | _ -> raise Marshal_error)
- | _ -> raise Marshal_error)
-
-let of_feedback_content = function
-| AddedAxiom -> constructor "feedback_content" "addedaxiom" []
-| Processed -> constructor "feedback_content" "processed" []
-| GlobRef(loc, filepath, modpath, ident, ty) ->
- constructor "feedback_content" "globref" [
- of_loc loc;
- of_string filepath;
- of_string modpath;
- of_string ident;
- of_string ty
- ]
-
-let of_feedback msg =
- let content = of_feedback_content msg.content in
- Element ("feedback", ["id",string_of_int msg.edit_id], [content])
-
-let to_feedback xml = match xml with
-| Element ("feedback", ["id",id], [content]) ->
- { edit_id = int_of_string id;
- content = to_feedback_content content }
-| _ -> raise Marshal_error
-
-let is_feedback = function
-| Element ("feedback", _, _) -> true
-| _ -> false
-
-(** Conversions between ['a value] and xml answers
-
- When decoding an xml answer, we dynamically check that it is compatible
- with the original call. For that we now rely on the fact that all
- sub-fonctions [to_xxx : xml -> xxx] check that the current xml element
- is "xxx", and raise [Marshal_error] if anything goes wrong. *)
-
-let of_answer (q : 'a call) (r : 'a value) : xml =
- let rec convert ty : 'a -> xml = match ty with
- | Unit -> Obj.magic of_unit
- | Bool -> Obj.magic of_bool
- | String -> Obj.magic of_string
- | Int -> Obj.magic of_int
- | State -> Obj.magic of_status
- | Option_state -> Obj.magic of_option_state
- | Coq_info -> Obj.magic of_coq_info
- | Goals -> Obj.magic of_goals
- | Evar -> Obj.magic of_evar
- | List t -> Obj.magic (of_list (convert t))
- | Option t -> Obj.magic (of_option (convert t))
- | Coq_object t -> Obj.magic (of_coq_object (convert t))
- | Pair (t1,t2) -> Obj.magic (of_pair (convert t1) (convert t2))
- | Union (t1,t2) -> Obj.magic (of_union (convert t1) (convert t2))
- in
- of_value (convert (expected_answer_type q)) r
-
-let to_answer xml (c : 'a call) : 'a value =
- let rec convert ty : xml -> 'a = match ty with
- | Unit -> Obj.magic to_unit
- | Bool -> Obj.magic to_bool
- | String -> Obj.magic to_string
- | Int -> Obj.magic to_int
- | State -> Obj.magic to_status
- | Option_state -> Obj.magic to_option_state
- | Coq_info -> Obj.magic to_coq_info
- | Goals -> Obj.magic to_goals
- | Evar -> Obj.magic to_evar
- | List t -> Obj.magic (to_list (convert t))
- | Option t -> Obj.magic (to_option (convert t))
- | Coq_object t -> Obj.magic (to_coq_object (convert t))
- | Pair (t1,t2) -> Obj.magic (to_pair (convert t1) (convert t2))
- | Union (t1,t2) -> Obj.magic (to_union (convert t1) (convert t2))
- in
- to_value (convert (expected_answer_type c)) xml
-
-(** * Debug printing *)
-
-let pr_unit () = ""
-let pr_string s = Printf.sprintf "%S" s
-let pr_int i = string_of_int i
-let pr_bool b = Printf.sprintf "%B" b
-let pr_goal (g : goals) =
- if g.fg_goals = [] then
- if g.bg_goals = [] then "Proof completed."
- else
- let rec pr_focus _ = function
- | [] -> assert false
- | [lg, rg] -> Printf.sprintf "%i" (List.length lg + List.length rg)
- | (lg, rg) :: l ->
- Printf.sprintf "%i:%a" (List.length lg + List.length rg) pr_focus l in
- Printf.sprintf "Still focussed: [%a]." pr_focus g.bg_goals
- else
- let pr_menu s = s in
- let pr_goal { goal_hyp = hyps; goal_ccl = goal } =
- "[" ^ String.concat "; " (List.map pr_menu hyps) ^ " |- " ^
- pr_menu goal ^ "]" in
- String.concat " " (List.map pr_goal g.fg_goals)
-let pr_evar (e : evar) = "[" ^ e.evar_info ^ "]"
-let pr_status (s : status) =
- let path =
- let l = String.concat "." s.status_path in
- "path=" ^ l ^ ";" in
- let name = match s.status_proofname with
- | None -> "no proof;"
- | Some n -> "proof = " ^ n ^ ";" in
- "Status: " ^ path ^ name
-let pr_coq_info (i : coq_info) = "FIXME"
-let pr_option_value = function
- | IntValue None -> "none"
- | IntValue (Some i) -> string_of_int i
- | StringValue s -> s
- | BoolValue b -> if b then "true" else "false"
-let pr_option_state (s : option_state) =
- Printf.sprintf "sync := %b; depr := %b; name := %s; value := %s\n"
- s.opt_sync s.opt_depr s.opt_name (pr_option_value s.opt_value)
-let pr_list pr l = "["^String.concat ";" (List.map pr l)^"]"
-let pr_option pr = function None -> "None" | Some x -> "Some("^pr x^")"
-let pr_coq_object (o : 'a coq_object) = "FIXME"
-let pr_pair pr1 pr2 (a,b) = "("^pr1 a^","^pr2 b^")"
-let pr_union pr1 pr2 = function Util.Inl x -> pr1 x | Util.Inr x -> pr2 x
-
-let pr_call = function
- | Interp (id,r,b,s) ->
- let raw = if r then "RAW" else "" in
- let verb = if b then "" else "SILENT" in
- "INTERP"^raw^verb^" "^string_of_int id^" ["^s^"]"
- | Rewind i -> "REWIND "^(string_of_int i)
- | Goal _ -> "GOALS"
- | Evars _ -> "EVARS"
- | Hints _ -> "HINTS"
- | Status _ -> "STATUS"
- | Search _ -> "SEARCH"
- | GetOptions _ -> "GETOPTIONS"
- | SetOptions l -> "SETOPTIONS" ^ " [" ^
- String.concat ";"
- (List.map (pr_pair (pr_list pr_string) pr_option_value) l) ^ "]"
- | InLoadPath s -> "INLOADPATH "^s
- | MkCases s -> "MKCASES "^s
- | Quit _ -> "QUIT"
- | About _ -> "ABOUT"
-let pr_value_gen pr = function
- | Good v -> "GOOD " ^ pr v
- | Fail (_,str) -> "FAIL ["^str^"]"
-let pr_value v = pr_value_gen (fun _ -> "FIXME") v
-let pr_full_value call value =
- let rec pr = function
- | Unit -> Obj.magic pr_unit
- | Bool -> Obj.magic pr_bool
- | String -> Obj.magic pr_string
- | Int -> Obj.magic pr_int
- | State -> Obj.magic pr_status
- | Option_state -> Obj.magic pr_option_state
- | Coq_info -> Obj.magic pr_coq_info
- | Goals -> Obj.magic pr_goal
- | Evar -> Obj.magic pr_evar
- | List t -> Obj.magic (pr_list (pr t))
- | Option t -> Obj.magic (pr_option (pr t))
- | Coq_object t -> Obj.magic pr_coq_object
- | Pair (t1,t2) -> Obj.magic (pr_pair (pr t1) (pr t2))
- | Union (t1,t2) -> Obj.magic (pr_union (pr t1) (pr t2))
- in
- pr_value_gen (pr (expected_answer_type call)) value
diff --git a/toplevel/ide_intf.mli b/toplevel/ide_intf.mli
deleted file mode 100644
index ab8ecc8e..00000000
--- a/toplevel/ide_intf.mli
+++ /dev/null
@@ -1,62 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(** * Applicative part of the interface of CoqIde calls to Coq *)
-
-open Interface
-
-type xml = Xml_parser.xml
-
-type 'a call
-
-type unknown
-
-val interp : interp_sty -> interp_rty call
-val rewind : rewind_sty -> rewind_rty call
-val goals : goals_sty -> goals_rty call
-val hints : hints_sty -> hints_rty call
-val status : status_sty -> status_rty call
-val inloadpath : inloadpath_sty -> inloadpath_rty call
-val mkcases : mkcases_sty -> mkcases_rty call
-val evars : evars_sty -> evars_rty call
-val search : search_sty -> search_rty call
-val get_options : get_options_sty -> get_options_rty call
-val set_options : set_options_sty -> set_options_rty call
-val quit : quit_sty -> quit_rty call
-
-val abstract_eval_call : handler -> 'a call -> 'a value
-
-(** * Protocol version *)
-
-val protocol_version : string
-
-(** * XML data marshalling *)
-
-exception Marshal_error
-
-val of_call : 'a call -> xml
-val to_call : xml -> unknown call
-
-val of_message : message -> xml
-val to_message : xml -> message
-val is_message : xml -> bool
-
-val of_value : ('a -> xml) -> 'a value -> xml
-
-val of_feedback : feedback -> xml
-val to_feedback : xml -> feedback
-val is_feedback : xml -> bool
-
-val of_answer : 'a call -> 'a value -> xml
-val to_answer : xml -> 'a call -> 'a value
-
-(** * Debug printing *)
-
-val pr_call : 'a call -> string
-val pr_value : 'a value -> string
-val pr_full_value : 'a call -> 'a value -> string
diff --git a/toplevel/ide_slave.ml b/toplevel/ide_slave.ml
deleted file mode 100644
index 9520a990..00000000
--- a/toplevel/ide_slave.ml
+++ /dev/null
@@ -1,466 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-open Vernacexpr
-open Names
-open Compat
-open Util
-open Pp
-open Printer
-open Namegen
-
-(** Ide_slave : an implementation of [Interface], i.e. mainly an interp
- function and a rewind function. This specialized loop is triggered
- when the -ideslave option is passed to Coqtop. Currently CoqIDE is
- the only one using this mode, but we try here to be as generic as
- possible, so this may change in the future... *)
-
-(** Signal handling: we postpone ^C during input and output phases,
- but make it directly raise a Sys.Break during evaluation of the request. *)
-
-let catch_break = ref false
-
-let init_signal_handler () =
- let f _ = if !catch_break then raise Sys.Break else Util.interrupt := true in
- Sys.set_signal Sys.sigint (Sys.Signal_handle f)
-
-
-(** Redirection of standard output to a printable buffer *)
-
-let orig_stdout = ref stdout
-
-let init_stdout,read_stdout =
- let out_buff = Buffer.create 100 in
- let out_ft = Format.formatter_of_buffer out_buff in
- let deep_out_ft = Format.formatter_of_buffer out_buff in
- let _ = Pp_control.set_gp deep_out_ft Pp_control.deep_gp in
- (fun () ->
- flush_all ();
- orig_stdout := Unix.out_channel_of_descr (Unix.dup Unix.stdout);
- Unix.dup2 Unix.stderr Unix.stdout;
- Pp_control.std_ft := out_ft;
- Pp_control.err_ft := out_ft;
- Pp_control.deep_ft := deep_out_ft;
- set_binary_mode_out !orig_stdout true;
- set_binary_mode_in stdin true;
- ),
- (fun () -> Format.pp_print_flush out_ft ();
- let r = Buffer.contents out_buff in
- Buffer.clear out_buff; r)
-
-let pr_debug s =
- if !Flags.debug then Printf.eprintf "[pid %d] %s\n%!" (Unix.getpid ()) s
-
-(** Categories of commands *)
-
-let coqide_known_option table = List.mem table [
- ["Printing";"Implicit"];
- ["Printing";"Coercions"];
- ["Printing";"Matching"];
- ["Printing";"Synth"];
- ["Printing";"Notations"];
- ["Printing";"All"];
- ["Printing";"Records"];
- ["Printing";"Existential";"Instances"];
- ["Printing";"Universes"]]
-
-let is_known_option cmd = match cmd with
- | VernacSetOption (_,o,BoolValue true)
- | VernacUnsetOption (_,o) -> coqide_known_option o
- | _ -> false
-
-let is_debug cmd = match cmd with
- | VernacSetOption (_,["Ltac";"Debug"], _) -> true
- | _ -> false
-
-let is_query cmd = match cmd with
- | VernacChdir None
- | VernacMemOption _
- | VernacPrintOption _
- | VernacCheckMayEval _
- | VernacGlobalCheck _
- | VernacPrint _
- | VernacSearch _
- | VernacLocate _ -> true
- | _ -> false
-
-let is_undo cmd = match cmd with
- | VernacUndo _ | VernacUndoTo _ -> true
- | _ -> false
-
-(** Check whether a command is forbidden by CoqIDE *)
-
-let coqide_cmd_checks (loc,ast) =
- let user_error s =
- raise (Loc.Exc_located (loc, Util.UserError ("CoqIde", str s)))
- in
- if is_debug ast then
- user_error "Debug mode not available within CoqIDE";
- if is_known_option ast then
- user_error "Use CoqIDE display menu instead";
- if is_navigation_vernac ast then
- user_error "Use CoqIDE navigation instead";
- if is_undo ast then
- msgerrnl (str "Warning: rather use CoqIDE navigation instead");
- if is_query ast then
- msgerrnl (str "Warning: query commands should not be inserted in scripts")
-
-(** Interpretation (cf. [Ide_intf.interp]) *)
-
-let interp (id,raw,verbosely,s) =
- let pa = Pcoq.Gram.parsable (Stream.of_string s) in
- let loc_ast = Vernac.parse_sentence (pa,None) in
- if not raw then coqide_cmd_checks loc_ast;
- Flags.make_silent (not verbosely);
- Vernac.eval_expr ~preserving:raw loc_ast;
- Flags.make_silent true;
- read_stdout ()
-
-(** Goal display *)
-
-let hyp_next_tac sigma env (id,_,ast) =
- let id_s = Names.string_of_id id in
- let type_s = string_of_ppcmds (pr_ltype_env env ast) in
- [
- ("clear "^id_s),("clear "^id_s^".");
- ("apply "^id_s),("apply "^id_s^".");
- ("exact "^id_s),("exact "^id_s^".");
- ("generalize "^id_s),("generalize "^id_s^".");
- ("absurd <"^id_s^">"),("absurd "^type_s^".")
- ] @ [
- ("discriminate "^id_s),("discriminate "^id_s^".");
- ("injection "^id_s),("injection "^id_s^".")
- ] @ [
- ("rewrite "^id_s),("rewrite "^id_s^".");
- ("rewrite <- "^id_s),("rewrite <- "^id_s^".")
- ] @ [
- ("elim "^id_s), ("elim "^id_s^".");
- ("inversion "^id_s), ("inversion "^id_s^".");
- ("inversion clear "^id_s), ("inversion_clear "^id_s^".")
- ]
-
-let concl_next_tac sigma concl =
- let expand s = (s,s^".") in
- List.map expand ([
- "intro";
- "intros";
- "intuition"
- ] @ [
- "reflexivity";
- "discriminate";
- "symmetry"
- ] @ [
- "assumption";
- "omega";
- "ring";
- "auto";
- "eauto";
- "tauto";
- "trivial";
- "decide equality";
- "simpl";
- "subst";
- "red";
- "split";
- "left";
- "right"
- ])
-
-let process_goal sigma g =
- let env = Goal.V82.env sigma g in
- let id = Goal.uid g in
- let ccl =
- let norm_constr = Reductionops.nf_evar sigma (Goal.V82.concl sigma g) in
- string_of_ppcmds (pr_goal_concl_style_env env norm_constr) in
- let process_hyp h_env d acc =
- let d = Term.map_named_declaration (Reductionops.nf_evar sigma) d in
- (string_of_ppcmds (pr_var_decl h_env d)) :: acc in
- let hyps =
- List.rev (Environ.fold_named_context process_hyp env ~init: []) in
- { Interface.goal_hyp = hyps; Interface.goal_ccl = ccl; Interface.goal_id = id; }
-
-let goals () =
- try
- let pfts = Proof_global.give_me_the_proof () in
- let (goals, zipper, sigma) = Proof.proof pfts in
- let fg = List.map (process_goal sigma) goals in
- let map_zip (lg, rg) =
- let lg = List.map (process_goal sigma) lg in
- let rg = List.map (process_goal sigma) rg in
- (lg, rg)
- in
- let bg = List.map map_zip zipper in
- Some { Interface.fg_goals = fg; Interface.bg_goals = bg; }
- with Proof_global.NoCurrentProof -> None
-
-let evars () =
- try
- let pfts = Proof_global.give_me_the_proof () in
- let { Evd.it = all_goals ; sigma = sigma } = Proof.V82.subgoals pfts in
- let exl = Evarutil.non_instantiated sigma in
- let map_evar ev = { Interface.evar_info = string_of_ppcmds (pr_evar ev); } in
- let el = List.map map_evar exl in
- Some el
- with Proof_global.NoCurrentProof -> None
-
-let hints () =
- try
- let pfts = Proof_global.give_me_the_proof () in
- let { Evd.it = all_goals ; sigma = sigma } = Proof.V82.subgoals pfts in
- match all_goals with
- | [] -> None
- | g :: _ ->
- let env = Goal.V82.env sigma g in
- let hint_goal = concl_next_tac sigma g in
- let get_hint_hyp env d accu = hyp_next_tac sigma env d :: accu in
- let hint_hyps = List.rev (Environ.fold_named_context get_hint_hyp env ~init: []) in
- Some (hint_hyps, hint_goal)
- with Proof_global.NoCurrentProof -> None
-
-(** Other API calls *)
-
-let inloadpath dir =
- Library.is_in_load_paths (System.physical_path_of_string dir)
-
-let status () =
- (** We remove the initial part of the current [dir_path]
- (usually Top in an interactive session, cf "coqtop -top"),
- and display the other parts (opened sections and modules) *)
- let path =
- let l = Names.repr_dirpath (Lib.cwd ()) in
- List.rev_map Names.string_of_id l
- in
- let proof =
- try Some (Names.string_of_id (Proof_global.get_current_proof_name ()))
- with Proof_global.NoCurrentProof -> None
- in
- let allproofs =
- let l = Proof_global.get_all_proof_names () in
- List.map Names.string_of_id l
- in
- {
- Interface.status_path = path;
- Interface.status_proofname = proof;
- Interface.status_allproofs = allproofs;
- Interface.status_statenum = Lib.current_command_label ();
- Interface.status_proofnum = Pfedit.current_proof_depth ();
- }
-
-(** This should be elsewhere... *)
-let search flags =
- let env = Global.env () in
- let rec extract_flags name tpe subtpe mods blacklist = function
- | [] -> (name, tpe, subtpe, mods, blacklist)
- | (Interface.Name_Pattern s, b) :: l ->
- let regexp =
- try Str.regexp s
- with e when Errors.noncritical e ->
- Util.error ("Invalid regexp: " ^ s)
- in
- extract_flags ((regexp, b) :: name) tpe subtpe mods blacklist l
- | (Interface.Type_Pattern s, b) :: l ->
- let constr = Pcoq.parse_string Pcoq.Constr.lconstr_pattern s in
- let (_, pat) = Constrintern.intern_constr_pattern Evd.empty env constr in
- extract_flags name ((pat, b) :: tpe) subtpe mods blacklist l
- | (Interface.SubType_Pattern s, b) :: l ->
- let constr = Pcoq.parse_string Pcoq.Constr.lconstr_pattern s in
- let (_, pat) = Constrintern.intern_constr_pattern Evd.empty env constr in
- extract_flags name tpe ((pat, b) :: subtpe) mods blacklist l
- | (Interface.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 ->
- Util.error ("Module " ^ path ^ " not found.")
- in
- extract_flags name tpe subtpe ((id, b) :: mods) blacklist l
- | (Interface.Include_Blacklist, b) :: l ->
- extract_flags name tpe subtpe mods b l
- in
- let (name, tpe, subtpe, mods, blacklist) =
- extract_flags [] [] [] [] false flags
- in
- let filter_function ref env constr =
- let id = Names.string_of_id (Nametab.basename_of_global ref) in
- let path = Libnames.dirpath (Nametab.path_of_global ref) in
- let toggle x b = if x then b else not b in
- let match_name (regexp, flag) =
- toggle (Str.string_match regexp id 0) flag
- in
- let match_type (pat, flag) =
- toggle (Matching.is_matching pat constr) flag
- in
- let match_subtype (pat, flag) =
- toggle (Matching.is_matching_appsubterm ~closed:false pat constr) flag
- in
- let match_module (mdl, flag) =
- toggle (Libnames.is_dirpath_prefix_of mdl path) flag
- in
- let in_blacklist =
- blacklist || (Search.filter_blacklist 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
- in
- let ans = ref [] in
- let print_function ref env constr =
- let fullpath = repr_dirpath (Nametab.dirpath_of_global ref) in
- let qualid = Nametab.shortest_qualid_of_global Idset.empty ref in
- let (shortpath, basename) = Libnames.repr_qualid qualid in
- let shortpath = repr_dirpath shortpath in
- (* [shortpath] is a suffix of [fullpath] and we're looking for the missing
- prefix *)
- let rec prefix full short accu = match full, short with
- | _, [] ->
- let full = List.rev_map string_of_id full in
- (full, accu)
- | _ :: full, m :: short ->
- prefix full short (string_of_id m :: accu)
- | _ -> assert false
- in
- let (prefix, qualid) = prefix fullpath shortpath [string_of_id basename] in
- let answer = {
- Interface.coq_object_prefix = prefix;
- Interface.coq_object_qualid = qualid;
- Interface.coq_object_object = string_of_ppcmds (pr_lconstr_env env constr);
- } in
- ans := answer :: !ans;
- in
- let () = Search.gen_filtered_search filter_function print_function in
- !ans
-
-let get_options () =
- let table = Goptions.get_tables () in
- let fold key state accu = (key, state) :: accu in
- Goptions.OptionMap.fold fold table []
-
-let set_options options =
- let iter (name, value) = match value with
- | BoolValue b -> Goptions.set_bool_option_value name b
- | IntValue i -> Goptions.set_int_option_value name i
- | StringValue s -> Goptions.set_string_option_value name s
- in
- List.iter iter options
-
-let about () = {
- Interface.coqtop_version = Coq_config.version;
- Interface.protocol_version = Ide_intf.protocol_version;
- Interface.release_date = Coq_config.date;
- Interface.compile_date = Coq_config.compile_date;
-}
-
-(** Grouping all call handlers together + error handling *)
-
-exception Quit
-
-let eval_call c =
- let rec handle_exn e =
- catch_break := false;
- let pr_exn e = (read_stdout ())^("\n"^(string_of_ppcmds (Errors.print e))) in
- match e with
- | Quit ->
- (* Here we do send an acknowledgement message to prove everything went
- OK. *)
- let dummy = Interface.Good () in
- let xml_answer = Ide_intf.of_answer (Ide_intf.quit ()) dummy in
- let () = Xml_utils.print_xml !orig_stdout xml_answer in
- let () = flush !orig_stdout in
- let () = pr_debug "Exiting gracefully." in
- exit 0
- | Vernacexpr.Drop -> None, "Drop is not allowed by coqide!"
- | Vernacexpr.Quit -> None, "Quit is not allowed by coqide!"
- | Vernac.DuringCommandInterp (_,inner) -> handle_exn inner
- | Error_in_file (_,_,inner) -> None, pr_exn inner
- | Loc.Exc_located (loc, inner) when loc = dummy_loc -> None, pr_exn inner
- | Loc.Exc_located (loc, inner) -> Some (Util.unloc loc), pr_exn inner
- | e -> None, pr_exn e
- in
- let interruptible f x =
- catch_break := true;
- Util.check_for_interrupt ();
- let r = f x in
- catch_break := false;
- r
- in
- let handler = {
- Interface.interp = interruptible interp;
- Interface.rewind = interruptible Backtrack.back;
- Interface.goals = interruptible goals;
- Interface.evars = interruptible evars;
- Interface.hints = interruptible hints;
- Interface.status = interruptible status;
- Interface.search = interruptible search;
- Interface.inloadpath = interruptible inloadpath;
- Interface.get_options = interruptible get_options;
- Interface.set_options = interruptible set_options;
- Interface.mkcases = interruptible Vernacentries.make_cases;
- Interface.quit = (fun () -> raise Quit);
- Interface.about = interruptible about;
- Interface.handle_exn = handle_exn; }
- in
- (* If the messages of last command are still there, we remove them *)
- ignore (read_stdout ());
- Ide_intf.abstract_eval_call handler c
-
-
-(** The main loop *)
-
-(** Exceptions during eval_call should be converted into [Interface.Fail]
- messages by [handle_exn] above. Otherwise, we die badly, after having
- tried to send a last message to the ide: trying to recover from errors
- with the current protocol would most probably bring desynchronisation
- between coqtop and ide. With marshalling, reading an answer to
- a different request could hang the ide... *)
-
-let fail err =
- Ide_intf.of_value (fun _ -> assert false) (Interface.Fail (None, err))
-
-let loop () =
- let p = Xml_parser.make () in
- let () = Xml_parser.check_eof p false in
- init_signal_handler ();
- catch_break := false;
- (* We'll handle goal fetching and display in our own way *)
- Vernacentries.enable_goal_printing := false;
- Vernacentries.qed_display_script := false;
- try
- while true do
- let xml_answer =
- try
- let xml_query = Xml_parser.parse p (Xml_parser.SChannel stdin) in
- let q = Ide_intf.to_call xml_query in
- let () = pr_debug ("<-- " ^ Ide_intf.pr_call q) in
- let r = eval_call q in
- let () = pr_debug ("--> " ^ Ide_intf.pr_full_value q r) in
- Ide_intf.of_answer q r
- with
- | Xml_parser.Error (Xml_parser.Empty, _) ->
- pr_debug ("End of input, exiting");
- exit 0
- | Xml_parser.Error (err, loc) ->
- let msg = "Syntax error in query: " ^ Xml_parser.error_msg err in
- fail msg
- | Ide_intf.Marshal_error ->
- fail "Incorrect query."
- in
- Xml_utils.print_xml !orig_stdout xml_answer;
- flush !orig_stdout
- done
- with any ->
- let msg = Printexc.to_string any in
- let r = "Fatal exception in coqtop:\n" ^ msg in
- pr_debug ("==> " ^ r);
- (try
- Xml_utils.print_xml !orig_stdout (fail r);
- flush !orig_stdout
- with any -> ());
- exit 1
diff --git a/toplevel/ide_slave.mli b/toplevel/ide_slave.mli
deleted file mode 100644
index 8b0ad168..00000000
--- a/toplevel/ide_slave.mli
+++ /dev/null
@@ -1,17 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(** [Ide_slave] : an implementation of [Ide_intf], i.e. mainly an interp
- function and a rewind function. This specialized loop is triggered
- when the -ideslave option is passed to Coqtop. Currently CoqIDE is
- the only one using this mode, but we try here to be as generic as
- possible, so this may change in the future... *)
-
-val init_stdout : unit -> unit
-
-val loop : unit -> unit
diff --git a/toplevel/ind_tables.ml b/toplevel/ind_tables.ml
index 932bdfe7..138e5189 100644
--- a/toplevel/ind_tables.ml
+++ b/toplevel/ind_tables.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -18,6 +18,7 @@ open Libobject
open Nameops
open Declarations
open Term
+open Errors
open Util
open Declare
open Entries
@@ -26,23 +27,28 @@ open Decl_kinds
(**********************************************************************)
(* Registering schemes in the environment *)
-type mutual_scheme_object_function = mutual_inductive -> constr array
-type individual_scheme_object_function = inductive -> constr
+
+type mutual_scheme_object_function =
+ mutual_inductive -> constr array Evd.in_evar_universe_context * Declareops.side_effects
+type individual_scheme_object_function =
+ inductive -> constr Evd.in_evar_universe_context * Declareops.side_effects
type 'a scheme_kind = string
-let scheme_map = ref Indmap.empty
+let scheme_map = Summary.ref Indmap.empty ~name:"Schemes"
+
+let pr_scheme_kind = Pp.str
let cache_one_scheme kind (ind,const) =
- let map = try Indmap.find ind !scheme_map with Not_found -> Stringmap.empty in
- scheme_map := Indmap.add ind (Stringmap.add kind const map) !scheme_map
+ let map = try Indmap.find ind !scheme_map with Not_found -> String.Map.empty in
+ scheme_map := Indmap.add ind (String.Map.add kind const map) !scheme_map
let cache_scheme (_,(kind,l)) =
Array.iter (cache_one_scheme kind) l
-let subst_one_scheme subst ((mind,i),const) =
+let subst_one_scheme subst (ind,const) =
(* Remark: const is a def: the result of substitution is a constant *)
- ((subst_ind subst mind,i),fst (subst_con subst const))
+ (subst_ind subst ind,subst_constant subst const)
let subst_scheme (subst,(kind,l)) =
(kind,Array.map (subst_one_scheme subst) l)
@@ -60,36 +66,24 @@ let inScheme : string * (inductive * constant) array -> obj =
discharge_function = discharge_scheme}
(**********************************************************************)
-(* Saving/restoring the table of scheme *)
-
-let freeze_schemes () = !scheme_map
-let unfreeze_schemes sch = scheme_map := sch
-let init_schemes () = scheme_map := Indmap.empty
-
-let _ =
- Summary.declare_summary "Schemes"
- { Summary.freeze_function = freeze_schemes;
- Summary.unfreeze_function = unfreeze_schemes;
- Summary.init_function = init_schemes }
-
-(**********************************************************************)
(* The table of scheme building functions *)
type individual
type mutual
type scheme_object_function =
- | MutualSchemeFunction of (mutual_inductive -> constr array)
- | IndividualSchemeFunction of (inductive -> constr)
+ | MutualSchemeFunction of mutual_scheme_object_function
+ | IndividualSchemeFunction of individual_scheme_object_function
let scheme_object_table =
(Hashtbl.create 17 : (string, string * scheme_object_function) Hashtbl.t)
let declare_scheme_object s aux f =
- (try check_ident ("ind"^s)
- with e when Errors.noncritical e ->
- error ("Illegal induction scheme suffix: "^s));
- let key = if aux = "" then s else aux in
+ let () =
+ if not (Id.is_valid ("ind" ^ s)) then
+ error ("Illegal induction scheme suffix: " ^ s)
+ in
+ let key = if String.is_empty aux then s else aux in
try
let _ = Hashtbl.find scheme_object_table key in
(* let aux_msg = if aux="" then "" else " (with key "^aux^")" in*)
@@ -110,6 +104,8 @@ let declare_individual_scheme_object s ?(aux="") f =
let declare_scheme kind indcl =
Lib.add_anonymous_leaf (inScheme (kind,indcl))
+let () = Declare.set_declare_scheme declare_scheme
+
let is_visible_name id =
try ignore (Nametab.locate (Libnames.qualid_of_ident id)); true
with Not_found -> false
@@ -120,30 +116,39 @@ let compute_name internal id =
| KernelSilent ->
Namegen.next_ident_away_from (add_prefix "internal_" id) is_visible_name
-let define internal id c =
+let define internal id c p univs =
let fd = declare_constant ~internal in
let id = compute_name internal id in
- let kn = fd id
- (DefinitionEntry
- { const_entry_body = c;
- const_entry_secctx = None;
- const_entry_type = None;
- const_entry_opaque = false },
- Decl_kinds.IsDefinition Scheme) in
- (match internal with
- | KernelSilent -> ()
- | _-> definition_message id);
+ let ctx = Evd.normalize_evar_universe_context univs in
+ let c = Vars.subst_univs_fn_constr
+ (Universes.make_opt_subst (Evd.evar_universe_context_subst ctx)) c in
+ let entry = {
+ const_entry_body = Future.from_val ((c,Univ.ContextSet.empty), Declareops.no_seff);
+ const_entry_secctx = None;
+ const_entry_type = None;
+ const_entry_polymorphic = p;
+ const_entry_universes = Evd.evar_context_universe_context ctx;
+ const_entry_opaque = false;
+ const_entry_inline_code = false;
+ const_entry_feedback = None;
+ } in
+ let kn = fd id (DefinitionEntry entry, Decl_kinds.IsDefinition Scheme) in
+ let () = match internal with
+ | KernelSilent -> ()
+ | _-> definition_message id
+ in
kn
let define_individual_scheme_base kind suff f internal idopt (mind,i as ind) =
- let c = f ind in
+ let (c, ctx), eff = f ind in
let mib = Global.lookup_mind mind in
let id = match idopt with
| Some id -> id
| None -> add_suffix mib.mind_packets.(i).mind_typename suff in
- let const = define internal id c in
+ let const = define internal id c mib.mind_polymorphic ctx in
declare_scheme kind [|ind,const|];
- const
+ const, Declareops.cons_side_effects
+ (Safe_typing.sideff_of_scheme kind (Global.safe_env()) [ind,const]) eff
let define_individual_scheme kind internal names (mind,i as ind) =
match Hashtbl.find scheme_object_table kind with
@@ -152,14 +157,21 @@ let define_individual_scheme kind internal names (mind,i as ind) =
define_individual_scheme_base kind s f internal names ind
let define_mutual_scheme_base kind suff f internal names mind =
- let cl = f mind in
+ let (cl, ctx), eff = f mind in
let mib = Global.lookup_mind mind in
let ids = Array.init (Array.length mib.mind_packets) (fun i ->
- try List.assoc i names
+ try Int.List.assoc i names
with Not_found -> add_suffix mib.mind_packets.(i).mind_typename suff) in
- let consts = array_map2 (define internal) ids cl in
- declare_scheme kind (Array.mapi (fun i cst -> ((mind,i),cst)) consts);
- consts
+
+ let consts = Array.map2 (fun id cl ->
+ define internal id cl mib.mind_polymorphic ctx) ids cl in
+ let schemes = Array.mapi (fun i cst -> ((mind,i),cst)) consts in
+ declare_scheme kind schemes;
+ consts,
+ Declareops.cons_side_effects
+ (Safe_typing.sideff_of_scheme
+ kind (Global.safe_env()) (Array.to_list schemes))
+ eff
let define_mutual_scheme kind internal names mind =
match Hashtbl.find scheme_object_table kind with
@@ -167,16 +179,23 @@ let define_mutual_scheme kind internal names mind =
| s,MutualSchemeFunction f ->
define_mutual_scheme_base kind s f internal names mind
+let find_scheme_on_env_too kind ind =
+ let s = String.Map.find kind (Indmap.find ind !scheme_map) in
+ s, Declareops.cons_side_effects
+ (Safe_typing.sideff_of_scheme
+ kind (Global.safe_env()) [ind, s])
+ Declareops.no_seff
+
let find_scheme kind (mind,i as ind) =
- try Stringmap.find kind (Indmap.find ind !scheme_map)
+ try find_scheme_on_env_too kind ind
with Not_found ->
match Hashtbl.find scheme_object_table kind with
| s,IndividualSchemeFunction f ->
define_individual_scheme_base kind s f KernelSilent None ind
| s,MutualSchemeFunction f ->
- (define_mutual_scheme_base kind s f KernelSilent [] mind).(i)
+ let ca, eff = define_mutual_scheme_base kind s f KernelSilent [] mind in
+ ca.(i), eff
let check_scheme kind ind =
- try let _ = Stringmap.find kind (Indmap.find ind !scheme_map) in true
+ try let _ = find_scheme_on_env_too kind ind in true
with Not_found -> false
-
diff --git a/toplevel/ind_tables.mli b/toplevel/ind_tables.mli
index 3cb3541e..98eaac09 100644
--- a/toplevel/ind_tables.mli
+++ b/toplevel/ind_tables.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,10 +8,6 @@
open Term
open Names
-open Libnames
-open Mod_subst
-open Sign
-open Declarations
(** This module provides support for registering inductive scheme builders,
declaring schemes and generating schemes on demand *)
@@ -22,8 +18,10 @@ type mutual
type individual
type 'a scheme_kind
-type mutual_scheme_object_function = mutual_inductive -> constr array
-type individual_scheme_object_function = inductive -> constr
+type mutual_scheme_object_function =
+ mutual_inductive -> constr array Evd.in_evar_universe_context * Declareops.side_effects
+type individual_scheme_object_function =
+ inductive -> constr Evd.in_evar_universe_context * Declareops.side_effects
(** Main functions to register a scheme builder *)
@@ -31,7 +29,8 @@ val declare_mutual_scheme_object : string -> ?aux:string ->
mutual_scheme_object_function -> mutual scheme_kind
val declare_individual_scheme_object : string -> ?aux:string ->
- individual_scheme_object_function -> individual scheme_kind
+ individual_scheme_object_function ->
+ individual scheme_kind
(*
val declare_scheme : 'a scheme_kind -> (inductive * constant) array -> unit
@@ -41,12 +40,15 @@ val declare_scheme : 'a scheme_kind -> (inductive * constant) array -> unit
val define_individual_scheme : individual scheme_kind ->
Declare.internal_flag (** internal *) ->
- identifier option -> inductive -> constant
+ Id.t option -> inductive -> constant * Declareops.side_effects
val define_mutual_scheme : mutual scheme_kind -> Declare.internal_flag (** internal *) ->
- (int * identifier) list -> mutual_inductive -> constant array
+ (int * Id.t) list -> mutual_inductive -> constant array * Declareops.side_effects
(** Main function to retrieve a scheme in the cache or to generate it *)
-val find_scheme : 'a scheme_kind -> inductive -> constant
+val find_scheme : 'a scheme_kind -> inductive -> constant * Declareops.side_effects
val check_scheme : 'a scheme_kind -> inductive -> bool
+
+
+val pr_scheme_kind : 'a scheme_kind -> Pp.std_ppcmds
diff --git a/toplevel/indschemes.ml b/toplevel/indschemes.ml
index b9888dcd..e6b23828 100644
--- a/toplevel/indschemes.ml
+++ b/toplevel/indschemes.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -15,7 +15,7 @@
declaring new schemes *)
open Pp
-open Flags
+open Errors
open Util
open Names
open Declarations
@@ -26,13 +26,11 @@ open Decl_kinds
open Indrec
open Declare
open Libnames
+open Globnames
open Goptions
open Nameops
open Termops
-open Typeops
-open Inductiveops
open Pretyping
-open Topconstr
open Nametab
open Smartlocate
open Vernacexpr
@@ -53,6 +51,24 @@ let _ =
optread = (fun () -> !elim_flag) ;
optwrite = (fun b -> elim_flag := b) }
+let bifinite_elim_flag = ref false
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "automatic declaration of induction schemes for non-recursive types";
+ optkey = ["Nonrecursive";"Elimination";"Schemes"];
+ optread = (fun () -> !bifinite_elim_flag) ;
+ optwrite = (fun b -> bifinite_elim_flag := b) }
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = true; (* compatibility 2014-09-03*)
+ optname = "automatic declaration of induction schemes for non-recursive types";
+ optkey = ["Record";"Elimination";"Schemes"];
+ optread = (fun () -> !bifinite_elim_flag) ;
+ optwrite = (fun b -> bifinite_elim_flag := b) }
+
let case_flag = ref false
let _ =
declare_bool_option
@@ -69,7 +85,7 @@ let _ =
{ optsync = true;
optdepr = false;
optname = "automatic declaration of boolean equality";
- optkey = ["Boolean";"Equality";"Schemes"];
+ optkey = ["Equality";"Schemes"];
optread = (fun () -> !eq_flag) ;
optwrite = (fun b -> eq_flag := b) }
let _ = (* compatibility *)
@@ -105,14 +121,19 @@ let _ =
(* Util *)
-let define id internal c t =
+let define id internal ctx c t =
let f = declare_constant ~internal in
let kn = f id
(DefinitionEntry
{ const_entry_body = c;
const_entry_secctx = None;
const_entry_type = t;
- const_entry_opaque = false },
+ const_entry_polymorphic = true;
+ const_entry_universes = Evd.universe_context ctx;
+ const_entry_opaque = false;
+ const_entry_inline_code = false;
+ const_entry_feedback = None;
+ },
Decl_kinds.IsDefinition Scheme) in
definition_message id;
kn
@@ -128,7 +149,7 @@ let alarm what internal msg =
| KernelVerbose
| KernelSilent ->
(if debug then
- Flags.if_warn Pp.msg_warning
+ msg_warning
(hov 0 msg ++ fnl () ++ what ++ str " not defined."))
| _ -> errorlabstrm "" msg
@@ -161,14 +182,15 @@ let try_declare_scheme what f internal names kn =
alarm what internal (msg ++ str ".")
| e when Errors.noncritical e ->
alarm what internal
- (str "Unknown exception during scheme creation.")
+ (str "Unknown exception during scheme creation: "++
+ str (Printexc.to_string e))
let beq_scheme_msg mind =
let mib = Global.lookup_mind mind in
(* TODO: mutual inductive case *)
str "Boolean equality on " ++
pr_enum (fun ind -> quote (Printer.pr_inductive (Global.env()) ind))
- (list_tabulate (fun i -> (mind,i)) (Array.length mib.mind_packets))
+ (List.init (Array.length mib.mind_packets) (fun i -> (mind,i)))
let declare_beq_scheme_with l kn =
try_declare_scheme (beq_scheme_msg kn) declare_beq_scheme_gen UserVerbose l kn
@@ -185,12 +207,12 @@ 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 case_dep_scheme_kind_from_type in
let kelim = elim_sorts (mib,mip) in
(* in case the inductive has a type elimination, generates only one
induction scheme, the other ones share the same code with the
apropriate type *)
- if List.mem InType kelim then
+ if Sorts.List.mem InType kelim then
ignore (define_individual_scheme dep KernelVerbose None ind)
(* Induction/recursion schemes *)
@@ -208,18 +230,18 @@ let kinds_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 from_prop = kind == InProp in
let kelim = elim_sorts (mib,mip) in
let elims =
- list_map_filter (fun (sort,kind) ->
- if List.mem sort kelim then Some kind else None)
+ 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
List.iter (fun kind -> ignore (define_individual_scheme kind KernelVerbose None ind))
elims
let declare_induction_schemes kn =
let mib = Global.lookup_mind kn in
- if mib.mind_finite then begin
+ if mib.mind_finite <> Decl_kinds.CoFinite then begin
for i = 0 to Array.length mib.mind_packets - 1 do
declare_one_induction_scheme (kn,i);
done;
@@ -229,7 +251,7 @@ let declare_induction_schemes kn =
let declare_eq_decidability_gen internal names kn =
let mib = Global.lookup_mind kn in
- if mib.mind_finite then
+ if mib.mind_finite <> Decl_kinds.CoFinite then
ignore (define_mutual_scheme eq_dec_scheme_kind internal names kn)
let eq_dec_scheme_msg ind = (* TODO: mutual inductive case *)
@@ -271,7 +293,7 @@ let declare_congr_scheme ind =
then
ignore (define_individual_scheme congr_scheme_kind KernelVerbose None ind)
else
- msg_warn "Cannot build congruence scheme because eq is not found"
+ msg_warning (strbrk "Cannot build congruence scheme because eq is not found")
end
let declare_sym_scheme ind =
@@ -281,6 +303,7 @@ let declare_sym_scheme ind =
(* Scheme command *)
+let smart_global_inductive y = smart_global_inductive y
let rec split_scheme l =
let env = Global.env() in
match l with
@@ -300,7 +323,7 @@ requested
let names inds recs isdep y z =
let ind = smart_global_inductive y in
let sort_of_ind = inductive_sort_family (snd (lookup_mind_specif env ind)) in
- let z' = family_of_sort (interp_sort z) in
+ let z' = interp_elimination_sort z in
let suffix = (
match sort_of_ind with
| InProp ->
@@ -323,7 +346,7 @@ requested
| InType -> recs ^ "t_nodep")
) in
let newid = add_suffix (basename_of_global (IndRef ind)) suffix in
- let newref = (dummy_loc,newid) in
+ let newref = (Loc.ghost,newid) in
((newref,isdep,ind,z)::l1),l2
in
match t with
@@ -334,18 +357,20 @@ requested
let do_mutual_induction_scheme lnamedepindsort =
let lrecnames = List.map (fun ((_,f),_,_,_) -> f) lnamedepindsort
- and sigma = Evd.empty
and env0 = Global.env() in
- let lrecspec =
- List.map
- (fun (_,dep,ind,sort) -> (ind,dep,interp_elimination_sort sort))
- lnamedepindsort
+ let sigma, lrecspec =
+ List.fold_right
+ (fun (_,dep,ind,sort) (evd, l) ->
+ let evd, indu = Evd.fresh_inductive_instance env0 evd ind in
+ (evd, (indu,dep,interp_elimination_sort sort) :: l))
+ lnamedepindsort (Evd.from_env env0,[])
in
- let listdecl = Indrec.build_mutual_induction_scheme env0 sigma lrecspec in
- let rec declare decl fi lrecref =
- let decltype = Retyping.get_type_of env0 Evd.empty decl in
- let decltype = refresh_universes decltype in
- let cst = define fi UserVerbose decl (Some decltype) in
+ let sigma, listdecl = Indrec.build_mutual_induction_scheme env0 sigma lrecspec in
+ let declare decl fi lrecref =
+ let decltype = Retyping.get_type_of env0 sigma decl in
+ (* let decltype = refresh_universes decltype in *)
+ let proof_output = Future.from_val ((decl,Univ.ContextSet.empty),Declareops.no_seff) in
+ let cst = define fi UserVerbose sigma proof_output (Some decltype) in
ConstRef cst :: lrecref
in
let _ = List.fold_right2 declare listdecl lrecnames [] in
@@ -354,25 +379,25 @@ let do_mutual_induction_scheme lnamedepindsort =
let get_common_underlying_mutual_inductive = function
| [] -> assert false
| (id,(mind,i as ind))::l as all ->
- match List.filter (fun (_,(mind',_)) -> mind <> mind') l with
+ match List.filter (fun (_,(mind',_)) -> not (eq_mind mind mind')) l with
| (_,ind')::_ ->
raise (RecursionSchemeError (NotMutualInScheme (ind,ind')))
| [] ->
- if not (list_distinct (List.map snd (List.map snd all))) then
- error "A type occurs twice";
+ if not (List.distinct_f Int.compare (List.map snd (List.map snd all)))
+ then error "A type occurs twice";
mind,
- list_map_filter
+ List.map_filter
(function (Some id,(_,i)) -> Some (i,snd id) | (None,_) -> None) all
let do_scheme l =
let ischeme,escheme = split_scheme l in
(* we want 1 kind of scheme at a time so we check if the user
tried to declare different schemes at once *)
- if (ischeme <> []) && (escheme <> [])
+ if not (List.is_empty ischeme) && not (List.is_empty escheme)
then
error "Do not declare equality and induction scheme at the same time."
else (
- if ischeme <> [] then do_mutual_induction_scheme ischeme
+ if not (List.is_empty ischeme) then do_mutual_induction_scheme ischeme
else
let mind,l = get_common_underlying_mutual_inductive escheme in
declare_beq_scheme_with l mind;
@@ -385,17 +410,19 @@ tried to declare different schemes at once *)
let list_split_rev_at index l =
let rec aux i acc = function
- hd :: tl when i = index -> acc, tl
+ hd :: tl when Int.equal i index -> acc, tl
| hd :: tl -> aux (succ i) (hd :: acc) tl
- | [] -> failwith "list_split_when: Invalid argument"
+ | [] -> failwith "List.split_when: Invalid argument"
in aux 0 [] l
let fold_left' f = function
- [] -> raise (Invalid_argument "fold_left'")
+ [] -> invalid_arg "fold_left'"
| hd :: tl -> List.fold_left f hd tl
let build_combined_scheme env schemes =
- let defs = List.map (fun cst -> (cst, Typeops.type_of_constant env cst)) schemes in
+ let defs = List.map (fun cst -> (* FIXME *)
+ let evd, c = Evd.fresh_constant_instance env Evd.empty cst in
+ (c, Typeops.type_of_constant_in env c)) schemes in
(* let nschemes = List.length schemes in *)
let find_inductive ty =
let (ctx, arity) = decompose_prod ty in
@@ -403,7 +430,7 @@ let build_combined_scheme env schemes =
match kind_of_term last with
| App (ind, args) ->
let ind = destInd ind in
- let (_,spec) = Inductive.lookup_mind_specif env ind in
+ let (_,spec) = Inductive.lookup_mind_specif env (fst ind) in
ctx, ind, spec.mind_nrealargs
| _ -> ctx, destInd last, 0
in
@@ -414,8 +441,8 @@ let build_combined_scheme env schemes =
let coqand = Coqlib.build_coq_and () and coqconj = Coqlib.build_coq_conj () in
let relargs = rel_vect 0 prods in
let concls = List.rev_map
- (fun (cst, t) ->
- mkApp(mkConst cst, relargs),
+ (fun (cst, t) -> (* FIXME *)
+ mkApp(mkConstU cst, relargs),
snd (decompose_prod_n prods t)) defs in
let concl_bod, concl_typ =
fold_left'
@@ -440,18 +467,19 @@ let do_combined_scheme name schemes =
schemes
in
let body,typ = build_combined_scheme (Global.env ()) csts in
- ignore (define (snd name) UserVerbose body (Some typ));
+ let proof_output = Future.from_val ((body,Univ.ContextSet.empty),Declareops.no_seff) in
+ ignore (define (snd name) UserVerbose Evd.empty proof_output (Some typ));
fixpoint_message None [snd name]
(**********************************************************************)
let map_inductive_block f kn n = for i=0 to n-1 do f (kn,i) done
-let mutual_inductive_size kn = Array.length (Global.lookup_mind kn).mind_packets
-
let declare_default_schemes kn =
- let n = mutual_inductive_size kn in
- if !elim_flag then declare_induction_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
+ declare_induction_schemes kn;
if !case_flag then map_inductive_block declare_one_case_analysis_scheme kn n;
if is_eq_flag() then try_declare_beq_scheme kn;
if !eq_dec_flag then try_declare_eq_decidability kn;
diff --git a/toplevel/indschemes.mli b/toplevel/indschemes.mli
index f6eb4fe1..98746107 100644
--- a/toplevel/indschemes.mli
+++ b/toplevel/indschemes.mli
@@ -1,20 +1,17 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Util
+open Loc
open Names
open Term
open Environ
-open Libnames
-open Glob_term
-open Genarg
open Vernacexpr
-open Ind_tables
+open Misctypes
(** See also Auto_ind_decl, Indrec, Eqscheme, Ind_tables, ... *)
@@ -35,17 +32,17 @@ val declare_rewriting_schemes : inductive -> unit
(** Mutual Minimality/Induction scheme *)
val do_mutual_induction_scheme :
- (identifier located * bool * inductive * glob_sort) list -> unit
+ (Id.t located * bool * inductive * glob_sort) list -> unit
(** Main calls to interpret the Scheme command *)
-val do_scheme : (identifier located option * scheme) list -> unit
+val do_scheme : (Id.t located option * scheme) list -> unit
(** Combine a list of schemes into a conjunction of them *)
val build_combined_scheme : env -> constant list -> constr * types
-val do_combined_scheme : identifier located -> identifier located list -> unit
+val do_combined_scheme : Id.t located -> Id.t located list -> unit
(** Hook called at each inductive type definition *)
diff --git a/toplevel/interface.mli b/toplevel/interface.mli
deleted file mode 100644
index bb338a96..00000000
--- a/toplevel/interface.mli
+++ /dev/null
@@ -1,231 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(** * Declarative part of the interface of CoqIde calls to Coq *)
-
-(** * Generic structures *)
-
-type raw = bool
-type verbose = bool
-
-(** The type of coqtop goals *)
-type goal = {
- goal_id : string;
- (** Unique goal identifier *)
- goal_hyp : string list;
- (** List of hypotheses *)
- goal_ccl : string;
- (** Goal conclusion *)
-}
-
-type evar = {
- evar_info : string;
- (** A string describing an evar: type, number, environment *)
-}
-
-type status = {
- status_path : string list;
- (** Module path of the current proof *)
- status_proofname : string option;
- (** Current proof name. [None] if no focussed proof is in progress *)
- status_allproofs : string list;
- (** List of all pending proofs. Order is not significant *)
- status_statenum : int;
- (** A unique id describing the state of coqtop *)
- status_proofnum : int;
- (** An id describing the state of the current proof. *)
-}
-
-type goals = {
- fg_goals : goal list;
- (** List of the focussed goals *)
- bg_goals : (goal list * goal list) list;
- (** Zipper representing the unfocussed background goals *)
-}
-
-type hint = (string * string) list
-(** A list of tactics applicable and their appearance *)
-
-type option_name = Goptionstyp.option_name
-
-type option_value = Goptionstyp.option_value =
- | BoolValue of bool
- | IntValue of int option
- | StringValue of string
-
-(** Summary of an option status *)
-type option_state = Goptionstyp.option_state = {
- opt_sync : bool;
- (** Whether an option is synchronous *)
- opt_depr : bool;
- (** Wheter an option is deprecated *)
- opt_name : string;
- (** A short string that is displayed when using [Test] *)
- opt_value : option_value;
- (** The current value of the option *)
-}
-
-type search_constraint =
-(** Whether the name satisfies a regexp (uses Ocaml Str syntax) *)
-| Name_Pattern of string
-(** Whether the object type satisfies a pattern *)
-| Type_Pattern of string
-(** Whether some subtype of object type satisfies a pattern *)
-| SubType_Pattern of string
-(** Whether the object pertains to a module *)
-| In_Module of string list
-(** Bypass the Search blacklist *)
-| Include_Blacklist
-
-(** A list of search constraints; the boolean flag is set to [false] whenever
- the flag should be negated. *)
-type search_flags = (search_constraint * bool) list
-
-(** A named object in Coq. [coq_object_qualid] is the shortest path defined for
- the user. [coq_object_prefix] is the missing part to recover the fully
- qualified name, i.e [fully_qualified = coq_object_prefix + coq_object_qualid].
- [coq_object_object] is the actual content of the object. *)
-type 'a coq_object = {
- coq_object_prefix : string list;
- coq_object_qualid : string list;
- coq_object_object : 'a;
-}
-
-type coq_info = {
- coqtop_version : string;
- protocol_version : string;
- release_date : string;
- compile_date : string;
-}
-
-(** Coq unstructured messages *)
-
-type message_level =
- | Debug of string
- | Info
- | Notice
- | Warning
- | Error
-
-type message = {
- message_level : message_level;
- message_content : string;
-}
-
-(** Coq "semantic" infos obtained during parsing/execution *)
-type edit_id = int
-
-type feedback_content =
- | AddedAxiom
- | Processed
- | GlobRef of (int*int) * string * string * string * string
-
-type feedback = {
- edit_id : edit_id;
- content : feedback_content;
-}
-
-(** Calls result *)
-
-type location = (int * int) option (* start and end of the error *)
-
-type 'a value =
- | Good of 'a
- | Fail of (location * string)
-
-(* Request/Reply message protocol between Coq and CoqIde *)
-
-(** Running a command (given as its id and its text).
- "raw" mode (less sanity checks, no impact on the undo stack)
- is suitable for queries, or for the Set/Unset
- of display options that coqide performs all the time.
- The returned string contains the messages produced
- but not the updated goals (they are
- to be fetch by a separated [current_goals]). *)
-type interp_sty = edit_id * raw * verbose * string
-type interp_rty = string
-
-(** Backtracking by at least a certain number of phrases.
- No finished proofs will be re-opened. Instead,
- we continue backtracking until before these proofs,
- and answer the amount of extra backtracking performed.
- Backtracking by more than the number of phrases already
- interpreted successfully (and not yet undone) will fail. *)
-type rewind_sty = int
-type rewind_rty = int
-
-(** Fetching the list of current goals. Return [None] if no proof is in
- progress, [Some gl] otherwise. *)
-type goals_sty = unit
-type goals_rty = goals option
-
-(** Retrieve the list of unintantiated evars in the current proof. [None] if no
- proof is in progress. *)
-type evars_sty = unit
-type evars_rty = evar list option
-
-(** Retrieving the tactics applicable to the current goal. [None] if there is
- no proof in progress. *)
-type hints_sty = unit
-type hints_rty = (hint list * hint) option
-
-(** The status, for instance "Ready in SomeSection, proving Foo" *)
-type status_sty = unit
-type status_rty = status
-
-(** Search for objects satisfying the given search flags. *)
-type search_sty = search_flags
-type search_rty = string coq_object list
-
-(** Retrieve the list of options of the current toplevel *)
-type get_options_sty = unit
-type get_options_rty = (option_name * option_state) list
-
-(** Set the options to the given value. Warning: this is not atomic, so whenever
- the call fails, the option state can be messed up... This is the caller duty
- to check that everything is correct. *)
-type set_options_sty = (option_name * option_value) list
-type set_options_rty = unit
-
-(** Is a directory part of Coq's loadpath ? *)
-type inloadpath_sty = string
-type inloadpath_rty = bool
-
-(** Create a "match" template for a given inductive type.
- For each branch of the match, we list the constructor name
- followed by enough pattern variables. *)
-type mkcases_sty = string
-type mkcases_rty = string list list
-
-(** Quit gracefully the interpreter. *)
-type quit_sty = unit
-type quit_rty = unit
-
-type about_sty = unit
-type about_rty = coq_info
-
-type handle_exn_rty = location * string
-type handle_exn_sty = exn
-
-type handler = {
- interp : interp_sty -> interp_rty;
- rewind : rewind_sty -> rewind_rty;
- goals : goals_sty -> goals_rty;
- evars : evars_sty -> evars_rty;
- hints : hints_sty -> hints_rty;
- status : status_sty -> status_rty;
- search : search_sty -> search_rty;
- get_options : get_options_sty -> get_options_rty;
- set_options : set_options_sty -> set_options_rty;
- inloadpath : inloadpath_sty -> inloadpath_rty;
- mkcases : mkcases_sty -> mkcases_rty;
- quit : quit_sty -> quit_rty;
- about : about_sty -> about_rty;
- handle_exn : handle_exn_sty -> handle_exn_rty;
-}
-
diff --git a/toplevel/lemmas.ml b/toplevel/lemmas.ml
deleted file mode 100644
index 4bc9cdb0..00000000
--- a/toplevel/lemmas.ml
+++ /dev/null
@@ -1,353 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(* Created by Hugo Herbelin from contents related to lemma proofs in
- file command.ml, Aug 2009 *)
-
-open Util
-open Flags
-open Pp
-open Names
-open Term
-open Declarations
-open Entries
-open Environ
-open Nameops
-open Libnames
-open Decls
-open Decl_kinds
-open Declare
-open Pretyping
-open Termops
-open Namegen
-open Evd
-open Evarutil
-open Reductionops
-open Topconstr
-open Constrintern
-open Impargs
-open Tacticals
-
-(* Support for mutually proved theorems *)
-
-let retrieve_first_recthm = function
- | VarRef id ->
- (pi2 (Global.lookup_named id),variable_opacity id)
- | ConstRef cst ->
- let cb = Global.lookup_constant cst in
- (Option.map Declarations.force (body_of_constant cb), is_opaque cb)
- | _ -> assert false
-
-let adjust_guardness_conditions const = function
- | [] -> const (* Not a recursive statement *)
- | possible_indexes ->
- (* Try all combinations... not optimal *)
- match kind_of_term const.const_entry_body with
- | Fix ((nv,0),(_,_,fixdefs as fixdecls)) ->
-(* let possible_indexes =
- List.map2 (fun i c -> match i with Some i -> i | None ->
- interval 0 (List.length ((lam_assum c))))
- lemma_guard (Array.to_list fixdefs) in
-*)
- let indexes =
- search_guard dummy_loc (Global.env()) possible_indexes fixdecls in
- { const with const_entry_body = mkFix ((indexes,0),fixdecls) }
- | c -> const
-
-let find_mutually_recursive_statements thms =
- let n = List.length thms in
- let inds = List.map (fun (id,(t,impls,annot)) ->
- let (hyps,ccl) = decompose_prod_assum t in
- let x = (id,(t,impls)) in
- match annot with
- (* Explicit fixpoint decreasing argument is given *)
- | Some (Some (_,id),CStructRec) ->
- let i,b,typ = lookup_rel_id id hyps in
- (match kind_of_term t with
- | Ind (kn,_ as ind) when
- let mind = Global.lookup_mind kn in
- mind.mind_finite & b = None ->
- [ind,x,i],[]
- | _ ->
- error "Decreasing argument is not an inductive assumption.")
- (* Unsupported cases *)
- | Some (_,(CWfRec _|CMeasureRec _)) ->
- error "Only structural decreasing is supported for mutual statements."
- (* Cofixpoint or fixpoint w/o explicit decreasing argument *)
- | None | Some (None, CStructRec) ->
- let whnf_hyp_hds = map_rel_context_in_env
- (fun env c -> fst (whd_betadeltaiota_stack env Evd.empty c))
- (Global.env()) hyps in
- let ind_hyps =
- List.flatten (list_map_i (fun i (_,b,t) ->
- match kind_of_term t with
- | Ind (kn,_ as ind) when
- let mind = Global.lookup_mind kn in
- mind.mind_finite & b = None ->
- [ind,x,i]
- | _ ->
- []) 0 (List.rev whnf_hyp_hds)) in
- let ind_ccl =
- let cclenv = push_rel_context hyps (Global.env()) in
- let whnf_ccl,_ = whd_betadeltaiota_stack cclenv Evd.empty ccl in
- match kind_of_term whnf_ccl with
- | Ind (kn,_ as ind) when
- let mind = Global.lookup_mind kn in
- mind.mind_ntypes = n & not mind.mind_finite ->
- [ind,x,0]
- | _ ->
- [] in
- ind_hyps,ind_ccl) thms in
- let inds_hyps,ind_ccls = List.split inds in
- let of_same_mutind ((kn,_),_,_) = function ((kn',_),_,_) -> kn = kn' in
- (* Check if all conclusions are coinductive in the same type *)
- (* (degenerated cartesian product since there is at most one coind ccl) *)
- let same_indccl =
- list_cartesians_filter (fun hyp oks ->
- if List.for_all (of_same_mutind hyp) oks
- then Some (hyp::oks) else None) [] ind_ccls in
- let ordered_same_indccl =
- List.filter (list_for_all_i (fun i ((kn,j),_,_) -> i=j) 0) same_indccl in
- (* Check if some hypotheses are inductive in the same type *)
- let common_same_indhyp =
- list_cartesians_filter (fun hyp oks ->
- if List.for_all (of_same_mutind hyp) oks
- then Some (hyp::oks) else None) [] inds_hyps in
- let ordered_inds,finite,guard =
- match ordered_same_indccl, common_same_indhyp with
- | indccl::rest, _ ->
- assert (rest=[]);
- (* One occ. of common coind ccls and no common inductive hyps *)
- if common_same_indhyp <> [] then
- if_verbose msgnl (str "Assuming mutual coinductive statements.");
- flush_all ();
- indccl, true, []
- | [], _::_ ->
- if same_indccl <> [] &&
- list_distinct (List.map pi1 (List.hd same_indccl)) then
- if_verbose msgnl (strbrk "Coinductive statements do not follow the order of definition, assuming the proof to be by induction."); flush_all ();
- let possible_guards = List.map (List.map pi3) inds_hyps in
- (* assume the largest indices as possible *)
- list_last common_same_indhyp, false, possible_guards
- | _, [] ->
- error
- ("Cannot find common (mutual) inductive premises or coinductive" ^
- " conclusions in the statements.")
- in
- (finite,guard,None), ordered_inds
-
-let look_for_possibly_mutual_statements = function
- | [id,(t,impls,None)] ->
- (* One non recursively proved theorem *)
- None,[id,(t,impls)],None
- | _::_ as thms ->
- (* More than one statement and/or an explicit decreasing mark: *)
- (* we look for a common inductive hyp or a common coinductive conclusion *)
- let recguard,ordered_inds = find_mutually_recursive_statements thms in
- let thms = List.map pi2 ordered_inds in
- Some recguard,thms, Some (List.map (fun (_,_,i) -> succ i) ordered_inds)
- | [] -> anomaly "Empty list of theorems."
-
-(* Saving a goal *)
-
-let save id const do_guard (locality,kind) hook =
- let const = adjust_guardness_conditions const do_guard in
- let {const_entry_body = pft;
- const_entry_type = tpo;
- const_entry_opaque = opacity } = const in
- let k = logical_kind_of_goal_kind kind in
- let l,r = match locality with
- | Local when Lib.sections_are_opened () ->
- let c = SectionLocalDef (pft, tpo, opacity) in
- let _ = declare_variable id (Lib.cwd(), c, k) in
- (Local, VarRef id)
- | Local | Global ->
- let kn = declare_constant id (DefinitionEntry const, k) in
- Autoinstance.search_declaration (ConstRef kn);
- (Global, ConstRef kn) in
- Pfedit.delete_current_proof ();
- definition_message id;
- hook l r
-
-let default_thm_id = id_of_string "Unnamed_thm"
-
-let compute_proof_name locality = function
- | Some (loc,id) ->
- (* We check existence here: it's a bit late at Qed time *)
- if Nametab.exists_cci (Lib.make_path id) || is_section_variable id ||
- locality=Global && Nametab.exists_cci (Lib.make_path_except_section id)
- then
- user_err_loc (loc,"",pr_id id ++ str " already exists.");
- id
- | None ->
- next_global_ident_away default_thm_id (Pfedit.get_all_proof_names ())
-
-let save_remaining_recthms (local,kind) body opaq i (id,(t_i,(_,imps))) =
- match body with
- | None ->
- (match local with
- | Local ->
- let impl=false in (* copy values from Vernacentries *)
- let k = IsAssumption Conjectural in
- let c = SectionLocalAssum (t_i,impl) in
- let _ = declare_variable id (Lib.cwd(),c,k) in
- (Local,VarRef id,imps)
- | Global ->
- let k = IsAssumption Conjectural in
- let kn = declare_constant id (ParameterEntry (None,t_i,None), k) in
- (Global,ConstRef kn,imps))
- | Some body ->
- let k = logical_kind_of_goal_kind kind in
- let body_i = match kind_of_term body with
- | Fix ((nv,0),decls) -> mkFix ((nv,i),decls)
- | CoFix (0,decls) -> mkCoFix (i,decls)
- | _ -> anomaly "Not a proof by induction" in
- match local with
- | Local ->
- let c = SectionLocalDef (body_i, Some t_i, opaq) in
- let _ = declare_variable id (Lib.cwd(), c, k) in
- (Local,VarRef id,imps)
- | Global ->
- let const =
- { const_entry_body = body_i;
- const_entry_secctx = None;
- const_entry_type = Some t_i;
- const_entry_opaque = opaq } in
- let kn = declare_constant id (DefinitionEntry const, k) in
- (Global,ConstRef kn,imps)
-
-let save_hook = ref ignore
-let set_save_hook f = save_hook := f
-
-let get_proof opacity =
- let id,(const,do_guard,persistence,hook) = Pfedit.cook_proof !save_hook in
- id,{const with const_entry_opaque = opacity},do_guard,persistence,hook
-
-let save_named opacity =
- let p = Proof_global.give_me_the_proof () in
- Proof.transaction p begin fun () ->
- let id,const,do_guard,persistence,hook = get_proof opacity in
- save id const do_guard persistence hook
- end
-
-let check_anonymity id save_ident =
- if atompart_of_id id <> string_of_id (default_thm_id) then
- error "This command can only be used for unnamed theorem."
-
-let save_anonymous opacity save_ident =
- let p = Proof_global.give_me_the_proof () in
- Proof.transaction p begin fun () ->
- let id,const,do_guard,persistence,hook = get_proof opacity in
- check_anonymity id save_ident;
- save save_ident const do_guard persistence hook
- end
-
-let save_anonymous_with_strength kind opacity save_ident =
- let p = Proof_global.give_me_the_proof () in
- Proof.transaction p begin fun () ->
- let id,const,do_guard,_,hook = get_proof opacity in
- check_anonymity id save_ident;
- (* we consider that non opaque behaves as local for discharge *)
- save save_ident const do_guard (Global, Proof kind) hook
- end
-
-(* Starting a goal *)
-
-let start_hook = ref ignore
-let set_start_hook = (:=) start_hook
-
-let start_proof id kind c ?init_tac ?(compute_guard=[]) hook =
- let sign = initialize_named_context_for_proof () in
- !start_hook c;
- Pfedit.start_proof id kind sign c ?init_tac ~compute_guard hook
-
-let rec_tac_initializer finite guard thms snl =
- if finite then
- match List.map (fun (id,(t,_)) -> (id,t)) thms with
- | (id,_)::l -> Hiddentac.h_mutual_cofix true id l
- | _ -> assert false
- else
- (* nl is dummy: it will be recomputed at Qed-time *)
- let nl = match snl with
- | None -> List.map succ (List.map list_last guard)
- | Some nl -> nl
- in match List.map2 (fun (id,(t,_)) n -> (id,n,t)) thms nl with
- | (id,n,_)::l -> Hiddentac.h_mutual_fix true id n l
- | _ -> assert false
-
-let start_proof_with_initialization kind recguard thms snl hook =
- let intro_tac (_, (_, (ids, _))) =
- Refiner.tclMAP (function
- | Name id -> Tactics.intro_mustbe_force id
- | Anonymous -> Tactics.intro) (List.rev ids) in
- let init_tac,guard = match recguard with
- | Some (finite,guard,init_tac) ->
- let rec_tac = rec_tac_initializer finite guard thms snl in
- Some (match init_tac with
- | None ->
- if Flags.is_auto_intros () then
- tclTHENS rec_tac (List.map intro_tac thms)
- else
- rec_tac
- | Some tacl ->
- tclTHENS rec_tac
- (if Flags.is_auto_intros () then
- List.map2 (fun tac thm -> tclTHEN tac (intro_tac thm)) tacl thms
- else
- tacl)),guard
- | None ->
- assert (List.length thms = 1);
- (if Flags.is_auto_intros () then Some (intro_tac (List.hd thms)) else None), [] in
- match thms with
- | [] -> anomaly "No proof to start"
- | (id,(t,(_,imps)))::other_thms ->
- let hook strength ref =
- let other_thms_data =
- if other_thms = [] then [] else
- (* there are several theorems defined mutually *)
- let body,opaq = retrieve_first_recthm ref in
- list_map_i (save_remaining_recthms kind body opaq) 1 other_thms in
- let thms_data = (strength,ref,imps)::other_thms_data in
- List.iter (fun (strength,ref,imps) ->
- maybe_declare_manual_implicits false ref imps;
- hook strength ref) thms_data in
- start_proof id kind t ?init_tac hook ~compute_guard:guard
-
-let start_proof_com kind thms hook =
- let evdref = ref Evd.empty in
- let env0 = Global.env () in
- let thms = List.map (fun (sopt,(bl,t,guard)) ->
- let impls, ((env, ctx), imps) = interp_context_evars evdref env0 bl in
- let t', imps' = interp_type_evars_impls ~impls ~evdref env t in
- Sign.iter_rel_context (check_evars env Evd.empty !evdref) ctx;
- let ids = List.map pi1 ctx in
- (compute_proof_name (fst kind) sopt,
- (nf_evar !evdref (it_mkProd_or_LetIn t' ctx),
- (ids, imps @ lift_implicits (List.length ids) imps'),
- guard)))
- thms in
- let recguard,thms,snl = look_for_possibly_mutual_statements thms in
- start_proof_with_initialization kind recguard thms snl hook
-
-(* Admitted *)
-
-let admit () =
- let (id,k,typ,hook) = Pfedit.current_proof_statement () in
- let e = Pfedit.get_used_variables(), typ, None in
- let kn =
- declare_constant id (ParameterEntry e,IsAssumption Conjectural) in
- Pfedit.delete_current_proof ();
- assumption_message id;
- hook Global (ConstRef kn)
-
-(* Miscellaneous *)
-
-let get_current_context () =
- try Pfedit.get_current_goal_context ()
- with e when Logic.catchable_exception e ->
- (Evd.empty, Global.env())
diff --git a/toplevel/lemmas.mli b/toplevel/lemmas.mli
deleted file mode 100644
index 1e68f846..00000000
--- a/toplevel/lemmas.mli
+++ /dev/null
@@ -1,63 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-open Names
-open Term
-open Decl_kinds
-open Topconstr
-open Tacexpr
-open Vernacexpr
-open Proof_type
-open Pfedit
-
-(** A hook start_proof calls on the type of the definition being started *)
-val set_start_hook : (types -> unit) -> unit
-
-val start_proof : identifier -> goal_kind -> types ->
- ?init_tac:tactic -> ?compute_guard:lemma_possible_guards ->
- declaration_hook -> unit
-
-val start_proof_com : goal_kind ->
- (lident option * (local_binder list * constr_expr * (lident option * recursion_order_expr) option)) list ->
- declaration_hook -> unit
-
-val start_proof_with_initialization :
- goal_kind -> (bool * lemma_possible_guards * tactic list option) option ->
- (identifier * (types * (name list * Impargs.manual_explicitation list))) list
- -> int list option -> declaration_hook -> unit
-
-(** A hook the next three functions pass to cook_proof *)
-val set_save_hook : (Proof.proof -> unit) -> unit
-
-(** {6 ... } *)
-(** [save_named b] saves the current completed proof under the name it
-was started; boolean [b] tells if the theorem is declared opaque; it
-fails if the proof is not completed *)
-
-val save_named : bool -> unit
-
-(** [save_anonymous b name] behaves as [save_named] but declares the theorem
-under the name [name] and respects the strength of the declaration *)
-
-val save_anonymous : bool -> identifier -> unit
-
-(** [save_anonymous_with_strength s b name] behaves as [save_anonymous] but
- declares the theorem under the name [name] and gives it the
- strength [strength] *)
-
-val save_anonymous_with_strength : theorem_kind -> bool -> identifier -> unit
-
-(** [admit ()] aborts the current goal and save it as an assmumption *)
-
-val admit : unit -> unit
-
-(** [get_current_context ()] returns the evar context and env of the
- current open proof if any, otherwise returns the empty evar context
- and the current global env *)
-
-val get_current_context : unit -> Evd.evar_map * Environ.env
diff --git a/toplevel/libtypes.ml b/toplevel/libtypes.ml
deleted file mode 100644
index b392f82c..00000000
--- a/toplevel/libtypes.ml
+++ /dev/null
@@ -1,111 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-open Term
-open Summary
-open Libobject
-open Libnames
-open Names
-(*
- * Module construction
- *)
-
-(* let reduce c = Reductionops.head_unfold_under_prod *)
-(* (Auto.Hint_db.transparent_state (Auto.searchtable_map "typeclass_instances")) *)
-(* (Global.env()) Evd.empty c *)
-
-let reduce c = c
-
-module TypeDnet = Term_dnet.Make
- (struct
- type t = Libnames.global_reference
- let compare = RefOrdered.compare
- let subst s gr = fst (Libnames.subst_global s gr)
- let constr_of = Global.type_of_global
- end)
- (struct let reduce = reduce
- let direction = false
- end)
-
-type result = Libnames.global_reference * (constr*existential_key) * Termops.subst
-
-let all_types = ref TypeDnet.empty
-let defined_types = ref TypeDnet.empty
-
-(*
- * Bookeeping & States
- *)
-
-let freeze () =
- (!all_types,!defined_types)
-
-let unfreeze (lt,dt) =
- all_types := lt;
- defined_types := dt
-
-let init () =
- all_types := TypeDnet.empty;
- defined_types := TypeDnet.empty
-
-let _ =
- declare_summary "type-library-state"
- { freeze_function = freeze;
- unfreeze_function = unfreeze;
- init_function = init }
-
-let load (_,d) =
-(* Profile.print_logical_stats !all_types;
- Profile.print_logical_stats d;*)
- all_types := TypeDnet.union d !all_types
-
-let subst s t = TypeDnet.subst s t
-(*
-let subst_key = Profile.declare_profile "subst"
-let subst a b = Profile.profile2 subst_key TypeDnet.subst a b
-
-let load_key = Profile.declare_profile "load"
-let load a = Profile.profile1 load_key load a
-*)
-let input : TypeDnet.t -> obj =
- declare_object
- { (default_object "LIBTYPES") with
- load_function = (fun _ -> load);
- subst_function = (fun (s,t) -> subst s t);
- classify_function = (fun x -> Substitute x)
- }
-
-let update () = Lib.add_anonymous_leaf (input !defined_types)
-
-(*
- * Search interface
- *)
-
-let search_pattern pat = TypeDnet.search_pattern !all_types pat
-let search_concl pat = TypeDnet.search_concl !all_types pat
-let search_head_concl pat = TypeDnet.search_head_concl !all_types pat
-let search_eq_concl eq pat = TypeDnet.search_eq_concl !all_types eq pat
-
-let add typ gr =
- defined_types := TypeDnet.add typ gr !defined_types;
- all_types := TypeDnet.add typ gr !all_types
-(*
-let add_key = Profile.declare_profile "add"
-let add a b = Profile.profile1 add_key add a b
-*)
-
-(*
- * Hooks declaration
- *)
-
-let _ = Declare.add_cache_hook
- ( fun sp ->
- let gr = Nametab.global_of_path sp in
- let ty = Global.type_of_global gr in
- add ty gr )
-
-let _ = Declaremods.set_end_library_hook update
diff --git a/toplevel/libtypes.mli b/toplevel/libtypes.mli
deleted file mode 100644
index f8d31edb..00000000
--- a/toplevel/libtypes.mli
+++ /dev/null
@@ -1,26 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-open Term
-
-(** Persistent library of all declared object, indexed by their types
- (uses Dnets) *)
-
-(** results are the reference of the object, together with a context
-(constr+evar) and a substitution under this context *)
-type result = Libnames.global_reference * (constr*existential_key) * Termops.subst
-
-(** this is the reduction function used in the indexing process *)
-val reduce : types -> types
-
-(** The different types of search available.
- See term_dnet.mli for more explanations *)
-val search_pattern : types -> result list
-val search_concl : types -> result list
-val search_head_concl : types -> result list
-val search_eq_concl : constr -> types -> result list
diff --git a/toplevel/locality.ml b/toplevel/locality.ml
new file mode 100644
index 00000000..f711dad9
--- /dev/null
+++ b/toplevel/locality.ml
@@ -0,0 +1,99 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** * Managing locality *)
+
+let local_of_bool = function
+ | true -> Decl_kinds.Local
+ | false -> Decl_kinds.Global
+
+let check_locality locality_flag =
+ match locality_flag with
+ | Some b ->
+ let s = if b then "Local" else "Global" in
+ Errors.error ("This command does not support the \""^s^"\" prefix.")
+ | None -> ()
+
+(** Extracting the locality flag *)
+
+(* Commands which supported an inlined Local flag *)
+
+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."
+ | Some true when local ->
+ Errors.error "Use only prefix \"Local\"."
+ | None ->
+ if local then begin
+ Pp.msg_warning (Pp.str "Obsolete syntax: use \"Local\" as a prefix.");
+ Some true
+ end else
+ None
+ | Some b -> Some b in
+ local
+
+(** Positioning locality for commands supporting discharging and export
+ outside of modules *)
+
+(* For commands whose default is to discharge and export:
+ Global is the default and is neutral;
+ Local in a section deactivates discharge,
+ Local not in a section deactivates export *)
+let make_non_locality = function Some false -> false | _ -> true
+
+let make_locality = function Some true -> true | _ -> false
+
+let enforce_locality locality_flag local =
+ make_locality (enforce_locality_full locality_flag local)
+
+let enforce_locality_exp locality_flag local =
+ match locality_flag, local with
+ | None, Some local -> local
+ | Some b, None -> local_of_bool b
+ | None, None -> Decl_kinds.Global
+ | Some _, Some _ -> Errors.error "Local non allowed in this case"
+
+(* For commands whose default is to not discharge but to export:
+ Global in sections forces discharge, Global not in section is the default;
+ Local in sections is the default, Local not in section forces non-export *)
+
+let make_section_locality =
+ function Some b -> b | None -> Lib.sections_are_opened ()
+
+let enforce_section_locality locality_flag local =
+ make_section_locality (enforce_locality_full locality_flag local)
+
+(** Positioning locality for commands supporting export but not discharge *)
+
+(* For commands whose default is to export (if not in section):
+ Global in sections is forbidden, Global not in section is neutral;
+ Local in sections is the default, Local not in section forces non-export *)
+
+let make_module_locality = function
+ | Some false ->
+ if Lib.sections_are_opened () then
+ Errors.error
+ "This command does not support the Global option in sections.";
+ false
+ | Some true -> true
+ | None -> false
+
+let enforce_module_locality locality_flag local =
+ make_module_locality (enforce_locality_full locality_flag local)
+
+module LocalityFixme = struct
+ let locality = ref None
+ let set l = locality := l
+ let consume () =
+ let l = !locality in
+ locality := None;
+ l
+ let assert_consumed () = check_locality !locality
+end
diff --git a/toplevel/locality.mli b/toplevel/locality.mli
new file mode 100644
index 00000000..c395fe92
--- /dev/null
+++ b/toplevel/locality.mli
@@ -0,0 +1,51 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** * Managing locality *)
+
+(** Commands which supported an inlined Local flag *)
+
+val enforce_locality_full : bool option -> bool -> bool option
+
+(** * Positioning locality for commands supporting discharging and export
+ outside of modules *)
+
+(** For commands whose default is to discharge and export:
+ Global is the default and is neutral;
+ Local in a section deactivates discharge,
+ Local not in a section deactivates export *)
+
+val make_locality : bool option -> bool
+val make_non_locality : bool option -> bool
+val enforce_locality : bool option -> bool -> bool
+val enforce_locality_exp :
+ bool option -> Decl_kinds.locality option -> Decl_kinds.locality
+
+(** For commands whose default is to not discharge but to export:
+ Global in sections forces discharge, Global not in section is the default;
+ Local in sections is the default, Local not in section forces non-export *)
+
+val make_section_locality : bool option -> bool
+val enforce_section_locality : bool option -> bool -> bool
+
+(** * Positioning locality for commands supporting export but not discharge *)
+
+(** For commands whose default is to export (if not in section):
+ Global in sections is forbidden, Global not in section is neutral;
+ Local in sections is the default, Local not in section forces non-export *)
+
+val make_module_locality : bool option -> bool
+val enforce_module_locality : bool option -> bool -> bool
+
+(* This is the old imperative interface that is still used for
+ * VernacExtend vernaculars. Time permitting this could be trashed too *)
+module LocalityFixme : sig
+ val set : bool option -> unit
+ val consume : unit -> bool option
+ val assert_consumed : unit -> unit
+end
diff --git a/toplevel/metasyntax.ml b/toplevel/metasyntax.ml
index 5a20f1c1..161cf824 100644
--- a/toplevel/metasyntax.ml
+++ b/toplevel/metasyntax.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,33 +8,34 @@
open Pp
open Flags
-open Compat
+open Errors
open Util
open Names
-open Topconstr
+open Constrexpr
+open Constrexpr_ops
+open Notation_term
+open Notation_ops
open Ppextend
open Extend
open Libobject
-open Summary
open Constrintern
open Vernacexpr
open Pcoq
-open Glob_term
open Libnames
open Tok
-open Lexer
-open Egrammar
+open Egramml
+open Egramcoq
open Notation
open Nameops
(**********************************************************************)
(* Tokens *)
-let cache_token (_,s) = add_keyword s
+let cache_token (_,s) = Lexer.add_keyword s
let inToken : string -> obj =
declare_object {(default_object "TOKEN") with
- open_function = (fun i o -> if i=1 then cache_token o);
+ open_function = (fun i o -> if Int.equal i 1 then cache_token o);
cache_function = cache_token;
subst_function = Libobject.ident_subst_function;
classify_function = (fun o -> Substitute o)}
@@ -60,113 +61,184 @@ let rec make_tags = function
| GramNonTerminal (loc, etyp, _, po) :: l -> etyp :: make_tags l
| [] -> []
-let cache_tactic_notation (_,(pa,pp)) =
- Egrammar.extend_grammar (Egrammar.TacticGrammar pa);
- Pptactic.declare_extra_tactic_pprule pp
+type tactic_grammar_obj = {
+ tacobj_local : locality_flag;
+ tacobj_tacgram : tactic_grammar;
+ tacobj_tacpp : Pptactic.pp_tactic;
+ tacobj_body : Tacexpr.glob_tactic_expr
+}
-let subst_tactic_parule subst (key,n,p,(d,tac)) =
- (key,n,p,(d,Tacinterp.subst_tactic subst tac))
+let cache_tactic_notation ((_, key), tobj) =
+ 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 subst_tactic_notation (subst,(pa,pp)) =
- (subst_tactic_parule subst pa,pp)
+let open_tactic_notation i ((_, key), tobj) =
+ if Int.equal i 1 && not tobj.tacobj_local then
+ Egramcoq.extend_tactic_grammar key tobj.tacobj_tacgram
-type tactic_grammar_obj =
- (string * int * grammar_prod_item list *
- (dir_path * Tacexpr.glob_tactic_expr))
- * (string * Genarg.argument_type list * (int * Pptactic.grammar_terminals))
+let load_tactic_notation i ((_, key), tobj) =
+ (** 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_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 = (fun i o -> if i=1 then cache_tactic_notation o);
+ open_function = open_tactic_notation;
+ load_function = load_tactic_notation;
cache_function = cache_tactic_notation;
subst_function = subst_tactic_notation;
- classify_function = (fun o -> Substitute o)}
+ classify_function = classify_tactic_notation}
let cons_production_parameter l = function
| GramTerminal _ -> l
| GramNonTerminal (_,_,_,ido) -> Option.List.cons ido l
-let rec tactic_notation_key = function
- | GramTerminal id :: _ -> id
- | _ :: l -> tactic_notation_key l
- | [] -> "terminal_free_notation"
-
-let rec next_key_away key t =
- if Pptactic.exists_extra_tactic_pprule key t then next_key_away (key^"'") t
- else key
-
-let add_tactic_notation (n,prods,e) =
+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 key = next_key_away (tactic_notation_key prods) tags in
- let pprule = (key,tags,(n,List.map make_terminal_status 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 = Tacinterp.glob_tactic_env ids (Global.env()) e in
- let parule = (key,n,prods,(Lib.cwd(),tac)) in
- Lib.add_anonymous_leaf (inTacticGrammar (parule,pprule))
+ let tac = Tacintern.glob_tactic_env ids (Global.env()) e in
+ let parule = {
+ tacgram_level = n;
+ tacgram_prods = prods;
+ } in
+ let tacobj = {
+ 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 print_grammar = function
+let entry_buf = Buffer.create 64
+
+let pr_entry e =
+ let () = Buffer.clear entry_buf in
+ let ft = Format.formatter_of_buffer entry_buf in
+ let () = Gram.entry_print ft e in
+ str (Buffer.contents entry_buf)
+
+let pr_grammar = function
| "constr" | "operconstr" | "binder_constr" ->
- msgnl (str "Entry constr is");
- Gram.entry_print Pcoq.Constr.constr;
- msgnl (str "and lconstr is");
- Gram.entry_print Pcoq.Constr.lconstr;
- msgnl (str "where binder_constr is");
- Gram.entry_print Pcoq.Constr.binder_constr;
- msgnl (str "and operconstr is");
- Gram.entry_print Pcoq.Constr.operconstr;
+ str "Entry constr is" ++ fnl () ++
+ pr_entry Pcoq.Constr.constr ++
+ str "and lconstr is" ++ fnl () ++
+ pr_entry Pcoq.Constr.lconstr ++
+ str "where binder_constr is" ++ fnl () ++
+ pr_entry Pcoq.Constr.binder_constr ++
+ str "and operconstr is" ++ fnl () ++
+ pr_entry Pcoq.Constr.operconstr
| "pattern" ->
- Gram.entry_print Pcoq.Constr.pattern
+ pr_entry Pcoq.Constr.pattern
| "tactic" ->
- msgnl (str "Entry tactic_expr is");
- Gram.entry_print Pcoq.Tactic.tactic_expr;
- msgnl (str "Entry binder_tactic is");
- Gram.entry_print Pcoq.Tactic.binder_tactic;
- msgnl (str "Entry simple_tactic is");
- Gram.entry_print Pcoq.Tactic.simple_tactic;
+ str "Entry tactic_expr is" ++ fnl () ++
+ pr_entry Pcoq.Tactic.tactic_expr ++
+ str "Entry binder_tactic is" ++ fnl () ++
+ pr_entry Pcoq.Tactic.binder_tactic ++
+ str "Entry simple_tactic is" ++ fnl () ++
+ pr_entry Pcoq.Tactic.simple_tactic ++
+ str "Entry tactic_arg is" ++ fnl () ++
+ pr_entry Pcoq.Tactic.tactic_arg
| "vernac" ->
- msgnl (str "Entry vernac is");
- Gram.entry_print Pcoq.Vernac_.vernac;
- msgnl (str "Entry command is");
- Gram.entry_print Pcoq.Vernac_.command;
- msgnl (str "Entry syntax is");
- Gram.entry_print Pcoq.Vernac_.syntax;
- msgnl (str "Entry gallina is");
- Gram.entry_print Pcoq.Vernac_.gallina;
- msgnl (str "Entry gallina_ext is");
- Gram.entry_print Pcoq.Vernac_.gallina_ext;
+ str "Entry vernac is" ++ fnl () ++
+ pr_entry Pcoq.Vernac_.vernac ++
+ str "Entry command is" ++ fnl () ++
+ pr_entry Pcoq.Vernac_.command ++
+ str "Entry syntax is" ++ fnl () ++
+ pr_entry Pcoq.Vernac_.syntax ++
+ str "Entry gallina is" ++ fnl () ++
+ pr_entry Pcoq.Vernac_.gallina ++
+ str "Entry gallina_ext is" ++ fnl () ++
+ pr_entry Pcoq.Vernac_.gallina_ext
| _ -> error "Unknown or unprintable grammar entry."
(**********************************************************************)
(* Parse a format (every terminal starting with a letter or a single
quote (except a single quote alone) must be quoted) *)
-let parse_format (loc,str) =
+let parse_format ((loc, str) : lstring) =
let str = " "^str in
let l = String.length str in
let push_token a = function
| cur::l -> (a::cur)::l
| [] -> [[a]] in
let push_white n l =
- if n = 0 then l else push_token (UnpTerminal (String.make n ' ')) l in
+ if Int.equal n 0 then l else push_token (UnpTerminal (String.make n ' ')) l in
let close_box i b = function
| a::(_::_ as l) -> push_token (UnpBox (b,a)) l
| _ -> error "Non terminated box in format." in
let close_quotation i =
- if i < String.length str & str.[i] = '\'' & (i+1 = l or str.[i+1] = ' ')
+ if i < String.length str && str.[i] == '\'' && (Int.equal (i+1) l || str.[i+1] == ' ')
then i+1
else error "Incorrectly terminated quoted expression." in
let rec spaces n i =
- if i < String.length str & str.[i] = ' ' then spaces (n+1) (i+1)
+ if i < String.length str && str.[i] == ' ' then spaces (n+1) (i+1)
else n in
let rec nonspaces quoted n i =
- if i < String.length str & str.[i] <> ' ' then
- if str.[i] = '\'' & quoted &
- (i+1 >= String.length str or str.[i+1] = ' ')
- then if n=0 then error "Empty quoted token." else n
+ if i < String.length str && str.[i] != ' ' then
+ if str.[i] == '\'' && quoted &&
+ (i+1 >= String.length str || str.[i+1] == ' ')
+ then if Int.equal n 0 then error "Empty quoted token." else n
else nonspaces quoted (n+1) (i+1)
else
if quoted then error "Spaces are not allowed in (quoted) symbols."
@@ -177,7 +249,7 @@ let parse_format (loc,str) =
and parse_quoted n i =
if i < String.length str then match str.[i] with
(* Parse " // " *)
- | '/' when i <= String.length str & str.[i+1] = '/' ->
+ | '/' when i <= String.length str && str.[i+1] == '/' ->
(* We forget the useless n spaces... *)
push_token (UnpCut PpFnl)
(parse_token (close_quotation (i+2)))
@@ -192,7 +264,7 @@ let parse_format (loc,str) =
| '[' ->
if i <= String.length str then match str.[i+1] with
(* Parse " [h .. ", *)
- | 'h' when i+1 <= String.length str & str.[i+2] = 'v' ->
+ | 'h' when i+1 <= String.length str && str.[i+2] == 'v' ->
(parse_box (fun n -> PpHVB n) (i+3))
(* Parse " [v .. ", *)
| 'v' ->
@@ -211,7 +283,7 @@ let parse_format (loc,str) =
push_token (UnpTerminal (String.sub str (i-1) (n+2)))
(parse_token (close_quotation (i+n))))
else
- if n = 0 then []
+ if Int.equal n 0 then []
else error "Ending spaces non part of a format annotation."
and parse_box box i =
let n = spaces 0 i in
@@ -221,7 +293,7 @@ let parse_format (loc,str) =
let i = i+n in
if i < l then match str.[i] with
(* Parse a ' *)
- | '\'' when i+1 >= String.length str or str.[i+1] = ' ' ->
+ | '\'' when i+1 >= String.length str || str.[i+1] == ' ' ->
push_white (n-1) (push_token (UnpTerminal "'") (parse_token (i+1)))
(* Parse the beginning of a quoted expression *)
| '\'' ->
@@ -232,13 +304,15 @@ let parse_format (loc,str) =
else push_white n [[]]
in
try
- if str <> "" then match parse_token 0 with
+ if not (String.is_empty str) then match parse_token 0 with
| [l] -> l
| _ -> error "Box closed without being opened in format."
else
error "Empty format."
- with e when Errors.noncritical e ->
- Loc.raise loc e
+ with reraise ->
+ let (e, info) = Errors.push reraise in
+ let info = Loc.add_loc info loc in
+ iraise (e, info)
(***********************)
(* Analyzing notations *)
@@ -247,16 +321,16 @@ type symbol_token = WhiteSpace of int | String of string
let split_notation_string str =
let push_token beg i l =
- if beg = i then l else
+ if Int.equal beg i then l else
let s = String.sub str beg (i - beg) in
String s :: l
in
let push_whitespace beg i l =
- if beg = i then l else WhiteSpace (i-beg) :: l
+ if Int.equal beg i then l else WhiteSpace (i-beg) :: l
in
let rec loop beg i =
if i < String.length str then
- if str.[i] = ' ' then
+ if str.[i] == ' ' then
push_token beg i (loop_on_whitespace (i+1) (i+1))
else
loop beg (i+1)
@@ -264,7 +338,7 @@ let split_notation_string str =
push_token beg i []
and loop_on_whitespace beg i =
if i < String.length str then
- if str.[i] <> ' ' then
+ if str.[i] != ' ' then
push_whitespace beg i (loop i (i+1))
else
loop_on_whitespace beg (i+1)
@@ -281,25 +355,25 @@ let msg_expected_form_of_recursive_notation =
"In the notation, the special symbol \"..\" must occur in\na configuration of the form \"x symbs .. symbs y\"."
let rec find_pattern nt xl = function
- | Break n as x :: l, Break n' :: l' when n=n' ->
+ | Break n as x :: l, Break n' :: l' when Int.equal n n' ->
find_pattern nt (x::xl) (l,l')
- | Terminal s as x :: l, Terminal s' :: l' when s = s' ->
+ | Terminal s as x :: l, Terminal s' :: l' when String.equal s s' ->
find_pattern nt (x::xl) (l,l')
| [], NonTerminal x' :: l' ->
(out_nt nt,x',List.rev xl),l'
- | _, Terminal s :: _ | Terminal s :: _, _ ->
- error ("The token \""^s^"\" occurs on one side of \"..\" but not on the other side.")
| _, Break s :: _ | Break s :: _, _ ->
error ("A break occurs on one side of \"..\" but not on the other side.")
+ | _, Terminal s :: _ | Terminal s :: _, _ ->
+ error ("The token \""^s^"\" occurs on one side of \"..\" but not on the other side.")
| _, [] ->
error msg_expected_form_of_recursive_notation
| ((SProdList _ | NonTerminal _) :: _), _ | _, (SProdList _ :: _) ->
- anomaly "Only Terminal or Break expected on left, non-SProdList on right"
+ anomaly (Pp.str "Only Terminal or Break expected on left, non-SProdList on right")
let rec interp_list_parser hd = function
| [] -> [], List.rev hd
- | NonTerminal id :: tl when id = ldots_var ->
- if hd = [] then error msg_expected_form_of_recursive_notation;
+ | NonTerminal id :: tl when Id.equal id ldots_var ->
+ if List.is_empty hd then error msg_expected_form_of_recursive_notation;
let hd = List.rev hd in
let ((x,y,sl),tl') = find_pattern (List.hd hd) [] (List.tl hd,tl) in
let xyl,tl'' = interp_list_parser [] tl' in
@@ -307,7 +381,7 @@ let rec interp_list_parser hd = function
(* remove the second copy of it afterwards *)
(x,y)::xyl, SProdList (x,sl) :: tl''
| (Terminal _ | Break _) as s :: tl ->
- if hd = [] then
+ if List.is_empty hd then
let yl,tl' = interp_list_parser [] tl in
yl, s :: tl'
else
@@ -315,7 +389,7 @@ let rec interp_list_parser hd = function
| NonTerminal _ as x :: tl ->
let xyl,tl' = interp_list_parser [x] tl in
xyl, List.rev_append hd tl'
- | SProdList _ :: _ -> anomaly "Unexpected SProdList in interp_list_parser"
+ | SProdList _ :: _ -> anomaly (Pp.str "Unexpected SProdList in interp_list_parser")
(* Find non-terminal tokens of notation *)
@@ -323,26 +397,25 @@ 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 = is_ident x in
- if (n > 0 & norm) or (n > 2 & x.[0] = '\'') then "'"^x^"'"
+ let norm = Lexer.is_ident x in
+ if (n > 0 && norm) || (n > 2 && x.[0] == '\'') then "'"^x^"'"
else x
let rec raw_analyze_notation_tokens = function
| [] -> []
| String ".." :: sl -> NonTerminal ldots_var :: raw_analyze_notation_tokens sl
| String "_" :: _ -> error "_ must be quoted."
- | String x :: sl when is_ident x ->
- NonTerminal (Names.id_of_string x) :: raw_analyze_notation_tokens sl
+ | String x :: sl when Lexer.is_ident x ->
+ NonTerminal (Names.Id.of_string x) :: raw_analyze_notation_tokens sl
| String s :: sl ->
- Terminal (drop_simple_quotes s) :: raw_analyze_notation_tokens sl
+ Terminal (String.drop_simple_quotes s) :: raw_analyze_notation_tokens sl
| WhiteSpace n :: sl ->
Break n :: raw_analyze_notation_tokens sl
let is_numeral symbs =
match List.filter (function Break _ -> false | _ -> true) symbs with
| ([Terminal "-"; Terminal x] | [Terminal x]) ->
- (try let _ = Bigint.of_string x in true
- with e when Errors.noncritical e -> false)
+ (try let _ = Bigint.of_string x in true with Failure _ -> false)
| _ ->
false
@@ -350,9 +423,9 @@ let rec get_notation_vars = function
| [] -> []
| NonTerminal id :: sl ->
let vars = get_notation_vars sl in
- if id = ldots_var then vars else
- if List.mem id vars then
- error ("Variable "^string_of_id id^" occurs more than once.")
+ if Id.equal id ldots_var then vars else
+ if Id.List.mem id vars then
+ error ("Variable "^Id.to_string id^" occurs more than once.")
else
id::vars
| (Terminal _ | Break _) :: sl -> get_notation_vars sl
@@ -362,18 +435,15 @@ let analyze_notation_tokens l =
let l = raw_analyze_notation_tokens l in
let vars = get_notation_vars l in
let recvars,l = interp_list_parser [] l in
- recvars, list_subtract vars (List.map snd recvars), l
+ recvars, List.subtract Id.equal vars (List.map snd recvars), l
let error_not_same_scope x y =
- error ("Variables "^string_of_id x^" and "^string_of_id y^
+ error ("Variables "^Id.to_string x^" and "^Id.to_string y^
" must be in the same scope.")
(**********************************************************************)
(* Build pretty-printing rules *)
-type printing_precedence = int * parenRelation
-type parsing_precedence = int option
-
let prec_assoc = function
| RightA -> (L,E)
| LeftA -> (E,L)
@@ -382,7 +452,7 @@ let prec_assoc = function
let precedence_of_entry_type from = function
| ETConstr (NumLevel n,BorderProd (_,None)) -> n, Prec n
| ETConstr (NumLevel n,BorderProd (b,Some a)) ->
- n, let (lp,rp) = prec_assoc a in if b=Left then lp else rp
+ n, let (lp,rp) = prec_assoc a in if b == Left then lp else rp
| ETConstr (NumLevel n,InternalProd) -> n, Prec n
| ETConstr (NextLevel,_) -> from, L
| _ -> 0, E (* ?? *)
@@ -395,36 +465,50 @@ let precedence_of_entry_type from = function
(* "{ x } + { y }" : "{ x } / + { y }" *)
(* "< x , y > { z , t }" : "< x , / y > / { z , / t }" *)
+let starts_with_left_bracket s =
+ let l = String.length s in not (Int.equal l 0) &&
+ (s.[0] == '{' || s.[0] == '[' || s.[0] == '(')
+
+let ends_with_right_bracket s =
+ let l = String.length s in not (Int.equal l 0) &&
+ (s.[l-1] == '}' || s.[l-1] == ']' || s.[l-1] == ')')
+
let is_left_bracket s =
- let l = String.length s in l <> 0 &
- (s.[0] = '{' or s.[0] = '[' or s.[0] = '(')
+ starts_with_left_bracket s && not (ends_with_right_bracket s)
let is_right_bracket s =
- let l = String.length s in l <> 0 &
- (s.[l-1] = '}' or s.[l-1] = ']' or s.[l-1] = ')')
+ not (starts_with_left_bracket s) && ends_with_right_bracket s
let is_comma s =
- let l = String.length s in l <> 0 &
- (s.[0] = ',' or s.[0] = ';')
+ let l = String.length s in not (Int.equal l 0) &&
+ (s.[0] == ',' || s.[0] == ';')
let is_operator s =
- let l = String.length s in l <> 0 &
- (s.[0] = '+' or s.[0] = '*' or s.[0] = '=' or
- s.[0] = '-' or s.[0] = '/' or s.[0] = '<' or s.[0] = '>' or
- s.[0] = '@' or s.[0] = '\\' or s.[0] = '&' or s.[0] = '~' or s.[0] = '$')
+ let l = String.length s in not (Int.equal l 0) &&
+ (s.[0] == '+' || s.[0] == '*' || s.[0] == '=' ||
+ s.[0] == '-' || s.[0] == '/' || s.[0] == '<' || s.[0] == '>' ||
+ s.[0] == '@' || s.[0] == '\\' || s.[0] == '&' || s.[0] == '~' || s.[0] == '$')
-let is_prod_ident = function
- | Terminal s when is_letter s.[0] or s.[0] = '_' -> true
- | _ -> false
-
-let rec is_non_terminal = function
+let is_non_terminal = function
| NonTerminal _ | SProdList _ -> true
| _ -> false
+let is_next_non_terminal = function
+| [] -> false
+| pr :: _ -> is_non_terminal pr
+
+let is_next_terminal = function Terminal _ :: _ -> true | _ -> false
+
+let is_next_break = function Break _ :: _ -> true | _ -> false
+
let add_break n l = UnpCut (PpBrk(n,0)) :: l
+let add_break_if_none n = function
+ | ((UnpCut (PpBrk _) :: _) | []) as l -> l
+ | l -> UnpCut (PpBrk(n,0)) :: l
+
let check_open_binder isopen sl m =
- if isopen & sl <> [] then
+ 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
@@ -432,84 +516,93 @@ let check_open_binder isopen sl m =
(* Heuristics for building default printing rules *)
-type previous_prod_status = NoBreak | CanBreak
+let index_id id l = List.index Id.equal id l
let make_hunks etyps symbols from =
let vars,typs = List.split etyps in
- let rec make ws = function
+ let rec make = function
| NonTerminal m :: prods ->
- let i = list_index m vars in
+ let i = index_id m vars in
let _,prec = precedence_of_entry_type from (List.nth typs (i-1)) in
let u = UnpMetaVar (i,prec) in
- if prods <> [] && is_non_terminal (List.hd prods) then
- u :: add_break 1 (make CanBreak prods)
+ if is_next_non_terminal prods then
+ u :: add_break_if_none 1 (make prods)
else
- u :: make CanBreak prods
-
+ u :: make_with_space prods
| Terminal s :: prods when List.exists is_non_terminal prods ->
- if is_comma s then
- UnpTerminal s :: add_break 1 (make NoBreak prods)
- else if is_right_bracket s then
- UnpTerminal s :: add_break 0 (make NoBreak prods)
- else if is_left_bracket s then
- if ws = CanBreak then
- add_break 1 (UnpTerminal s :: make CanBreak prods)
- else
- UnpTerminal s :: make CanBreak prods
- else if is_operator s then
- if ws = CanBreak then
- UnpTerminal (" "^s) :: add_break 1 (make NoBreak prods)
- else
- UnpTerminal s :: add_break 1 (make NoBreak prods)
- else if is_ident_tail s.[String.length s - 1] then
- let sep = if is_prod_ident (List.hd prods) then "" else " " in
- if ws = CanBreak then
- add_break 1 (UnpTerminal (s^sep) :: make CanBreak prods)
- else
- UnpTerminal (s^sep) :: make CanBreak prods
- else if ws = CanBreak then
- add_break 1 (UnpTerminal (s^" ") :: make CanBreak prods)
- else
- UnpTerminal s :: make CanBreak prods
+ if (is_comma s || is_operator s) then
+ (* Always a breakable space after comma or separator *)
+ UnpTerminal s :: add_break_if_none 1 (make prods)
+ else if is_right_bracket s && is_next_terminal prods then
+ (* Always no space after right bracked, but possibly a break *)
+ UnpTerminal s :: add_break_if_none 0 (make prods)
+ else if is_left_bracket s && is_next_non_terminal prods then
+ UnpTerminal s :: make prods
+ else if not (is_next_break prods) then
+ (* Add rigid space, no break, unless user asked for something *)
+ UnpTerminal (s^" ") :: make prods
+ else
+ (* Rely on user spaces *)
+ UnpTerminal s :: make prods
| Terminal s :: prods ->
- if is_right_bracket s then
- UnpTerminal s :: make NoBreak prods
- else if ws = CanBreak then
- add_break 1 (UnpTerminal s :: make NoBreak prods)
- else
- UnpTerminal s :: make NoBreak prods
+ (* Separate but do not cut a trailing sequence of terminal *)
+ (match prods with
+ | Terminal _ :: _ -> UnpTerminal (s^" ") :: make prods
+ | _ -> UnpTerminal s :: make prods)
| Break n :: prods ->
- add_break n (make NoBreak prods)
+ add_break n (make prods)
| SProdList (m,sl) :: prods ->
- let i = list_index m vars in
+ let i = index_id m vars in
let typ = List.nth typs (i-1) in
let _,prec = precedence_of_entry_type from typ in
let sl' =
(* If no separator: add a break *)
- if sl = [] then add_break 1 []
+ if List.is_empty sl then add_break 1 []
(* We add NonTerminal for simulation but remove it afterwards *)
- else snd (list_sep_last (make NoBreak (sl@[NonTerminal m]))) in
+ else snd (List.sep_last (make (sl@[NonTerminal m]))) in
let hunk = match typ with
| ETConstr _ -> UnpListMetaVar (i,prec,sl')
| ETBinder isopen ->
check_open_binder isopen sl m;
UnpBinderListMetaVar (i,isopen,sl')
| _ -> assert false in
- hunk :: make CanBreak prods
+ hunk :: make_with_space prods
| [] -> []
- in make NoBreak symbols
+ and make_with_space prods =
+ match prods with
+ | Terminal s' :: prods'->
+ if is_operator s' then
+ (* A rigid space before operator and a breakable after *)
+ UnpTerminal (" "^s') :: add_break_if_none 1 (make prods')
+ else if is_comma s' then
+ (* No space whatsoever before comma *)
+ make prods
+ else if is_right_bracket s' then
+ make prods
+ else
+ (* A breakable space between any other two terminals *)
+ add_break_if_none 1 (make prods)
+ | (NonTerminal _ | SProdList _) :: _ ->
+ (* A breakable space before a non-terminal *)
+ add_break_if_none 1 (make prods)
+ | Break _ :: _ ->
+ (* Rely on user wish *)
+ make prods
+ | [] -> []
+
+ in make symbols
(* Build default printing rules from explicit format *)
let error_format () = error "The format does not match the notation."
let rec split_format_at_ldots hd = function
- | UnpTerminal s :: fmt when s = string_of_id ldots_var -> List.rev hd, fmt
+ | UnpTerminal s :: fmt when String.equal s (Id.to_string ldots_var) -> List.rev hd, fmt
| u :: fmt ->
check_no_ldots_in_box u;
split_format_at_ldots (u::hd) fmt
@@ -535,7 +628,7 @@ let read_recursive_format sl fmt =
let sl = skip_var_in_recursive_format fmt in
try split_format_at_ldots [] sl with Exit -> error_format () in
let rec get_tail = function
- | a :: sepfmt, b :: fmt when a = b -> get_tail (sepfmt, fmt)
+ | a :: sepfmt, b :: fmt when Pervasives.(=) a b -> get_tail (sepfmt, fmt) (* FIXME *)
| [], tail -> skip_var_in_recursive_format tail
| _ -> error "The format is not the same on the right and left hand side of the special token \"..\"." in
let slfmt, fmt = get_head fmt in
@@ -544,13 +637,13 @@ let read_recursive_format sl fmt =
let hunks_of_format (from,(vars,typs)) symfmt =
let rec aux = function
| symbs, (UnpTerminal s' as u) :: fmt
- when s' = String.make (String.length s') ' ' ->
+ when String.equal s' (String.make (String.length s') ' ') ->
let symbs, l = aux (symbs,fmt) in symbs, u :: l
| Terminal s :: symbs, (UnpTerminal s') :: fmt
- when s = drop_simple_quotes s' ->
+ when String.equal s (String.drop_simple_quotes s') ->
let symbs, l = aux (symbs,fmt) in symbs, UnpTerminal s :: l
- | NonTerminal s :: symbs, UnpTerminal s' :: fmt when s = id_of_string s' ->
- let i = list_index s vars in
+ | NonTerminal s :: symbs, UnpTerminal s' :: fmt when Id.equal s (Id.of_string s') ->
+ let i = index_id s vars in
let _,prec = precedence_of_entry_type from (List.nth typs (i-1)) in
let symbs, l = aux (symbs,fmt) in symbs, UnpMetaVar (i,prec) :: l
| symbs, UnpBox (a,b) :: fmt ->
@@ -560,12 +653,12 @@ let hunks_of_format (from,(vars,typs)) symfmt =
| symbs, (UnpCut _ as u) :: fmt ->
let symbs, l = aux (symbs,fmt) in symbs, u :: l
| SProdList (m,sl) :: symbs, fmt ->
- let i = list_index m vars in
+ let i = index_id m vars in
let typ = List.nth typs (i-1) in
let _,prec = precedence_of_entry_type from typ in
let slfmt,fmt = read_recursive_format sl fmt in
let sl, slfmt = aux (sl,slfmt) in
- if sl <> [] then error_format ();
+ if not (List.is_empty sl) then error_format ();
let symbs, l = aux (symbs,fmt) in
let hunk = match typ with
| ETConstr _ -> UnpListMetaVar (i,prec,slfmt)
@@ -594,7 +687,7 @@ 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 ->
- message ("Identifier '"^k^"' now a keyword");
+ Flags.if_verbose msg_info (strbrk ("Identifier '"^k^"' now a keyword"));
Lexer.add_keyword k;
n1 :: GramConstrTerminal(KEYWORD k) :: define_keywords_aux l
| n :: l -> n :: define_keywords_aux l
@@ -603,7 +696,7 @@ let rec define_keywords_aux = function
(* Ensure that IDENT articulation terminal symbols are keywords *)
let define_keywords = function
| GramConstrTerminal(IDENT k)::l ->
- message ("Identifier '"^k^"' now a keyword");
+ Flags.if_verbose msg_info (strbrk ("Identifier '"^k^"' now a keyword"));
Lexer.add_keyword k;
GramConstrTerminal(KEYWORD k) :: define_keywords_aux l
| l -> define_keywords_aux l
@@ -613,7 +706,7 @@ let distribute a ll = List.map (fun l -> a @ l) ll
(* Expand LIST1(t,sep) into the combination of t and t;sep;LIST1(t,sep)
as many times as expected in [n] argument *)
let rec expand_list_rule typ tkl x n i hds ll =
- if i = n then
+ if Int.equal i n then
let hds =
GramConstrListMark (n,true) :: hds
@ [GramConstrNonTerminal (ETConstrList (typ,tkl), Some x)] in
@@ -633,14 +726,14 @@ let make_production etyps symbols =
let typ = List.assoc m etyps in
distribute [GramConstrNonTerminal (typ, Some m)] ll
| Terminal s ->
- distribute [GramConstrTerminal (terminal s)] ll
+ distribute [GramConstrTerminal (Lexer.terminal s)] ll
| Break _ ->
ll
| SProdList (x,sl) ->
let tkl = List.flatten
- (List.map (function Terminal s -> [terminal s]
+ (List.map (function Terminal s -> [Lexer.terminal s]
| Break _ -> []
- | _ -> anomaly "Found a non terminal token in recursive notation separator") sl) in
+ | _ -> anomaly (Pp.str "Found a non terminal token in recursive notation separator")) sl) in
match List.assoc x etyps with
| ETConstr typ -> expand_list_rule typ tkl x 1 0 [] ll
| ETBinder o ->
@@ -654,7 +747,7 @@ let make_production etyps symbols =
let rec find_symbols c_current c_next c_last = function
| [] -> []
| NonTerminal id :: sl ->
- let prec = if sl <> [] then c_current else c_last in
+ let prec = if not (List.is_empty sl) then c_current else c_last in
(id, prec) :: (find_symbols c_next c_next c_last sl)
| Terminal s :: sl -> find_symbols c_next c_next c_last sl
| Break n :: sl -> find_symbols c_current c_next c_last sl
@@ -676,10 +769,10 @@ let recompute_assoc typs =
(* Registration of syntax extensions (parsing/printing, no interpretation)*)
let pr_arg_level from = function
- | (n,L) when n=from -> str "at next level"
+ | (n,L) when Int.equal n from -> str "at next level"
| (n,E) -> str "at level " ++ int n
| (n,L) -> str "at level below " ++ int n
- | (n,Prec m) when m=n -> str "at level " ++ int n
+ | (n,Prec m) when Int.equal m n -> str "at level " ++ int n
| (n,_) -> str "Unknown level"
let pr_level ntn (from,args) =
@@ -693,42 +786,51 @@ let error_incompatible_level ntn oldprec prec =
spc() ++ str "while it is now required to be" ++ spc() ++
pr_level ntn prec ++ str ".")
-let cache_one_syntax_extension (typs,prec,ntn,gr,pp) =
+type syntax_extension = {
+ synext_level : Notation.level;
+ synext_notation : notation;
+ synext_notgram : notation_grammar;
+ synext_unparsing : unparsing list;
+ synext_extra : (string * string) list;
+}
+
+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
try
let oldprec = Notation.level_of_notation ntn in
- if prec <> oldprec then error_incompatible_level ntn oldprec prec
+ 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 *)
- Egrammar.extend_grammar (Egrammar.Notation (prec,typs,gr));
+ Egramcoq.extend_constr_grammar prec se.synext_notgram;
(* Declare the printing rule *)
- Notation.declare_notation_printing_rule ntn (pp,fst prec)
+ Notation.declare_notation_printing_rule ntn
+ ~extra:se.synext_extra (se.synext_unparsing, fst prec)
-let cache_syntax_extension (_,(_,sy_rules)) =
- List.iter cache_one_syntax_extension sy_rules
+let cache_syntax_extension (_, (_, sy)) =
+ List.iter cache_one_syntax_extension sy
let subst_parsing_rule subst x = x
let subst_printing_rule subst x = x
-let subst_syntax_extension (subst,(local,sy)) =
- (local, List.map (fun (typs,prec,ntn,gr,pp) ->
- (typs,prec,ntn,subst_parsing_rule subst gr,subst_printing_rule subst pp))
- sy)
+let subst_syntax_extension (subst, (local, sy)) =
+ let map sy = { sy with
+ synext_notgram = subst_parsing_rule subst sy.synext_notgram;
+ synext_unparsing = subst_printing_rule subst sy.synext_unparsing;
+ } in
+ (local, List.map map sy)
-let classify_syntax_definition (local,_ as o) =
+let classify_syntax_definition (local, _ as o) =
if local then Dispose else Substitute o
-type syntax_extension_obj =
- bool *
- (notation_var_internalization_type list * Notation.level *
- notation * notation_grammar * unparsing list)
- list
-
let inSyntaxExtension : syntax_extension_obj -> obj =
declare_object {(default_object "SYNTAX-EXTENSION") with
- open_function = (fun i o -> if i=1 then cache_syntax_extension o);
+ open_function = (fun i o -> if Int.equal i 1 then cache_syntax_extension o);
cache_function = cache_syntax_extension;
subst_function = subst_syntax_extension;
classify_function = classify_syntax_definition}
@@ -740,41 +842,49 @@ let inSyntaxExtension : syntax_extension_obj -> obj =
let interp_modifiers modl =
let onlyparsing = ref false in
- let rec interp assoc level etyps format = function
+ let rec interp assoc level etyps format extra = function
| [] ->
- (assoc,level,etyps,!onlyparsing,format)
+ (assoc,level,etyps,!onlyparsing,format,extra)
| SetEntryType (s,typ) :: l ->
- let id = id_of_string s in
- if List.mem_assoc id etyps then
+ let id = Id.of_string s in
+ if Id.List.mem_assoc id etyps then
error (s^" is already assigned to an entry or constr level.");
- interp assoc level ((id,typ)::etyps) format l
+ interp assoc level ((id,typ)::etyps) format extra l
| SetItemLevel ([],n) :: l ->
- interp assoc level etyps format l
+ interp assoc level etyps format extra l
| SetItemLevel (s::idl,n) :: l ->
- let id = id_of_string s in
- if List.mem_assoc id etyps then
+ let id = Id.of_string s in
+ if Id.List.mem_assoc id etyps then
error (s^" is already assigned to an entry or constr level.");
let typ = ETConstr (n,()) in
- interp assoc level ((id,typ)::etyps) format (SetItemLevel (idl,n)::l)
+ interp assoc level ((id,typ)::etyps) format extra (SetItemLevel (idl,n)::l)
| SetLevel n :: l ->
- if level <> None then error "A level is given more than once.";
- interp assoc (Some n) etyps format l
+ if not (Option.is_empty level) then error "A level is given more than once.";
+ interp assoc (Some n) etyps format extra l
| SetAssoc a :: l ->
- if assoc <> None then error"An associativity is given more than once.";
- interp (Some a) level etyps format 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 ->
onlyparsing := true;
- interp assoc level etyps format l
- | SetFormat s :: l ->
- if format <> None then error "A format is given more than once.";
- interp assoc level etyps (Some s) l
- in interp None None [] None modl
+ 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
+ | SetFormat (k,(_,s)) :: l ->
+ interp assoc level etyps format ((k,s) :: extra) l
+ in interp None None [] None [] modl
let check_infix_modifiers modifiers =
- let (assoc,level,t,b,fmt) = interp_modifiers modifiers in
- if t <> [] then
+ let (assoc,level,t,b,fmt,extra) = interp_modifiers modifiers in
+ if not (List.is_empty t) then
error "Explicit entry level or type unexpected in infix notation."
+let check_useless_entry_types recvars mainvars etyps =
+ let vars = let (l1,l2) = List.split recvars in l1@l2@mainvars in
+ match List.filter (fun (x,etyp) -> not (List.mem x vars)) etyps with
+ | (x,_)::_ -> error (Id.to_string x ^ " is unbound in the notation.")
+ | _ -> ()
+
let no_syntax_modifiers = function
| [] | [SetOnlyParsing _] -> true
| _ -> false
@@ -805,7 +915,7 @@ let join_auxiliary_recursive_types recvars etyps =
| None, None -> typs
| Some _, None -> typs
| None, Some ytyp -> (x,ytyp)::typs
- | Some xtyp, Some ytyp when xtyp = ytyp -> typs
+ | Some xtyp, Some ytyp when Pervasives.(=) xtyp ytyp -> typs (* FIXME *)
| Some xtyp, Some ytyp ->
errorlabstrm ""
(strbrk "In " ++ pr_id x ++ str " .. " ++ pr_id y ++
@@ -821,12 +931,12 @@ let internalization_type_of_entry_type = function
| ETBinderList _ | ETConstrList _ -> assert false
let set_internalization_type typs =
- List.map (down_snd internalization_type_of_entry_type) typs
+ List.map (fun (_, e) -> internalization_type_of_entry_type e) typs
let make_internalization_vars recvars mainvars typs =
let maintyps = List.combine mainvars typs in
let extratyps = List.map (fun (x,y) -> (y,List.assoc x maintyps)) recvars in
- maintyps@extratyps
+ maintyps @ extratyps
let make_interpretation_type isrec = function
| NtnInternTypeConstr when isrec -> NtnTypeConstrList
@@ -835,14 +945,21 @@ let make_interpretation_type isrec = function
| NtnInternTypeBinder -> error "Type not allowed in recursive notation."
let make_interpretation_vars recvars allvars =
- List.iter (fun (x,y) ->
- if fst (List.assoc x allvars) <> fst (List.assoc y allvars) then
- error_not_same_scope x y) recvars;
+ let eq_subscope (sc1, l1) (sc2, l2) =
+ Option.equal String.equal sc1 sc2 &&
+ List.equal String.equal l1 l2
+ in
+ let check (x, y) =
+ let (scope1, _) = Id.Map.find x allvars in
+ let (scope2, _) = Id.Map.find y allvars in
+ if not (eq_subscope scope1 scope2) then error_not_same_scope x y
+ in
+ let () = List.iter check recvars in
let useless_recvars = List.map snd recvars in
let mainvars =
- List.filter (fun (x,_) -> not (List.mem x useless_recvars)) allvars in
- List.map (fun (x,(sc,typ)) ->
- (x,(sc,make_interpretation_type (List.mem_assoc x recvars) typ))) 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
let check_rule_productivity l =
if List.for_all (function NonTerminal _ -> true | _ -> false) l then
@@ -850,9 +967,17 @@ let check_rule_productivity l =
if (match l with SProdList _ :: _ -> true | _ -> false) then
error "A recursive notation must start with at least one symbol."
-let is_not_printable = function
- | AVar _ -> msg_warn "This notation will not be used for printing as it is bound to a \nsingle variable"; true
- | _ -> false
+let is_not_printable onlyparse noninjective = 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
+ 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
+ else onlyparse
let find_precedence lev etyps symbols =
match symbols with
@@ -861,31 +986,32 @@ let find_precedence lev etyps symbols =
| ETConstr _ ->
error "The level of the leftmost non-terminal cannot be changed."
| ETName | ETBigint | ETReference ->
- if lev = None then
- ([msgnl,str "Setting notation at level 0."],0)
- else
- if lev <> Some 0 then
+ begin match lev with
+ | None ->
+ ([msg_info,strbrk "Setting notation at level 0."],0)
+ | Some 0 ->
+ ([],0)
+ | _ ->
error "A notation starting with an atomic expression must be at level 0."
- else
- ([],0)
+ end
| ETPattern | ETBinder _ | ETOther _ -> (* Give a default ? *)
- if lev = None then
+ if Option.is_empty lev then
error "Need an explicit level."
else [],Option.get lev
| ETConstrList _ | ETBinderList _ ->
assert false (* internally used in grammar only *)
with Not_found ->
- if lev = None then
+ 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)
+ (match List.last symbols with Terminal _ -> true |_ -> false)
->
- if lev = None then
- ([msgnl,str "Setting notation at level 0."], 0)
+ if Option.is_empty lev then
+ ([msg_info,strbrk "Setting notation at level 0."], 0)
else [],Option.get lev
| _ ->
- if lev = None then error "Cannot determine the level.";
+ if Option.is_empty lev then error "Cannot determine the level.";
[],Option.get lev
let check_curly_brackets_notation_exists () =
@@ -908,9 +1034,9 @@ let remove_curly_brackets l =
let br',next' = skip_break [] l' in
(match next' with
| Terminal "}" as t2 :: l'' as l1 ->
- if l <> l0 or l' <> l1 then
- msg_warn "Skipping spaces inside curly brackets";
- if deb & 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
+ msg_warning (strbrk "Skipping spaces inside curly brackets");
+ if deb && List.is_empty l'' then [t1;x;t2] else begin
check_curly_brackets_notation_exists ();
x :: aux false l''
end
@@ -919,14 +1045,15 @@ let remove_curly_brackets l =
| x :: l -> x :: aux false l
in aux true l
-let compute_syntax_data (df,modifiers) =
- let (assoc,n,etyps,onlyparse,fmt) = interp_modifiers modifiers in
+let compute_syntax_data df modifiers =
+ let (assoc,n,etyps,onlyparse,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
+ let _ = check_useless_entry_types recvars mainvars etyps in
let ntn_for_interp = make_notation_key symbols in
let symbols' = remove_curly_brackets symbols in
- let need_squash = (symbols <> symbols') in
+ let need_squash = not (List.equal Notation.symbol_eq symbols symbols') in
let ntn_for_grammar = make_notation_key symbols' in
check_rule_productivity symbols';
let msgs,n = find_precedence n etyps symbols' in
@@ -947,45 +1074,52 @@ let compute_syntax_data (df,modifiers) =
let df' = ((Lib.library_dp(),Lib.current_dirpath true),df) in
let i_data = (onlyparse,recvars,mainvars,(ntn_for_interp,df')) in
(* Return relevant data for interpretation and for parsing/printing *)
- (msgs,i_data,i_typs,sy_fulldata)
+ (msgs,i_data,i_typs,sy_fulldata,extra)
-let compute_pure_syntax_data (df,mods) =
- let (msgs,(onlyparse,_,_,_),_,sy_data) = compute_syntax_data (df,mods) in
+let compute_pure_syntax_data df mods =
+ let (msgs,(onlyparse,_,_,_),_,sy_data,extra) = compute_syntax_data df mods in
let msgs =
if onlyparse then
(msg_warning,
- str "The only parsing modifier has no effect in Reserved Notation.")::msgs
+ strbrk "The only parsing modifier has no effect in Reserved Notation.")::msgs
else msgs in
- msgs, sy_data
+ msgs, sy_data, extra
(**********************************************************************)
(* Registration of notations interpretation *)
-let load_notation _ (_,(_,scope,pat,onlyparse,_)) =
- Option.iter Notation.declare_scope scope
-
-let open_notation i (_,(_,scope,pat,onlyparse,(ntn,df))) =
- if i=1 & not (Notation.exists_notation_in_scope scope ntn pat) then begin
+type notation_obj = {
+ notobj_local : bool;
+ notobj_scope : scope_name option;
+ notobj_interp : interpretation;
+ notobj_onlyparse : bool;
+ notobj_notation : notation * notation_location;
+}
+
+let load_notation _ (_, nobj) =
+ Option.iter Notation.declare_scope nobj.notobj_scope
+
+let open_notation i (_, nobj) =
+ let scope = nobj.notobj_scope in
+ let (ntn, df) = nobj.notobj_notation in
+ let pat = nobj.notobj_interp in
+ if Int.equal i 1 then begin
(* Declare the interpretation *)
Notation.declare_notation_interpretation ntn scope pat df;
(* Declare the uninterpretation *)
- if not onlyparse then
- Notation.declare_uninterpretation (NotationRule (scope,ntn)) pat
+ if not nobj.notobj_onlyparse then
+ Notation.declare_uninterpretation (NotationRule (scope, ntn)) pat
end
let cache_notation o =
load_notation 1 o;
open_notation 1 o
-let subst_notation (subst,(lc,scope,pat,b,ndf)) =
- (lc,scope,subst_interpretation subst pat,b,ndf)
-
-let classify_notation (local,_,_,_,_ as o) =
- if local then Dispose else Substitute o
+let subst_notation (subst, nobj) =
+ { nobj with notobj_interp = subst_interpretation subst nobj.notobj_interp; }
-type notation_obj =
- bool * scope_name option * interpretation * bool *
- (notation * notation_location)
+let classify_notation nobj =
+ if nobj.notobj_local then Dispose else Substitute nobj
let inNotation : notation_obj -> obj =
declare_object {(default_object "NOTATION") with
@@ -998,9 +1132,12 @@ let inNotation : notation_obj -> obj =
(**********************************************************************)
let with_lib_stk_protection f x =
- let fs = Lib.freeze () in
+ let fs = Lib.freeze `No in
try let a = f x in Lib.unfreeze fs; a
- with reraise -> Lib.unfreeze fs; raise reraise
+ with reraise ->
+ let reraise = Errors.push reraise in
+ let () = Lib.unfreeze fs in
+ iraise reraise
let with_syntax_protection f x =
with_lib_stk_protection
@@ -1011,11 +1148,11 @@ let with_syntax_protection f x =
(* Recovering existing syntax *)
let contract_notation ntn =
- if ntn = "{ _ }" then ntn else
+ if String.equal ntn "{ _ }" then ntn else
let rec aux ntn i =
if i <= String.length ntn - 5 then
let ntn' =
- if String.sub ntn i 5 = "{ _ }" then
+ if String.is_sub "{ _ }" ntn i then
String.sub ntn 0 i ^ "_" ^
String.sub ntn (i+5) (String.length ntn -i-5)
else ntn in
@@ -1029,58 +1166,94 @@ let recover_syntax ntn =
try
let prec = Notation.level_of_notation ntn in
let pp_rule,_ = Notation.find_notation_printing_rule ntn in
- let typs,pa_rule = Egrammar.recover_notation_grammar ntn prec in
- (typs,prec,ntn,pa_rule,pp_rule)
+ let pp_extra_rules = Notation.find_notation_extra_printing_rules ntn in
+ let pa_rule = Egramcoq.recover_constr_grammar ntn prec in
+ { synext_level = prec;
+ synext_notation = ntn;
+ synext_notgram = pa_rule;
+ synext_unparsing = pp_rule;
+ synext_extra = pp_extra_rules }
with Not_found ->
raise NoSyntaxRule
-let recover_squash_syntax () = recover_syntax "{ _ }"
+let recover_squash_syntax sy =
+ let sq = recover_syntax "{ _ }" in
+ [sy; sq]
let recover_notation_syntax rawntn =
let ntn = contract_notation rawntn in
- let (typs,_,_,_,_ as sy_rule) = recover_syntax ntn in
- let need_squash = ntn<>rawntn in
- typs,if need_squash then [sy_rule; recover_squash_syntax ()] else [sy_rule]
+ 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
(**********************************************************************)
(* Main entry point for building parsing and printing rules *)
-let make_pa_rule (n,typs,symbols,_) ntn =
+let make_pa_rule i_typs (n,typs,symbols,_) ntn =
let assoc = recompute_assoc typs in
let prod = make_production typs symbols in
- (n,assoc,ntn,prod)
+ { notgram_level = n;
+ notgram_assoc = assoc;
+ notgram_notation = ntn;
+ notgram_prods = prod;
+ notgram_typs = i_typs; }
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)
+ | 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) =
- let pa_rule = make_pa_rule sy_data ntn in
+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 pp_rule = make_pp_rule sy_data in
- let sy_rule = (i_typs,prec,ntn,pa_rule,pp_rule) in
+ let sy = {
+ synext_level = prec;
+ synext_notation = ntn;
+ synext_notgram = pa_rule;
+ synext_unparsing = pp_rule;
+ synext_extra = extra;
+ } in
(* By construction, the rule for "{ _ }" is declared, but we need to
redeclare it because the file where it is declared needs not be open
when the current file opens (especially in presence of -nois) *)
- if need_squash then [sy_rule; recover_squash_syntax ()] else [sy_rule]
+ if need_squash then recover_squash_syntax sy else [sy]
(**********************************************************************)
(* Main functions about notations *)
+let to_map l =
+ let fold accu (x, v) = Id.Map.add x v accu in
+ List.fold_left fold Id.Map.empty l
+
let add_notation_in_scope local df c mods scope =
- let (msgs,i_data,i_typs,sy_data) = compute_syntax_data (df,mods) in
+ 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 in
+ let sy_rules = make_syntax_rules sy_data extra in
(* Prepare the interpretation *)
- let (onlyparse,recvars,mainvars,df') = i_data in
+ let (onlyparse, recvars,mainvars, df') = i_data in
let i_vars = make_internalization_vars recvars mainvars i_typs in
- let (acvars,ac) = interp_aconstr i_vars recvars c in
- let a = (make_interpretation_vars recvars acvars,ac) in
- let onlyparse = onlyparse or is_not_printable ac 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 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 notation = {
+ notobj_local = local;
+ notobj_scope = scope;
+ notobj_interp = (List.map_filter map i_vars, ac);
+ (** Order is important here! *)
+ notobj_onlyparse = onlyparse;
+ notobj_notation = df';
+ } in
(* Ready to change the global state *)
Flags.if_verbose (List.iter (fun (f,x) -> f x)) msgs;
- Lib.add_anonymous_leaf (inSyntaxExtension(local,sy_rules));
- Lib.add_anonymous_leaf (inNotation (local,scope,a,onlyparse,df'));
+ Lib.add_anonymous_leaf (inSyntaxExtension (local, sy_rules));
+ Lib.add_anonymous_leaf (inNotation notation);
df'
let add_notation_interpretation_core local df ?(impls=empty_internalization_env) c scope onlyparse =
@@ -1095,17 +1268,31 @@ let add_notation_interpretation_core local df ?(impls=empty_internalization_env)
let path = (Lib.library_dp(),Lib.current_dirpath true) in
let df' = (make_notation_key symbs,(path,df)) in
let i_vars = make_internalization_vars recvars mainvars i_typs in
- let (acvars,ac) = interp_aconstr ~impls i_vars recvars c in
- let a = (make_interpretation_vars recvars acvars,ac) in
- let onlyparse = onlyparse or is_not_printable ac in
- Lib.add_anonymous_leaf (inNotation (local,scope,a,onlyparse,df'));
+ 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 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 notation = {
+ notobj_local = local;
+ notobj_scope = scope;
+ notobj_interp = (List.map_filter map i_vars, ac);
+ (** Order is important here! *)
+ notobj_onlyparse = onlyparse;
+ notobj_notation = df';
+ } in
+ Lib.add_anonymous_leaf (inNotation notation);
df'
(* Notations without interpretation (Reserved Notation) *)
let add_syntax_extension local ((loc,df),mods) =
- let msgs,sy_data = compute_pure_syntax_data (df,mods) in
- let sy_rules = make_syntax_rules sy_data in
+ let msgs, sy_data, extra = compute_pure_syntax_data df mods in
+ let sy_rules = make_syntax_rules sy_data extra in
Flags.if_verbose (List.iter (fun (f,x) -> f x)) msgs;
Lib.add_anonymous_leaf (inSyntaxExtension(local,sy_rules))
@@ -1139,9 +1326,16 @@ let add_notation local c ((loc,df),modifiers) sc =
in
Dumpglob.dump_notation (loc,df') sc true
+let add_notation_extra_printing_rule df k v =
+ let notk =
+ let dfs = split_notation_string df in
+ let _,_, symbs = analyze_notation_tokens dfs in
+ make_notation_key symbs in
+ Notation.add_notation_extra_printing_rule notk k v
+
(* Infix notations *)
-let inject_var x = CRef (Ident (dummy_loc, id_of_string x))
+let inject_var x = CRef (Ident (Loc.ghost, Id.of_string x),None)
let add_infix local ((loc,inf),modifiers) pr sc =
check_infix_modifiers modifiers;
@@ -1154,16 +1348,16 @@ let add_infix local ((loc,inf),modifiers) pr sc =
(**********************************************************************)
(* Delimiters and classes bound to scopes *)
-type scope_command = ScopeDelim of string | ScopeClasses of Classops.cl_typ
+type scope_command = ScopeDelim of string | ScopeClasses of scope_class list
let load_scope_command _ (_,(scope,dlm)) =
Notation.declare_scope scope
let open_scope_command i (_,(scope,o)) =
- if i=1 then
+ if Int.equal i 1 then
match o with
| ScopeDelim dlm -> Notation.declare_delimiters scope dlm
- | ScopeClasses cl -> Notation.declare_class_scope scope cl
+ | ScopeClasses cl -> List.iter (Notation.declare_scope_class scope) cl
let cache_scope_command o =
load_scope_command 1 o;
@@ -1171,7 +1365,10 @@ let cache_scope_command o =
let subst_scope_command (subst,(scope,o as x)) = match o with
| ScopeClasses cl ->
- let cl' = Classops.subst_cl_typ subst cl in if cl'==cl then x else
+ let cl' = List.map_filter (subst_scope_class subst) cl in
+ let cl' =
+ if List.for_all2eq (==) cl cl' then cl
+ else cl' in
scope, ScopeClasses cl'
| _ -> x
@@ -1192,19 +1389,28 @@ let add_class_scope scope cl =
(* Check if abbreviation to a name and avoid early insertion of
maximal implicit arguments *)
let try_interp_name_alias = function
- | [], CRef ref -> intern_reference ref
+ | [], CRef (ref,_) -> intern_reference ref
| _ -> raise Not_found
let add_syntactic_definition ident (vars,c) local onlyparse =
+ let nonprintable = ref false in
let vars,pat =
- try [], ARef (try_interp_name_alias (vars,c))
+ try [], NRef (try_interp_name_alias (vars,c))
with Not_found ->
- let i_vars = List.map (fun id -> (id,NtnInternTypeConstr)) vars in
- let vars,pat = interp_aconstr i_vars [] c in
- List.map (fun (id,(sc,kind)) -> (id,sc)) vars, pat
+ let fold accu id = Id.Map.add id NtnInternTypeConstr accu in
+ let i_vars = List.fold_left fold Id.Map.empty vars in
+ let nenv = {
+ ninterp_var_type = i_vars;
+ ninterp_rec_vars = Id.Map.empty;
+ 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
+ List.map map vars, pat
in
let onlyparse = match onlyparse with
- | None when (is_not_printable pat) -> Some Flags.Current
+ | None when (is_not_printable false !nonprintable pat) -> Some Flags.Current
| p -> p
in
Syntax_def.declare_syntactic_definition local ident onlyparse (vars,pat)
diff --git a/toplevel/metasyntax.mli b/toplevel/metasyntax.mli
index f3bdf833..38a37757 100644
--- a/toplevel/metasyntax.mli
+++ b/toplevel/metasyntax.mli
@@ -1,27 +1,30 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Util
open Names
-open Libnames
-open Ppextend
-open Extend
open Tacexpr
open Vernacexpr
open Notation
-open Topconstr
+open Constrexpr
+open Notation_term
val add_token_obj : string -> unit
(** Adding a tactic notation in the environment *)
val add_tactic_notation :
- int * grammar_tactic_prod_item_expr list * raw_tactic_expr -> unit
+ 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*)
@@ -31,10 +34,12 @@ val add_infix : locality_flag -> (lstring * syntax_modifier list) ->
val add_notation : locality_flag -> constr_expr ->
(lstring * syntax_modifier list) -> scope_name option -> unit
+val add_notation_extra_printing_rule : string -> string -> string -> unit
+
(** Declaring delimiter keys and default scopes *)
val add_delimiters : scope_name -> string -> unit
-val add_class_scope : scope_name -> Classops.cl_typ -> unit
+val add_class_scope : scope_name -> scope_class list -> unit
(** Add only the interpretation of a notation that already has pa/pp rules *)
@@ -53,12 +58,12 @@ val add_syntax_extension :
(** Add a syntactic definition (as in "Notation f := ...") *)
-val add_syntactic_definition : identifier -> identifier list * constr_expr ->
+val add_syntactic_definition : Id.t -> Id.t list * constr_expr ->
bool -> Flags.compat_version option -> unit
(** Print the Camlp4 state of a grammar *)
-val print_grammar : string -> unit
+val pr_grammar : string -> Pp.std_ppcmds
val check_infix_modifiers : syntax_modifier list -> unit
diff --git a/toplevel/mltop.ml b/toplevel/mltop.ml
new file mode 100644
index 00000000..9dc1dd5b
--- /dev/null
+++ b/toplevel/mltop.ml
@@ -0,0 +1,439 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Errors
+open Util
+open Pp
+open Flags
+open Libobject
+open System
+
+(* Code to hook Coq into the ML toplevel -- depends on having the
+ objective-caml compiler mostly visible. The functions implemented here are:
+ \begin{itemize}
+ \item [dir_ml_load name]: Loads the ML module fname from the current ML
+ path.
+ \item [dir_ml_use]: Directive #use of Ocaml toplevel
+ \item [add_ml_dir]: Directive #directory of Ocaml toplevel
+ \end{itemize}
+
+ How to build an ML module interface with these functions.
+ The idea is that the ML directory path is like the Coq directory
+ path. So we can maintain the two in parallel.
+ In the same way, we can use the "ml_env" as a kind of ML
+ environment, which we freeze, unfreeze, and add things to just like
+ to the other environments.
+ Finally, we can create an object which is an ML module, and require
+ that the "caching" of the ML module cause the loading of the
+ associated ML file, if that file has not been yet loaded. Of
+ course, the problem is how to record dependencies between ML
+ modules.
+ (I do not know of a solution to this problem, other than to
+ put all the needed names into the ML Module object.) *)
+
+
+(* NB: this module relies on OCaml's Dynlink library. The bytecode version
+ of Dynlink is always available, but there are some architectures
+ with native compilation and no dynlink.cmxa. Instead of nasty IFDEF's
+ here for hiding the calls to Dynlink, we now simply reject this rather
+ rare situation during ./configure, and give instructions there about how
+ to build a dummy dynlink.cmxa, cf. dev/dynlink.ml. *)
+
+(* This path is where we look for .cmo *)
+let coq_mlpath_copy = ref ["."]
+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'))
+ in
+ coq_mlpath_copy := path :: List.filter filter !coq_mlpath_copy
+
+(* If there is a toplevel under Coq *)
+type toplevel = {
+ load_obj : string -> unit;
+ use_file : string -> unit;
+ add_dir : string -> unit;
+ ml_loop : unit -> unit }
+
+(* Determines the behaviour of Coq with respect to ML files (compiled
+ or not) *)
+type kind_load =
+ | WithTop of toplevel
+ | WithoutTop
+
+(* Must be always initialized *)
+let load = ref WithoutTop
+
+(* Are we in a native version of Coq? *)
+let is_native = Dynlink.is_native
+
+(* Sets and initializes a toplevel (if any) *)
+let set_top toplevel = load :=
+ WithTop toplevel;
+ Nativelib.load_obj := toplevel.load_obj
+
+(* Removes the toplevel (if any) *)
+let remove () =
+ load := WithoutTop;
+ Nativelib.load_obj := (fun x -> () : string -> unit)
+
+(* Tests if an Ocaml toplevel runs under Coq *)
+let is_ocaml_top () =
+ match !load with
+ | WithTop _ -> true
+ |_ -> false
+
+(* Tests if we can load ML files *)
+let has_dynlink = Coq_config.has_natdynlink || not is_native
+
+(* Runs the toplevel loop of Ocaml *)
+let ocaml_toploop () =
+ match !load with
+ | WithTop t -> Printexc.catch t.ml_loop ()
+ | _ -> ()
+
+(* Try to interpret load_obj's (internal) errors *)
+let report_on_load_obj_error exc =
+ let x = Obj.repr exc in
+ (* Try an horrible (fragile) hack to report on Symtable dynlink errors *)
+ (* (we follow ocaml's Printexc.to_string decoding of exceptions) *)
+ if Obj.is_block x && String.equal (Obj.magic (Obj.field (Obj.field x 0) 0)) "Symtable.Error"
+ then
+ let err_block = Obj.field x 1 in
+ if Int.equal (Obj.tag err_block) 0 then
+ (* Symtable.Undefined_global of string *)
+ str "reference to undefined global " ++
+ str (Obj.magic (Obj.field err_block 0))
+ else str (Printexc.to_string exc)
+ else str (Printexc.to_string exc)
+
+(* Dynamic loading of .cmo/.cma *)
+
+let ml_load s =
+ match !load with
+ | WithTop t ->
+ (try t.load_obj s; s
+ with
+ | e when Errors.noncritical e ->
+ let e = Errors.push e in
+ match fst e with
+ | (UserError _ | Failure _ | Not_found as u) -> Exninfo.iraise (u, snd e)
+ | exc ->
+ let msg = report_on_load_obj_error exc in
+ errorlabstrm "Mltop.load_object" (str"Cannot link ml-object " ++
+ str s ++ str" to Coq code (" ++ msg ++ str ")."))
+ | WithoutTop ->
+ try
+ Dynlink.loadfile s; s
+ with Dynlink.Error a ->
+ errorlabstrm "Mltop.load_object"
+ (strbrk "while loading " ++ str s ++
+ strbrk ": " ++ str (Dynlink.error_message a))
+
+let dir_ml_load s =
+ match !load with
+ | WithTop _ -> ml_load s
+ | WithoutTop ->
+ let warn = Flags.is_verbose() in
+ let _,gname = find_file_in_path ~warn !coq_mlpath_copy s in
+ ml_load gname
+
+(* Dynamic interpretation of .ml *)
+let dir_ml_use s =
+ match !load with
+ | WithTop t -> t.use_file s
+ | _ -> msg_warning (str "Cannot access the ML compiler")
+
+(* Adds a path to the ML paths *)
+let add_ml_dir s =
+ match !load with
+ | WithTop t -> t.add_dir s; keep_copy_mlpath s
+ | WithoutTop when has_dynlink -> keep_copy_mlpath s
+ | _ -> ()
+
+(* For Rec Add ML Path *)
+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 add_path ~unix_path:dir ~coq_root:coq_dirpath ~implicit =
+ if exists_dir dir then
+ begin
+ add_ml_dir dir;
+ Loadpath.add_load_path dir
+ (if implicit then Loadpath.ImplicitRootPath else Loadpath.RootPath)
+ coq_dirpath
+ end
+ else
+ msg_warning (str ("Cannot open " ^ dir))
+
+let convert_string d =
+ try Names.Id.of_string d
+ with UserError _ ->
+ msg_warning (str ("Directory "^d^" cannot be used as a Coq identifier (skipped)"));
+ raise Exit
+
+let add_rec_path ~unix_path ~coq_root ~implicit =
+ if exists_dir unix_path then
+ let dirs = all_subdirs ~unix_path in
+ let prefix = Names.DirPath.repr coq_root in
+ let convert_dirs (lp, cp) =
+ try
+ let path = List.rev_map convert_string cp @ prefix in
+ Some (lp, Names.DirPath.make path)
+ with Exit -> None
+ in
+ let dirs = List.map_filter convert_dirs dirs in
+ let () = add_ml_dir unix_path in
+ let add (path, dir) =
+ Loadpath.add_load_path path Loadpath.ImplicitPath dir in
+ let () = if implicit then List.iter add dirs in
+ Loadpath.add_load_path unix_path
+ (if implicit then Loadpath.ImplicitRootPath else Loadpath.RootPath)
+ coq_root
+ else
+ msg_warning (str ("Cannot open " ^ unix_path))
+
+(* convertit un nom quelconque en nom de fichier ou de module *)
+let mod_of_name name =
+ if Filename.check_suffix name ".cmo" then
+ Filename.chop_suffix name ".cmo"
+ else
+ name
+
+let get_ml_object_suffix name =
+ if Filename.check_suffix name ".cmo" then
+ Some ".cmo"
+ else if Filename.check_suffix name ".cma" then
+ Some ".cma"
+ else if Filename.check_suffix name ".cmxs" then
+ Some ".cmxs"
+ else
+ None
+
+let file_of_name name =
+ let suffix = get_ml_object_suffix name in
+ let fail s =
+ errorlabstrm "Mltop.load_object"
+ (str"File not found on loadpath : " ++ str s ++ str"\n" ++
+ str"Loadpath: " ++ str(String.concat ":" !coq_mlpath_copy)) in
+ if not (Filename.is_relative name) then
+ if Sys.file_exists name then name else fail name
+ else if is_native then
+ let name = match suffix with
+ | Some ((".cmo"|".cma") as suffix) ->
+ (Filename.chop_suffix name suffix) ^ ".cmxs"
+ | Some ".cmxs" -> name
+ | _ -> name ^ ".cmxs"
+ in
+ if is_in_path !coq_mlpath_copy name then name else fail name
+ else
+ let (full, base) = match suffix with
+ | Some ".cmo" | Some ".cma" -> true, name
+ | Some ".cmxs" -> false, Filename.chop_suffix name ".cmxs"
+ | _ -> false, name
+ in
+ if full then
+ if is_in_path !coq_mlpath_copy base then base else fail base
+ else
+ let name = base ^ ".cma" in
+ if is_in_path !coq_mlpath_copy name then name else
+ let name = base ^ ".cmo" in
+ if is_in_path !coq_mlpath_copy name then name else
+ fail (base ^ ".cm[ao]")
+
+(** Is the ML code of the standard library placed into loadable plugins
+ or statically compiled into coqtop ? For the moment this choice is
+ made according to the presence of native dynlink : even if bytecode
+ coqtop could always load plugins, we prefer to have uniformity between
+ bytecode and native versions. *)
+
+(* [known_loaded_module] contains the names of the loaded ML modules
+ * (linked or loaded with load_object). It is used not to load a
+ * module twice. It is NOT the list of ML modules Coq knows. *)
+
+let known_loaded_modules = ref String.Map.empty
+
+let add_known_module mname path =
+ if not (String.Map.mem mname !known_loaded_modules) ||
+ String.Map.find mname !known_loaded_modules = None then
+ known_loaded_modules := String.Map.add mname path !known_loaded_modules
+
+let module_is_known mname =
+ String.Map.mem mname !known_loaded_modules
+
+let known_module_path mname =
+ String.Map.find mname !known_loaded_modules
+
+(** A plugin is just an ML module with an initialization function. *)
+
+let known_loaded_plugins = ref String.Map.empty
+
+let add_known_plugin init name =
+ add_known_module name None;
+ known_loaded_plugins := String.Map.add name init !known_loaded_plugins
+
+let init_known_plugins () =
+ String.Map.iter (fun _ f -> f()) !known_loaded_plugins
+
+(** Registering functions to be used at caching time, that is when the Declare
+ ML module command is issued. *)
+
+let cache_objs = ref String.Map.empty
+
+let declare_cache_obj f name =
+ let objs = try String.Map.find name !cache_objs with Not_found -> [] in
+ let objs = f :: objs in
+ cache_objs := String.Map.add name objs !cache_objs
+
+let perform_cache_obj name =
+ let objs = try String.Map.find name !cache_objs with Not_found -> [] in
+ let objs = List.rev objs in
+ List.iter (fun f -> f ()) objs
+
+(** ml object = ml module or plugin *)
+
+let init_ml_object mname =
+ try String.Map.find mname !known_loaded_plugins ()
+ with Not_found -> ()
+
+let load_ml_object mname ?path fname=
+ let path = match path with
+ | None -> dir_ml_load fname
+ | Some p -> ml_load p in
+ add_known_module mname (Some path);
+ init_ml_object mname;
+ path
+
+let dir_ml_load m = ignore(dir_ml_load m)
+let add_known_module m = add_known_module m None
+let load_ml_object_raw fname = dir_ml_load (file_of_name fname)
+let load_ml_objects_raw_rex rex =
+ List.iter (fun (_,fp) ->
+ let name = file_of_name (Filename.basename fp) in
+ try dir_ml_load name
+ with e -> prerr_endline (Printexc.to_string e))
+ (System.where_in_path_rex !coq_mlpath_copy rex)
+
+(* Summary of declared ML Modules *)
+
+(* List and not String.Set because order is important: most recent first. *)
+
+let loaded_modules = ref []
+let get_loaded_modules () = List.rev !loaded_modules
+let add_loaded_module md path =
+ if not (List.mem_assoc md !loaded_modules) then
+ loaded_modules := (md,path) :: !loaded_modules
+let reset_loaded_modules () = loaded_modules := []
+
+let if_verbose_load verb f name ?path fname =
+ if not verb then f name ?path fname
+ else
+ let info = "[Loading ML file "^fname^" ..." in
+ try
+ let path = f name ?path fname in
+ msg_info (str (info^" done]"));
+ path
+ with reraise ->
+ msg_info (str (info^" failed]"));
+ raise reraise
+
+(** Load a module for the first time (i.e. dynlink it)
+ or simulate its reload (i.e. doing nothing except maybe
+ an initialization function). *)
+
+let trigger_ml_object verb cache reinit ?path name =
+ if module_is_known name then begin
+ if reinit then init_ml_object name;
+ add_loaded_module name (known_module_path name);
+ if cache then perform_cache_obj name
+ end else if not has_dynlink then
+ error ("Dynamic link not supported (module "^name^")")
+ else begin
+ let file = file_of_name (Option.default name path) in
+ let path =
+ if_verbose_load (verb && is_verbose ()) load_ml_object name ?path file in
+ add_loaded_module name (Some path);
+ if cache then perform_cache_obj name
+ end
+
+let load_ml_object n m = ignore(load_ml_object n m)
+
+let unfreeze_ml_modules x =
+ reset_loaded_modules ();
+ List.iter
+ (fun (name,path) -> trigger_ml_object false false false ?path name) x
+
+let _ =
+ Summary.declare_summary Summary.ml_modules
+ { Summary.freeze_function = (fun _ -> get_loaded_modules ());
+ Summary.unfreeze_function = unfreeze_ml_modules;
+ Summary.init_function = reset_loaded_modules }
+
+(* Liboject entries of declared ML Modules *)
+
+type ml_module_object = {
+ mlocal : Vernacexpr.locality_flag;
+ mnames : string list
+}
+
+let cache_ml_objects (_,{mnames=mnames}) =
+ let iter obj = trigger_ml_object true true true obj in
+ List.iter iter mnames
+
+let load_ml_objects _ (_,{mnames=mnames}) =
+ let iter obj = trigger_ml_object true false true obj in
+ List.iter iter mnames
+
+let classify_ml_objects ({mlocal=mlocal} as o) =
+ if mlocal then Dispose else Substitute o
+
+let inMLModule : ml_module_object -> obj =
+ declare_object
+ {(default_object "ML-MODULE") with
+ cache_function = cache_ml_objects;
+ load_function = load_ml_objects;
+ subst_function = (fun (_,o) -> o);
+ classify_function = classify_ml_objects }
+
+let declare_ml_modules local l =
+ let l = List.map mod_of_name l in
+ Lib.add_anonymous_leaf (inMLModule {mlocal=local; mnames=l})
+
+let print_ml_path () =
+ let l = !coq_mlpath_copy in
+ str"ML Load Path:" ++ fnl () ++ str" " ++
+ hv 0 (prlist_with_sep fnl str l)
+
+(* Printing of loaded ML modules *)
+
+let print_ml_modules () =
+ let l = get_loaded_modules () in
+ str"Loaded ML Modules: " ++ pr_vertical_list str (List.map fst l)
+
+let print_gc () =
+ let stat = Gc.stat () in
+ let msg =
+ str "minor words: " ++ real stat.Gc.minor_words ++ fnl () ++
+ str "promoted words: " ++ real stat.Gc.promoted_words ++ fnl () ++
+ str "major words: " ++ real stat.Gc.major_words ++ fnl () ++
+ str "minor_collections: " ++ int stat.Gc.minor_collections ++ fnl () ++
+ str "major_collections: " ++ int stat.Gc.major_collections ++ fnl () ++
+ str "heap_words: " ++ int stat.Gc.heap_words ++ fnl () ++
+ str "heap_chunks: " ++ int stat.Gc.heap_chunks ++ fnl () ++
+ str "live_words: " ++ int stat.Gc.live_words ++ fnl () ++
+ str "live_blocks: " ++ int stat.Gc.live_blocks ++ fnl () ++
+ str "free_words: " ++ int stat.Gc.free_words ++ fnl () ++
+ str "free_blocks: " ++ int stat.Gc.free_blocks ++ fnl () ++
+ str "largest_free: " ++ int stat.Gc.largest_free ++ fnl () ++
+ str "fragments: " ++ int stat.Gc.fragments ++ fnl () ++
+ str "compactions: " ++ int stat.Gc.compactions ++ fnl () ++
+ str "top_heap_words: " ++ int stat.Gc.top_heap_words ++ fnl () ++
+ str "stack_size: " ++ int stat.Gc.stack_size
+ in
+ hv 0 msg
diff --git a/toplevel/mltop.ml4 b/toplevel/mltop.ml4
deleted file mode 100644
index bf7b995a..00000000
--- a/toplevel/mltop.ml4
+++ /dev/null
@@ -1,337 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-open Util
-open Pp
-open Flags
-open System
-open Libobject
-open Library
-open System
-open Vernacinterp
-
-(* Code to hook Coq into the ML toplevel -- depends on having the
- objective-caml compiler mostly visible. The functions implemented here are:
- \begin{itemize}
- \item [dir_ml_load name]: Loads the ML module fname from the current ML
- path.
- \item [dir_ml_use]: Directive #use of Ocaml toplevel
- \item [add_ml_dir]: Directive #directory of Ocaml toplevel
- \end{itemize}
-
- How to build an ML module interface with these functions.
- The idea is that the ML directory path is like the Coq directory
- path. So we can maintain the two in parallel.
- In the same way, we can use the "ml_env" as a kind of ML
- environment, which we freeze, unfreeze, and add things to just like
- to the other environments.
- Finally, we can create an object which is an ML module, and require
- that the "caching" of the ML module cause the loading of the
- associated ML file, if that file has not been yet loaded. Of
- course, the problem is how to record dependencies between ML
- modules.
- (I do not know of a solution to this problem, other than to
- put all the needed names into the ML Module object.) *)
-
-(* This path is where we look for .cmo *)
-let coq_mlpath_copy = ref ["."]
-let keep_copy_mlpath path =
- let cpath = canonical_path_name path in
- let filter path' = (cpath <> canonical_path_name path') in
- coq_mlpath_copy := path :: List.filter filter !coq_mlpath_copy
-
-(* If there is a toplevel under Coq *)
-type toplevel = {
- load_obj : string -> unit;
- use_file : string -> unit;
- add_dir : string -> unit;
- ml_loop : unit -> unit }
-
-(* Determines the behaviour of Coq with respect to ML files (compiled
- or not) *)
-type kind_load =
- | WithTop of toplevel
- | WithoutTop
-
-(* Must be always initialized *)
-let load = ref WithoutTop
-
-(* Are we in a native version of Coq? *)
-let is_native = IFDEF Byte THEN false ELSE true END
-
-(* Sets and initializes a toplevel (if any) *)
-let set_top toplevel = load := WithTop toplevel
-
-(* Removes the toplevel (if any) *)
-let remove ()= load := WithoutTop
-
-(* Tests if an Ocaml toplevel runs under Coq *)
-let is_ocaml_top () =
- match !load with
- | WithTop _ -> true
- |_ -> false
-
-(* Tests if we can load ML files *)
-let has_dynlink = IFDEF HasDynlink THEN true ELSE false END
-
-(* Runs the toplevel loop of Ocaml *)
-let ocaml_toploop () =
- match !load with
- | WithTop t -> Printexc.catch t.ml_loop ()
- | _ -> ()
-
-(* Dynamic loading of .cmo/.cma *)
-let dir_ml_load s =
- match !load with
- | WithTop t ->
- (try t.load_obj s
- with
- | (UserError _ | Failure _ | Anomaly _ | Not_found as u) -> raise u
- | e when Errors.noncritical e ->
- errorlabstrm "Mltop.load_object"
- (str"Cannot link ml-object " ++ str s ++ str" to Coq code."))
-(* TO DO: .cma loading without toplevel *)
- | WithoutTop ->
- IFDEF HasDynlink THEN
- (* WARNING
- * if this code section starts to use a module not used elsewhere
- * in this file, the Makefile dependency logic needs to be updated.
- *)
- let warn = Flags.is_verbose() in
- let _,gname = find_file_in_path ~warn !coq_mlpath_copy s in
- try
- Dynlink.loadfile gname;
- with | Dynlink.Error a ->
- errorlabstrm "Mltop.load_object" (str (Dynlink.error_message a))
- ELSE
- errorlabstrm "Mltop.no_load_object"
- (str"Loading of ML object file forbidden in a native Coq.")
- END
-
-(* Dynamic interpretation of .ml *)
-let dir_ml_use s =
- match !load with
- | WithTop t -> t.use_file s
- | _ -> msg_warn "Cannot access the ML compiler"
-
-(* Adds a path to the ML paths *)
-let add_ml_dir s =
- match !load with
- | WithTop t -> t.add_dir s; keep_copy_mlpath s
- | WithoutTop when has_dynlink -> keep_copy_mlpath s
- | _ -> ()
-
-(* For Rec Add ML Path *)
-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 add_path ~unix_path:dir ~coq_root:coq_dirpath =
- if exists_dir dir then
- begin
- add_ml_dir dir;
- Library.add_load_path true (dir,coq_dirpath)
- end
- else
- msg_warning (str ("Cannot open " ^ dir))
-
-let convert_string d =
- try Names.id_of_string d
- with e when Errors.noncritical e ->
- if_warn msg_warning
- (str ("Directory "^d^" cannot be used as a Coq identifier (skipped)"));
- flush_all ();
- failwith "caught"
-
-let add_rec_path ~unix_path ~coq_root =
- if exists_dir unix_path then
- let dirs = all_subdirs ~unix_path in
- let prefix = Names.repr_dirpath coq_root in
- let convert_dirs (lp,cp) =
- (lp,Names.make_dirpath (List.map convert_string (List.rev cp)@prefix)) in
- let dirs = map_succeed convert_dirs dirs in
- List.iter (fun lpe -> add_ml_dir (fst lpe)) dirs;
- add_ml_dir unix_path;
- List.iter (Library.add_load_path false) dirs;
- Library.add_load_path true (unix_path, coq_root)
- else
- msg_warning (str ("Cannot open " ^ unix_path))
-
-(* convertit un nom quelconque en nom de fichier ou de module *)
-let mod_of_name name =
- let base =
- if Filename.check_suffix name ".cmo" then
- Filename.chop_suffix name ".cmo"
- else
- name
- in
- String.capitalize base
-
-let get_ml_object_suffix name =
- if Filename.check_suffix name ".cmo" then
- Some ".cmo"
- else if Filename.check_suffix name ".cma" then
- Some ".cma"
- else if Filename.check_suffix name ".cmxs" then
- Some ".cmxs"
- else
- None
-
-let file_of_name name =
- let name = String.uncapitalize name in
- let suffix = get_ml_object_suffix name in
- let fail s =
- errorlabstrm "Mltop.load_object"
- (str"File not found on loadpath : " ++ str s) in
- if is_native then
- let name = match suffix with
- | Some ((".cmo"|".cma") as suffix) ->
- (Filename.chop_suffix name suffix) ^ ".cmxs"
- | Some ".cmxs" -> name
- | _ -> name ^ ".cmxs"
- in
- if is_in_path !coq_mlpath_copy name then name else fail name
- else
- let (full, base) = match suffix with
- | Some ".cmo" | Some ".cma" -> true, name
- | Some ".cmxs" -> false, Filename.chop_suffix name ".cmxs"
- | _ -> false, name
- in
- if full then
- if is_in_path !coq_mlpath_copy base then base else fail base
- else
- let name = base ^ ".cmo" in
- if is_in_path !coq_mlpath_copy name then name else
- let name = base ^ ".cma" in
- if is_in_path !coq_mlpath_copy name then name else
- fail (base ^ ".cm[oa]")
-
-(** Is the ML code of the standard library placed into loadable plugins
- or statically compiled into coqtop ? For the moment this choice is
- made according to the presence of native dynlink : even if bytecode
- coqtop could always load plugins, we prefer to have uniformity between
- bytecode and native versions. *)
-
-(* [known_loaded_module] contains the names of the loaded ML modules
- * (linked or loaded with load_object). It is used not to load a
- * module twice. It is NOT the list of ML modules Coq knows. *)
-
-let known_loaded_modules = ref Stringset.empty
-
-let add_known_module mname =
- let mname = String.capitalize mname in
- known_loaded_modules := Stringset.add mname !known_loaded_modules
-
-let module_is_known mname =
- Stringset.mem (String.capitalize mname) !known_loaded_modules
-
-(** A plugin is just an ML module with an initialization function. *)
-
-let known_loaded_plugins = ref Stringmap.empty
-
-let add_known_plugin init name =
- let name = String.capitalize name in
- add_known_module name;
- known_loaded_plugins := Stringmap.add name init !known_loaded_plugins
-
-let init_known_plugins () =
- Stringmap.iter (fun _ f -> f()) !known_loaded_plugins
-
-(** ml object = ml module or plugin *)
-
-let init_ml_object mname =
- try Stringmap.find mname !known_loaded_plugins ()
- with Not_found -> ()
-
-let load_ml_object mname fname=
- dir_ml_load fname;
- add_known_module mname;
- init_ml_object mname
-
-(* Summary of declared ML Modules *)
-
-(* List and not Stringset because order is important: most recent first. *)
-
-let loaded_modules = ref []
-let get_loaded_modules () = List.rev !loaded_modules
-let add_loaded_module md = loaded_modules := md :: !loaded_modules
-let reset_loaded_modules () = loaded_modules := []
-
-let if_verbose_load verb f name fname =
- if not verb then f name fname
- else
- let info = "[Loading ML file "^fname^" ..." in
- try
- f name fname;
- msgnl (str (info^" done]"));
- with reraise ->
- msgnl (str (info^" failed]"));
- raise reraise
-
-(** Load a module for the first time (i.e. dynlink it)
- or simulate its reload (i.e. doing nothing except maybe
- an initialization function). *)
-
-let cache_ml_object verb reinit name =
- begin
- if module_is_known name then
- (if reinit then init_ml_object name)
- else if not has_dynlink then
- error ("Dynamic link not supported (module "^name^")")
- else
- if_verbose_load (verb && is_verbose ())
- load_ml_object name (file_of_name name)
- end;
- add_loaded_module name
-
-let unfreeze_ml_modules x =
- reset_loaded_modules ();
- List.iter (cache_ml_object false false) x
-
-let _ =
- Summary.declare_summary "ML-MODULES"
- { Summary.freeze_function = get_loaded_modules;
- Summary.unfreeze_function = unfreeze_ml_modules;
- Summary.init_function = reset_loaded_modules }
-
-(* Liboject entries of declared ML Modules *)
-
-type ml_module_object = {
- mlocal : Vernacexpr.locality_flag;
- mnames : string list
-}
-
-let cache_ml_objects (_,{mnames=mnames}) =
- List.iter (cache_ml_object true true) mnames
-
-let classify_ml_objects ({mlocal=mlocal} as o) =
- if mlocal then Dispose else Substitute o
-
-let inMLModule : ml_module_object -> obj =
- declare_object
- {(default_object "ML-MODULE") with
- load_function = (fun _ -> cache_ml_objects);
- cache_function = cache_ml_objects;
- subst_function = (fun (_,o) -> o);
- classify_function = classify_ml_objects }
-
-let declare_ml_modules local l =
- let l = List.map mod_of_name l in
- Lib.add_anonymous_leaf (inMLModule {mlocal=local; mnames=l})
-
-let print_ml_path () =
- let l = !coq_mlpath_copy in
- ppnl (str"ML Load Path:" ++ fnl () ++ str" " ++
- hv 0 (prlist_with_sep pr_fnl pr_str l))
-
-(* Printing of loaded ML modules *)
-
-let print_ml_modules () =
- let l = get_loaded_modules () in
- pp (str"Loaded ML Modules: " ++ pr_vertical_list pr_str l)
diff --git a/toplevel/mltop.mli b/toplevel/mltop.mli
index 4a6e1e0b..2a91afd8 100644
--- a/toplevel/mltop.mli
+++ b/toplevel/mltop.mli
@@ -1,11 +1,13 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+(** {5 Toplevel management} *)
+
(** If there is a toplevel under Coq, it is described by the following
record. *)
type toplevel = {
@@ -26,12 +28,14 @@ val remove : unit -> unit
(** Tests if an Ocaml toplevel runs under Coq *)
val is_ocaml_top : unit -> bool
-(** Tests if we can load ML files *)
-val has_dynlink : bool
-
(** Starts the Ocaml toplevel loop *)
val ocaml_toploop : unit -> unit
+(** {5 ML Dynlink} *)
+
+(** Tests if we can load ML files *)
+val has_dynlink : bool
+
(** Dynamic loading of .cmo *)
val dir_ml_load : string -> unit
@@ -43,13 +47,17 @@ val add_ml_dir : string -> unit
val add_rec_ml_dir : string -> unit
(** Adds a path to the Coq and ML paths *)
-val add_path : unix_path:string -> coq_root:Names.dir_path -> unit
-val add_rec_path : unix_path:string -> coq_root:Names.dir_path -> unit
+val add_path : unix_path:string -> coq_root:Names.DirPath.t -> implicit:bool -> unit
+val add_rec_path : unix_path:string -> coq_root:Names.DirPath.t -> implicit:bool -> unit
(** List of modules linked to the toplevel *)
val add_known_module : string -> unit
val module_is_known : string -> bool
val load_ml_object : string -> string -> unit
+val load_ml_object_raw : string -> unit
+val load_ml_objects_raw_rex : Str.regexp -> unit
+
+(** {5 Initialization functions} *)
(** Declare a plugin and its initialization function.
A plugin is just an ML module with an initialization function.
@@ -61,8 +69,19 @@ val add_known_plugin : (unit -> unit) -> string -> unit
(** Calls all initialization functions in a non-specified order *)
val init_known_plugins : unit -> unit
+(** Register a callback that will be called when the module is declared with
+ the Declare ML Module command. This is useful to define Coq objects at that
+ time only. Several functions can be defined for one module; they will be
+ called in the order of declaration, and after the ML module has been
+ properly initialized. *)
+val declare_cache_obj : (unit -> unit) -> string -> unit
+
+(** {5 Declaring modules} *)
+
val declare_ml_modules : Vernacexpr.locality_flag -> string list -> unit
-val print_ml_path : unit -> unit
+(** {5 Utilities} *)
-val print_ml_modules : unit -> unit
+val print_ml_path : unit -> Pp.std_ppcmds
+val print_ml_modules : unit -> Pp.std_ppcmds
+val print_gc : unit -> Pp.std_ppcmds
diff --git a/toplevel/obligations.ml b/toplevel/obligations.ml
new file mode 100644
index 00000000..aa068586
--- /dev/null
+++ b/toplevel/obligations.ml
@@ -0,0 +1,1075 @@
+open Printf
+open Globnames
+open Libobject
+open Entries
+open Decl_kinds
+open Declare
+
+(**
+ - Get types of existentials ;
+ - Flatten dependency tree (prefix order) ;
+ - Replace existentials by De Bruijn indices in term, applied to the right arguments ;
+ - Apply term prefixed by quantification on "existentials".
+*)
+
+open Term
+open Context
+open Vars
+open Names
+open Evd
+open Pp
+open Errors
+open Util
+
+let declare_fix_ref = ref (fun _ _ _ _ _ _ -> 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)
+
+let check_evars env evm =
+ Evar.Map.iter
+ (fun key evi ->
+ let (loc,k) = evar_source key evm in
+ match k with
+ | Evar_kinds.QuestionMark _
+ | Evar_kinds.ImplicitArg (_,_,false) -> ()
+ | _ ->
+ Pretype_errors.error_unsolvable_implicit loc env evm key None)
+ (Evd.undefined_map evm)
+
+type oblinfo =
+ { ev_name: int * Id.t;
+ ev_hyps: named_context;
+ ev_status: 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 seen = ref Int.Set.empty in
+ let transparent = ref Id.Set.empty in
+ let evar_info id = List.assoc_f Evar.equal id evs in
+ let rec substrec (depth, fixrels) c = match kind_of_term c with
+ | Evar (k, args) ->
+ let { ev_name = (id, idstr) ;
+ ev_hyps = hyps ; ev_chop = chop } =
+ try evar_info k
+ with Not_found ->
+ anomaly ~label:"eterm" (Pp.str "existential variable " ++ int (Evar.repr k) ++ str " not found")
+ in
+ seen := Int.Set.add id !seen;
+ (* Evar arguments are created in inverse order,
+ and we must not apply to defined ones (i.e. LetIn's)
+ *)
+ let args =
+ let n = match chop with None -> 0 | Some c -> c in
+ let (l, r) = List.chop n (List.rev (Array.to_list args)) in
+ List.rev r
+ in
+ let args =
+ let rec aux hyps args acc =
+ match hyps, args with
+ ((_, None, _) :: tlh), (c :: tla) ->
+ aux tlh tla ((substrec (depth, fixrels) c) :: acc)
+ | ((_, Some _, _) :: tlh), (_ :: tla) ->
+ aux tlh tla acc
+ | [], [] -> acc
+ | _, _ -> acc (*failwith "subst_evars: invalid argument"*)
+ in aux hyps args []
+ in
+ if List.exists
+ (fun x -> match kind_of_term x with
+ | Rel n -> Int.List.mem n fixrels
+ | _ -> false) args
+ then
+ transparent := Id.Set.add idstr !transparent;
+ mkApp (idf idstr, Array.of_list args)
+ | Fix _ ->
+ map_constr_with_binders succfix substrec (depth, 1 :: fixrels) c
+ | _ -> map_constr_with_binders succfix substrec (depth, fixrels) c
+ in
+ let t' = substrec (0, []) t in
+ t', !seen, !transparent
+
+
+(** Substitute variable references in t using De Bruijn indices,
+ where n binders were passed through. *)
+let subst_vars acc n t =
+ let var_index id = Util.List.index Id.equal id acc in
+ let rec substrec depth c = match kind_of_term c with
+ | Var v -> (try mkRel (depth + (var_index v)) with Not_found -> c)
+ | _ -> map_constr_with_binders succ substrec depth c
+ in
+ substrec 0 t
+
+(** Rewrite type of an evar ([ H1 : t1, ... Hn : tn |- concl ])
+ to a product : forall H1 : t1, ..., forall Hn : tn, concl.
+ Changes evars and hypothesis references to variable references.
+*)
+let etype_of_evar evs hyps concl =
+ let rec aux acc n = function
+ (id, copt, t) :: tl ->
+ let t', s, trans = subst_evar_constr evs n mkVar t in
+ let t'' = subst_vars acc 0 t' in
+ let rest, s', trans' = aux (id :: 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 ->
+ 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,
+ Int.Set.union s'' s',
+ Id.Set.union trans'' trans'
+ | None ->
+ mkNamedProd_or_LetIn (id, None, t'') rest, s', trans')
+ | [] ->
+ let t', s, trans = subst_evar_constr evs n mkVar concl in
+ subst_vars acc 0 t', s, trans
+ in aux [] 0 (List.rev hyps)
+
+let trunc_named_context n ctx =
+ let len = List.length ctx in
+ List.firstn (len - n) ctx
+
+let rec chop_product n t =
+ if Int.equal n 0 then Some t
+ else
+ match kind_of_term t with
+ | Prod (_, _, b) -> if noccurn 1 b then chop_product (pred n) (Termops.pop b) else None
+ | _ -> None
+
+let evar_dependencies evm oev =
+ let one_step deps =
+ Evar.Set.fold (fun ev s ->
+ let evi = Evd.find evm ev in
+ let deps' = evars_of_filtered_evar_info evi in
+ if Evar.Set.mem oev deps' then
+ invalid_arg ("Ill-formed evar map: cycle detected for evar " ^ string_of_existential oev)
+ else Evar.Set.union deps' s)
+ deps deps
+ in
+ let rec aux deps =
+ let deps' = one_step deps in
+ if Evar.Set.equal deps deps' then deps
+ else aux deps'
+ in aux (Evar.Set.singleton oev)
+
+let move_after (id, ev, deps as obl) l =
+ let rec aux restdeps = function
+ | (id', _, _) as obl' :: tl ->
+ let restdeps' = Evar.Set.remove id' restdeps in
+ if Evar.Set.is_empty restdeps' then
+ obl' :: obl :: tl
+ else obl' :: aux restdeps' tl
+ | [] -> [obl]
+ in aux (Evar.Set.remove id deps) l
+
+let sort_dependencies evl =
+ let rec aux l found list =
+ match l with
+ | (id, ev, deps) as obl :: tl ->
+ let found' = Evar.Set.union found (Evar.Set.singleton id) in
+ if Evar.Set.subset deps found' then
+ aux tl found' (obl :: list)
+ else aux (move_after obl tl) found list
+ | [] -> List.rev list
+ in aux evl Evar.Set.empty []
+
+open Environ
+
+let eterm_obligations env name evm fs ?status t ty =
+ (* 'Serialize' the evars *)
+ let nc = Environ.named_context env in
+ let nc_len = Context.named_context_length nc in
+ let evm = Evarutil.nf_evar_map_undefined evm in
+ let evl = Evarutil.non_instantiated evm in
+ let evl = Evar.Map.bindings evl in
+ let evl = List.map (fun (id, ev) -> (id, ev, evar_dependencies evm id)) evl in
+ let sevl = sort_dependencies evl in
+ let evl = List.map (fun (id, ev, _) -> id, ev) sevl in
+ let evn =
+ let i = ref (-1) in
+ List.rev_map (fun (id, ev) -> incr i;
+ (id, (!i, Id.of_string
+ (Id.to_string name ^ "_obligation_" ^ string_of_int (succ !i))),
+ ev)) evl
+ in
+ let evts =
+ (* Remove existential variables in types and build the corresponding products *)
+ List.fold_right
+ (fun (id, (n, nstr), ev) l ->
+ let hyps = Evd.evar_filtered_context ev in
+ let hyps = trunc_named_context nc_len hyps in
+ let evtyp, deps, transp = etype_of_evar l hyps ev.evar_concl in
+ let evtyp, hyps, chop =
+ match chop_product fs evtyp with
+ | Some t -> t, trunc_named_context fs hyps, fs
+ | None -> evtyp, hyps, 0
+ in
+ let loc, k = evar_source id evm in
+ let status = match k with Evar_kinds.QuestionMark o -> 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
+ 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 }
+ in (id, info) :: l)
+ evn []
+ in
+ let t', _, transparent = (* Substitute evar refs in the term by variables *)
+ subst_evar_constr evts 0 mkVar t
+ in
+ let ty, _, _ = subst_evar_constr evts 0 mkVar ty in
+ let evars =
+ List.map (fun (ev, info) ->
+ let { ev_name = (_, name); ev_status = status;
+ ev_src = src; ev_typ = typ; ev_deps = deps; ev_tac = tac } = info
+ in
+ let 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
+ in
+ let evnames = List.map (fun (ev, info) -> ev, snd info.ev_name) evts in
+ let evmap f c = pi1 (subst_evar_constr evts 0 f c) in
+ Array.of_list (List.rev evars), (evnames, evmap), t', ty
+
+let tactics_module = ["Program";"Tactics"]
+let safe_init_constant md name () =
+ Coqlib.check_required_library ("Coq"::md);
+ Coqlib.gen_constant "Obligations" md name
+let hide_obligation = safe_init_constant tactics_module "obligation"
+
+let pperror cmd = Errors.errorlabstrm "Program" cmd
+let error s = pperror (str s)
+
+let reduce c =
+ Reductionops.clos_norm_flags Closure.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)
+ | 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
+
+type 'a obligation_body =
+ | DefinedObl of 'a
+ | TermObl of constr
+
+type obligation =
+ { obl_name : Id.t;
+ obl_type : types;
+ obl_location : Evar_kinds.t Loc.located;
+ obl_body : constant obligation_body option;
+ obl_status : Evar_kinds.obligation_definition_status;
+ obl_deps : Int.Set.t;
+ obl_tac : unit Proofview.tactic option;
+ }
+
+type obligations = (obligation array * int)
+
+type fixpoint_kind =
+ | IsFixpoint of (Id.t Loc.located option * Constrexpr.recursion_order_expr) list
+ | IsCoFixpoint
+
+type notations = (Vernacexpr.lstring * Constrexpr.constr_expr * Notation_term.scope_name option) list
+
+type program_info = {
+ prg_name: Id.t;
+ prg_body: constr;
+ prg_type: constr;
+ prg_ctx: Evd.evar_universe_context;
+ prg_obligations: obligations;
+ prg_deps : Id.t list;
+ prg_fixkind : fixpoint_kind option ;
+ prg_implicits : (Constrexpr.explicitation * (bool * bool * bool)) list;
+ prg_notations : notations ;
+ prg_kind : definition_kind;
+ prg_reduce : constr -> constr;
+ prg_hook : unit Lemmas.declaration_hook;
+}
+
+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; }
+
+(* true = hide obligations *)
+let hide_obligations = ref false
+
+let set_hide_obligations = (:=) hide_obligations
+let get_hide_obligations () = !hide_obligations
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "Hidding of Program obligations";
+ optkey = ["Hide";"Obligations"];
+ optread = get_hide_obligations;
+ optwrite = set_hide_obligations; }
+
+let shrink_obligations = ref false
+
+let set_shrink_obligations = (:=) shrink_obligations
+let get_shrink_obligations () = !shrink_obligations
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "Shrinking of Program obligations";
+ optkey = ["Shrink";"Obligations"];
+ optread = get_shrink_obligations;
+ optwrite = set_shrink_obligations; }
+
+let evar_of_obligation o = make_evar (Global.named_context_val ()) o.obl_type
+
+let get_body obl =
+ match obl.obl_body with
+ | None -> assert false
+ | 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
+
+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'
+
+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)
+ deps []
+
+let subst_deps expand obls deps t =
+ let osubst = obl_substitution expand obls deps in
+ (Vars.replace_vars (List.map (fun (n, (_, b)) -> n, b) osubst) t)
+
+let rec prod_app t n =
+ match kind_of_term (strip_outer_cast t) with
+ | Prod (_,_,b) -> subst1 n b
+ | LetIn (_, b, t, b') -> prod_app (subst1 b b') n
+ | _ ->
+ errorlabstrm "prod_app"
+ (str"Needed a product, but didn't find one" ++ fnl ())
+
+
+(* prod_appvect T [| a1 ; ... ; an |] -> (T a1 ... an) *)
+let prod_applist t nL = List.fold_left prod_app t nL
+
+let replace_appvars subst =
+ let rec aux c =
+ let f, l = decompose_app c in
+ if isVar f then
+ try
+ let c' = List.map (map_constr aux) l in
+ let (t, b) = Id.List.assoc (destVar f) subst in
+ mkApp (delayed_force hide_obligation,
+ [| prod_applist t c'; applistc b c' |])
+ with Not_found -> map_constr aux c
+ else map_constr aux c
+ in map_constr aux
+
+let subst_prog expand obls ints prg =
+ let subst = obl_substitution expand obls ints in
+ if get_hide_obligations () then
+ (replace_appvars subst prg.prg_body,
+ replace_appvars subst ((* Termops.refresh_universes *) prg.prg_type))
+ else
+ let subst' = List.map (fun (n, (_, b)) -> n, b) subst in
+ (Vars.replace_vars subst' prg.prg_body,
+ Vars.replace_vars subst' ((* Termops.refresh_universes *) prg.prg_type))
+
+let subst_deps_obl obls obl =
+ let t' = subst_deps true obls obl.obl_deps obl.obl_type in
+ { obl with obl_type = t' }
+
+module ProgMap = Map.Make(Id)
+
+let map_replace k v m = ProgMap.add k v (ProgMap.remove k m)
+
+let map_keys m = ProgMap.fold (fun k _ l -> k :: l) m []
+
+let map_cardinal m =
+ let i = ref 0 in
+ ProgMap.iter (fun _ _ -> incr i) m;
+ !i
+
+exception Found of program_info
+
+let map_first m =
+ try
+ ProgMap.iter (fun _ v -> raise (Found v)) m;
+ assert(false)
+ with Found x -> x
+
+let from_prg : program_info ProgMap.t ref =
+ Summary.ref ProgMap.empty ~name:"program-tcc-table"
+
+let close sec =
+ if not (ProgMap.is_empty !from_prg) then
+ let keys = map_keys !from_prg in
+ errorlabstrm "Program"
+ (str "Unsolved obligations when closing " ++ str sec ++ str":" ++ spc () ++
+ prlist_with_sep spc (fun x -> Nameops.pr_id x) keys ++
+ (str (if Int.equal (List.length keys) 1 then " has " else " have ") ++
+ str "unsolved obligations"))
+
+let input : program_info ProgMap.t -> obj =
+ declare_object
+ { (default_object "Program state") with
+ cache_function = (fun (na, pi) -> from_prg := pi);
+ load_function = (fun _ (_, pi) -> from_prg := pi);
+ discharge_function = (fun _ -> close "section"; None);
+ classify_function = (fun _ -> close "module"; Dispose) }
+
+open Evd
+
+let progmap_remove prg =
+ Lib.add_anonymous_leaf (input (ProgMap.remove prg.prg_name !from_prg))
+
+let progmap_add n prg =
+ Lib.add_anonymous_leaf (input (ProgMap.add n prg !from_prg))
+
+let progmap_replace prg' =
+ Lib.add_anonymous_leaf (input (map_replace prg'.prg_name prg' !from_prg))
+
+let rec intset_to = function
+ -1 -> Int.Set.empty
+ | n -> Int.Set.add n (intset_to (pred n))
+
+let subst_body expand prg =
+ let obls, _ = prg.prg_obligations in
+ let ints = intset_to (pred (Array.length obls)) in
+ subst_prog expand obls ints prg
+
+let declare_definition prg =
+ let body, typ = subst_body true prg in
+ let nf = Universes.nf_evars_and_universes_opt_subst (fun x -> None)
+ (Evd.evar_universe_context_subst prg.prg_ctx) in
+ let ce =
+ definition_entry ~types:(nf typ) ~poly:(pi2 prg.prg_kind)
+ ~univs:(Evd.evar_context_universe_context prg.prg_ctx) (nf body)
+ in
+ progmap_remove prg;
+ !declare_definition_ref prg.prg_name
+ prg.prg_kind ce prg.prg_implicits
+ (Lemmas.mk_hook (fun l r ->
+ Lemmas.call_hook (fun exn -> exn) prg.prg_hook l r; r))
+
+open Pp
+
+let rec lam_index n t acc =
+ match kind_of_term t with
+ | Lambda (Name n', _, _) when Id.equal n n' ->
+ acc
+ | Lambda (_, _, b) ->
+ lam_index n b (succ acc)
+ | _ -> raise Not_found
+
+let compute_possible_guardness_evidences (n,_) fixbody fixtype =
+ match n with
+ | Some (loc, n) -> [lam_index n fixbody 0]
+ | None ->
+ (* If recursive argument was not given by user, we try all args.
+ An earlier approach was to look only for inductive arguments,
+ but doing it properly involves delta-reduction, and it finally
+ doesn't seem to worth the effort (except for huge mutual
+ fixpoints ?) *)
+ let m = nb_prod fixtype in
+ let ctx = fst (decompose_prod_n_assum m fixtype) in
+ List.map_i (fun i _ -> i) 0 ctx
+
+let mk_proof c = ((c, Univ.ContextSet.empty), Declareops.no_seff)
+
+let declare_mutual_definition l =
+ let len = List.length l in
+ let first = List.hd l in
+ let fixdefs, fixtypes, fiximps =
+ List.split3
+ (List.map (fun x ->
+ let subs, typ = (subst_body true x) in
+ let term = snd (Reductionops.splay_lam_n (Global.env ()) Evd.empty len subs) in
+ let typ = snd (Reductionops.splay_prod_n (Global.env ()) Evd.empty len typ) in
+ x.prg_reduce term, x.prg_reduce typ, x.prg_implicits) l)
+ in
+(* let fixdefs = List.map reduce_fix fixdefs in *)
+ let fixkind = Option.get first.prg_fixkind in
+ let arrrec, recvec = Array.of_list fixtypes, Array.of_list fixdefs in
+ let fixdecls = (Array.of_list (List.map (fun x -> Name x.prg_name) l), arrrec, recvec) in
+ let (local,poly,kind) = first.prg_kind in
+ let fixnames = first.prg_deps in
+ let kind = if fixkind != IsCoFixpoint then Fixpoint else CoFixpoint in
+ let indexes, fixdecls =
+ match fixkind with
+ | IsFixpoint wfl ->
+ let possible_indexes =
+ List.map3 compute_possible_guardness_evidences
+ wfl fixdefs fixtypes in
+ let indexes =
+ Pretyping.search_guard Loc.ghost (Global.env())
+ possible_indexes fixdecls in
+ Some indexes,
+ List.map_i (fun i _ ->
+ mk_proof (mkFix ((indexes,i),fixdecls))) 0 l
+ | IsCoFixpoint ->
+ None,
+ List.map_i (fun i _ ->
+ mk_proof (mkCoFix (i,fixdecls))) 0 l
+ in
+ (* Declare the recursive definitions *)
+ let ctx = Evd.evar_context_universe_context first.prg_ctx in
+ let kns = List.map4 (!declare_fix_ref (local, poly, kind) ctx)
+ fixnames fixdecls fixtypes fiximps in
+ (* Declare notations *)
+ List.iter Metasyntax.add_notation_interpretation first.prg_notations;
+ Declare.recursive_message (fixkind != IsCoFixpoint) indexes fixnames;
+ let gr = List.hd kns in
+ let kn = match gr with ConstRef kn -> kn | _ -> assert false in
+ Lemmas.call_hook (fun exn -> exn) first.prg_hook local gr;
+ 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 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 -> { obl with obl_body = Some (TermObl body) }
+ | Evar_kinds.Define opaque ->
+ let opaque = if get_proofs_transparency () then false else opaque in
+ let poly = pi2 prg.prg_kind in
+ let ctx, body, args =
+ if get_shrink_obligations () && not poly then
+ shrink_body body else [], body, [||]
+ in
+ let ce =
+ { const_entry_body = Future.from_val((body,Univ.ContextSet.empty),Declareops.no_seff);
+ const_entry_secctx = None;
+ const_entry_type = if List.is_empty ctx then ty else None;
+ const_entry_polymorphic = poly;
+ const_entry_universes = uctx;
+ const_entry_opaque = opaque;
+ const_entry_inline_code = false;
+ const_entry_feedback = None;
+ } in
+ (** ppedrot: seems legit to have obligations as local *)
+ let constant = Declare.declare_constant obl.obl_name ~local:true
+ (DefinitionEntry ce,IsProof Property)
+ in
+ if not opaque then
+ Hints.add_hints false [Id.to_string prg.prg_name]
+ (Hints.HintsUnfoldEntry [EvalConstRef constant]);
+ definition_message obl.obl_name;
+ { obl with obl_body =
+ if poly then
+ Some (DefinedObl constant)
+ else
+ Some (TermObl (it_mkLambda_or_LetIn (mkApp (mkConst constant, args)) ctx)) }
+
+let init_prog_info n 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_tac = None } |],
+ mkVar n
+ | Some b ->
+ Array.mapi
+ (fun i (n, t, l, o, d, tac) ->
+ { obl_name = n ; obl_body = None;
+ obl_location = l; obl_type = reduce t; obl_status = o;
+ obl_deps = d; obl_tac = tac })
+ obls, b
+ in
+ { prg_name = n ; prg_body = b; prg_type = reduce t;
+ prg_ctx = ctx;
+ prg_obligations = (obls', Array.length obls');
+ prg_deps = deps; prg_fixkind = fixkind ; prg_notations = notations ;
+ prg_implicits = impls; prg_kind = kind; prg_reduce = reduce;
+ prg_hook = hook; }
+
+let get_prog name =
+ let prg_infos = !from_prg in
+ match name with
+ Some n ->
+ (try ProgMap.find n prg_infos
+ with Not_found -> raise (NoObligations (Some n)))
+ | None ->
+ (let n = map_cardinal prg_infos in
+ match n with
+ 0 -> raise (NoObligations None)
+ | 1 -> 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 get_any_prog () =
+ let prg_infos = !from_prg in
+ let n = map_cardinal prg_infos in
+ if n > 0 then map_first prg_infos
+ else raise (NoObligations None)
+
+let get_prog_err n =
+ try get_prog n with NoObligations id -> pperror (explain_no_obligations id)
+
+let get_any_prog_err () =
+ try get_any_prog () with NoObligations id -> pperror (explain_no_obligations id)
+
+let obligations_solved prg = Int.equal (snd prg.prg_obligations) 0
+
+let all_programs () =
+ ProgMap.fold (fun k p l -> p :: l) !from_prg []
+
+type progress =
+ | Remain of int
+ | Dependent
+ | Defined of global_reference
+
+let obligations_message rem =
+ if rem > 0 then
+ if Int.equal rem 1 then
+ Flags.if_verbose msg_info (int rem ++ str " obligation remaining")
+ else
+ Flags.if_verbose msg_info (int rem ++ str " obligations remaining")
+ else
+ Flags.if_verbose msg_info (str "No more obligations remaining")
+
+let update_obls prg obls rem =
+ let prg' = { prg with prg_obligations = (obls, rem) } in
+ progmap_replace prg';
+ obligations_message rem;
+ if rem > 0 then Remain rem
+ else (
+ match prg'.prg_deps with
+ | [] ->
+ let kn = declare_definition prg' in
+ progmap_remove prg';
+ Defined kn
+ | l ->
+ let progs = List.map (fun x -> ProgMap.find x !from_prg) prg'.prg_deps in
+ if List.for_all (fun x -> obligations_solved x) progs then
+ let kn = declare_mutual_definition progs in
+ Defined (ConstRef kn)
+ else Dependent)
+
+let is_defined obls x = not (Option.is_empty obls.(x).obl_body)
+
+let deps_remaining obls deps =
+ Int.Set.fold
+ (fun x acc ->
+ if is_defined obls x then acc
+ else x :: acc)
+ deps []
+
+let dependencies obls n =
+ let res = ref Int.Set.empty in
+ Array.iteri
+ (fun i obl ->
+ if not (Int.equal i n) && Int.Set.mem n obl.obl_deps then
+ res := Int.Set.add i !res)
+ obls;
+ !res
+
+let goal_kind poly = Decl_kinds.Local, poly, Decl_kinds.DefinitionBody Decl_kinds.Definition
+
+let goal_proof_kind poly = Decl_kinds.Local, poly, Decl_kinds.Proof Decl_kinds.Lemma
+
+let kind_of_obligation poly o =
+ match o with
+ | Evar_kinds.Define false | Evar_kinds.Expand -> goal_kind poly
+ | _ -> goal_proof_kind poly
+
+let not_transp_msg =
+ str "Obligation should be transparent but was declared opaque." ++ spc () ++
+ str"Use 'Defined' instead."
+
+let error_not_transp () = pperror not_transp_msg
+
+let rec string_of_list sep f = function
+ [] -> ""
+ | x :: [] -> f x
+ | x :: ((y :: _) as tl) -> f x ^ sep ^ string_of_list sep f tl
+
+(* Solve an obligation using tactics, return the corresponding proof term *)
+
+let solve_by_tac name evi t poly ctx =
+ let id = name in
+ let concl = evi.evar_concl in
+ (* spiwack: the status is dropped. *)
+ let (entry,_,ctx') = Pfedit.build_constant_by_tactic
+ id ~goal_kind:(goal_kind poly) ctx evi.evar_hyps concl (Tacticals.New.tclCOMPLETE t) in
+ let env = Global.env () in
+ let entry = Term_typing.handle_entry_side_effects env entry in
+ let body, eff = Future.force entry.Entries.const_entry_body in
+ assert(Declareops.side_effects_is_empty eff);
+ assert(Univ.ContextSet.is_empty (snd body));
+ Inductiveops.control_only_guard (Global.env ()) (fst body);
+ (fst body), entry.Entries.const_entry_type, ctx'
+
+let rec solve_obligation prg num tac =
+ let user_num = succ num in
+ let obls, rem = prg.prg_obligations in
+ let obl = obls.(num) in
+ if not (Option.is_empty obl.obl_body) then
+ pperror (str "Obligation" ++ spc () ++ int user_num ++ str "already" ++ spc() ++ str "solved.")
+ else
+ match deps_remaining obls obl.obl_deps with
+ | [] ->
+ let obl = subst_deps_obl obls obl in
+ let kind = kind_of_obligation (pi2 prg.prg_kind) obl.obl_status in
+ let evd = Evd.from_env ~ctx:prg.prg_ctx Environ.empty_env in
+ Lemmas.start_proof_univs obl.obl_name kind evd obl.obl_type
+ (fun ctx' -> Lemmas.mk_hook (fun strength gr ->
+ let cst = match gr with ConstRef cst -> cst | _ -> assert false in
+ let obl =
+ let transparent = evaluable_constant cst (Global.env ()) in
+ let body =
+ match obl.obl_status with
+ | Evar_kinds.Expand ->
+ if not transparent then error_not_transp ()
+ else DefinedObl cst
+ | Evar_kinds.Define opaque ->
+ if not opaque && not transparent then error_not_transp ()
+ else DefinedObl cst
+ in
+ if transparent then
+ Hints.add_hints true [Id.to_string prg.prg_name]
+ (Hints.HintsUnfoldEntry [EvalConstRef cst]);
+ { obl with obl_body = Some body }
+ in
+ let obls = Array.copy obls in
+ let _ = obls.(num) <- obl in
+ let ctx' =
+ let ctx =
+ match ctx' with
+ | None -> prg.prg_ctx
+ | Some ctx' -> ctx'
+ in
+ if not (pi2 prg.prg_kind) (* Not polymorphic *) then
+ (* This context is already declared globally, we cannot
+ instantiate the rigid variables anymore *)
+ Evd.abstract_undefined_variables ctx
+ else ctx
+ in
+ let res =
+ try update_obls
+ {prg with prg_body = prg.prg_body;
+ prg_type = prg.prg_type;
+ prg_ctx = ctx' }
+
+ obls (pred rem)
+ with e when Errors.noncritical e ->
+ let e = Errors.push e in
+ pperror (Errors.iprint (Cerrors.process_vernac_interp_error e))
+ in
+ match res with
+ | Remain n when n > 0 ->
+ let deps = dependencies obls num in
+ if not (Int.Set.is_empty deps) then
+ ignore(auto_solve_obligations (Some prg.prg_name) None ~oblset:deps)
+ | _ -> ()));
+ trace (str "Started obligation " ++ int user_num ++ str " proof: " ++
+ Printer.pr_constr_env (Global.env ()) Evd.empty obl.obl_type);
+ ignore (Pfedit.by (snd (get_default_tactic ())));
+ Option.iter (fun tac -> Pfedit.set_end_tac tac) tac
+ | l -> pperror (str "Obligation " ++ int user_num ++ str " depends on obligation(s) "
+ ++ str (string_of_list ", " (fun x -> string_of_int (succ x)) l))
+
+and obligation (user_num, name, typ) tac =
+ let num = pred user_num in
+ let prg = get_prog_err name in
+ let obls, rem = prg.prg_obligations in
+ if num < Array.length obls then
+ let obl = obls.(num) in
+ match obl.obl_body with
+ None -> solve_obligation prg num tac
+ | Some r -> error "Obligation already solved"
+ else error (sprintf "Unknown obligation number %i" (succ num))
+
+
+and solve_obligation_by_tac prg obls i tac =
+ let obl = obls.(i) in
+ match obl.obl_body with
+ | Some _ -> false
+ | None ->
+ try
+ if List.is_empty (deps_remaining obls obl.obl_deps) then
+ let obl = subst_deps_obl obls obl in
+ let tac =
+ match tac with
+ | Some t -> t
+ | None ->
+ match obl.obl_tac with
+ | Some t -> t
+ | None -> snd (get_default_tactic ())
+ in
+ let t, ty, ctx =
+ solve_by_tac obl.obl_name (evar_of_obligation obl) tac
+ (pi2 !prg.prg_kind) !prg.prg_ctx
+ in
+ let uctx = Evd.evar_context_universe_context ctx in
+ prg := {!prg with prg_ctx = ctx};
+ obls.(i) <- declare_obligation !prg obl t ty uctx;
+ true
+ else false
+ with e when Errors.noncritical e ->
+ let (e, _) = Errors.push e in
+ match e with
+ | Refiner.FailError (_, s) ->
+ user_err_loc (fst obl.obl_location, "solve_obligation", Lazy.force s)
+ | e -> false
+
+and solve_prg_obligations prg ?oblset tac =
+ let obls, rem = prg.prg_obligations in
+ let rem = ref rem in
+ let obls' = Array.copy obls in
+ let set = ref Int.Set.empty in
+ let p = match oblset with
+ | None -> (fun _ -> true)
+ | Some s -> set := s;
+ (fun i -> Int.Set.mem i !set)
+ in
+ let prg = 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))
+ obls'
+ in
+ update_obls !prg obls' !rem
+
+and solve_obligations n tac =
+ let prg = get_prog_err n in
+ solve_prg_obligations prg tac
+
+and solve_all_obligations tac =
+ ProgMap.iter (fun k v -> ignore(solve_prg_obligations v tac)) !from_prg
+
+and try_solve_obligation n prg tac =
+ let prg = get_prog prg in
+ let obls, rem = prg.prg_obligations in
+ let obls' = Array.copy obls in
+ let prgref = ref prg in
+ if solve_obligation_by_tac prgref obls' n tac then
+ ignore(update_obls !prgref obls' (pred rem));
+
+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...");
+ try solve_prg_obligations (get_prog_err n) ?oblset tac with NoObligations _ -> Dependent
+
+open Pp
+let show_obligations_of_prg ?(msg=true) prg =
+ let n = prg.prg_name in
+ let obls, rem = prg.prg_obligations in
+ let showed = ref 5 in
+ if msg then 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 () ++
+ hov 1 (Printer.pr_constr_env (Global.env ()) Evd.empty x.obl_type ++
+ str "." ++ fnl ())))
+ | Some _ -> ())
+ obls
+
+let show_obligations ?(msg=true) n =
+ let progs = match n with
+ | None -> all_programs ()
+ | Some n ->
+ try [ProgMap.find n !from_prg]
+ with Not_found -> raise (NoObligations (Some n))
+ in List.iter (show_obligations_of_prg ~msg) progs
+
+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 () ++
+ 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
+ ?(reduce=reduce) ?(hook=Lemmas.mk_hook (fun _ _ -> ())) obls =
+ let info = str (Id.to_string n) ++ str " has type-checked" in
+ let prg = init_prog_info n 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 ".");
+ 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 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)
+ ?(hook=Lemmas.mk_hook (fun _ _ -> ())) notations fixkind =
+ 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 n (Some b) t ctx deps (Some fixkind)
+ notations obls imps kind reduce hook
+ in progmap_add n prg) l;
+ let _defined =
+ List.fold_left (fun finished x ->
+ if finished then finished
+ else
+ let res = auto_solve_obligations (Some x) tactic in
+ match res with
+ | Defined _ ->
+ (* If one definition is turned into a constant,
+ the whole block is defined. *) true
+ | _ -> false)
+ false deps
+ in ()
+
+let admit_prog prg =
+ let obls, rem = prg.prg_obligations in
+ let obls = Array.copy obls in
+ Array.iteri
+ (fun i x ->
+ match x.obl_body with
+ | None ->
+ let x = subst_deps_obl obls x in
+ let ctx = Evd.evar_context_universe_context prg.prg_ctx in
+ let kn = Declare.declare_constant x.obl_name ~local:true
+ (ParameterEntry (None,false,(x.obl_type,ctx),None), IsAssumption Conjectural)
+ in
+ assumption_message x.obl_name;
+ obls.(i) <- { x with obl_body = Some (DefinedObl kn) }
+ | Some _ -> ())
+ obls;
+ ignore(update_obls prg obls 0)
+
+let rec admit_all_obligations () =
+ let prg = try Some (get_any_prog ()) with NoObligations _ -> None in
+ match prg with
+ | None -> ()
+ | Some prg ->
+ admit_prog prg;
+ admit_all_obligations ()
+
+let admit_obligations n =
+ match n with
+ | None -> admit_all_obligations ()
+ | Some _ ->
+ let prg = get_prog_err n in
+ admit_prog prg
+
+let next_obligation n tac =
+ let prg = match n with
+ | None -> get_any_prog_err ()
+ | Some _ -> get_prog_err n
+ in
+ let obls, rem = prg.prg_obligations in
+ let is_open _ x = Option.is_empty x.obl_body && List.is_empty (deps_remaining obls x.obl_deps) in
+ let i = match Array.findi is_open obls with
+ | Some i -> i
+ | None -> anomaly (Pp.str "Could not find a solvable obligation.")
+ in
+ solve_obligation prg i tac
+
+let init_program () =
+ Coqlib.check_required_library Coqlib.datatypes_module_name;
+ Coqlib.check_required_library ["Coq";"Init";"Specif"];
+ Coqlib.check_required_library ["Coq";"Program";"Tactics"]
+
+
+let set_program_mode c =
+ if c then
+ if !Flags.program_mode then ()
+ else begin
+ init_program ();
+ Flags.program_mode := true;
+ end
diff --git a/toplevel/obligations.mli b/toplevel/obligations.mli
new file mode 100644
index 00000000..582b4935
--- /dev/null
+++ b/toplevel/obligations.mli
@@ -0,0 +1,116 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Environ
+open Term
+open Evd
+open Names
+open Pp
+open Globnames
+open Vernacexpr
+open Decl_kinds
+open Tacexpr
+
+(** Forward declaration. *)
+val declare_fix_ref : (definition_kind -> Univ.universe_context -> Id.t ->
+ Entries.proof_output -> types -> Impargs.manual_implicits -> global_reference) ref
+
+val declare_definition_ref :
+ (Id.t -> definition_kind ->
+ Entries.definition_entry -> Impargs.manual_implicits
+ -> global_reference Lemmas.declaration_hook -> global_reference) ref
+
+val check_evars : env -> evar_map -> unit
+
+val evar_dependencies : evar_map -> Evar.t -> Evar.Set.t
+val sort_dependencies : (Evar.t * evar_info * Evar.Set.t) list -> (Evar.t * evar_info * Evar.Set.t) list
+
+(* env, id, evars, number of function prototypes to try to clear from
+ evars contexts, object and type *)
+val eterm_obligations : env -> Id.t -> evar_map -> int ->
+ ?status:Evar_kinds.obligation_definition_status -> constr -> types ->
+ (Id.t * types * Evar_kinds.t Loc.located * Evar_kinds.obligation_definition_status * Int.Set.t *
+ unit Proofview.tactic option) array
+ (* Existential key, obl. name, type as product,
+ location of the original evar, associated tactic,
+ status and dependencies as indexes into the array *)
+ * ((existential_key * Id.t) list * ((Id.t -> constr) -> constr -> constr)) *
+ constr * types
+ (* Translations from existential identifiers to obligation identifiers
+ and for terms with existentials to closed terms, given a
+ translation from obligation identifiers to constrs, new term, new type *)
+
+type obligation_info =
+ (Id.t * Term.types * Evar_kinds.t Loc.located *
+ Evar_kinds.obligation_definition_status * Int.Set.t * unit Proofview.tactic option) array
+ (* ident, type, location, (opaque or transparent, expand or define),
+ dependencies, tactic to solve it *)
+
+type progress = (* Resolution status of a program *)
+ | Remain of int (* n obligations remaining *)
+ | Dependent (* Dependent on other definitions *)
+ | Defined of global_reference (* Defined as id *)
+
+val 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 add_definition : Names.Id.t -> ?term:Term.constr -> Term.types ->
+ Evd.evar_universe_context ->
+ ?implicits:(Constrexpr.explicitation * (bool * bool * bool)) list ->
+ ?kind:Decl_kinds.definition_kind ->
+ ?tactic:unit Proofview.tactic ->
+ ?reduce:(Term.constr -> Term.constr) ->
+ ?hook:unit Lemmas.declaration_hook -> obligation_info -> progress
+
+type notations =
+ (Vernacexpr.lstring * Constrexpr.constr_expr * Notation_term.scope_name option) list
+
+type fixpoint_kind =
+ | IsFixpoint of (Id.t Loc.located option * Constrexpr.recursion_order_expr) list
+ | IsCoFixpoint
+
+val add_mutual_definitions :
+ (Names.Id.t * Term.constr * Term.types *
+ (Constrexpr.explicitation * (bool * bool * bool)) list * obligation_info) list ->
+ Evd.evar_universe_context ->
+ ?tactic:unit Proofview.tactic ->
+ ?kind:Decl_kinds.definition_kind ->
+ ?reduce:(Term.constr -> Term.constr) ->
+ ?hook:unit Lemmas.declaration_hook ->
+ notations ->
+ fixpoint_kind -> unit
+
+val obligation : int * Names.Id.t option * Constrexpr.constr_expr option ->
+ Tacexpr.raw_tactic_expr option -> unit
+
+val next_obligation : Names.Id.t option -> Tacexpr.raw_tactic_expr option -> unit
+
+val solve_obligations : Names.Id.t option -> unit Proofview.tactic option -> progress
+(* Number of remaining obligations to be solved for this program *)
+
+val solve_all_obligations : unit Proofview.tactic option -> unit
+
+val try_solve_obligation : int -> Names.Id.t option -> unit Proofview.tactic option -> unit
+
+val try_solve_obligations : Names.Id.t option -> unit Proofview.tactic option -> unit
+
+val show_obligations : ?msg:bool -> Names.Id.t option -> unit
+
+val show_term : Names.Id.t option -> std_ppcmds
+
+val admit_obligations : Names.Id.t option -> unit
+
+exception NoObligations of Names.Id.t option
+
+val explain_no_obligations : Names.Id.t option -> Pp.std_ppcmds
+
+val set_program_mode : bool -> unit
diff --git a/toplevel/record.ml b/toplevel/record.ml
index 86a9411f..55f53351 100644
--- a/toplevel/record.ml
+++ b/toplevel/record.ml
@@ -1,107 +1,169 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Pp
+open Errors
open Util
open Names
-open Libnames
+open Globnames
open Nameops
open Term
+open Context
+open Vars
open Environ
open Declarations
open Entries
open Declare
-open Nametab
open Constrintern
-open Command
-open Inductive
-open Safe_typing
open Decl_kinds
-open Indtypes
open Type_errors
-open Topconstr
+open Constrexpr
+open Constrexpr_ops
+open Goptions
(********** definition d'un record (structure) **************)
-let interp_evars evdref env impls k typ =
- let typ' = intern_gen true ~impls !evdref env typ in
- let imps = Implicit_quantifiers.implicits_of_glob_constr typ' in
- imps, Pretyping.Default.understand_tcc_evars evdref env k typ'
+(** Flag governing use of primitive projections. Disabled by default. *)
+let primitive_flag = ref false
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "use of primitive projections";
+ optkey = ["Primitive";"Projections"];
+ optread = (fun () -> !primitive_flag) ;
+ optwrite = (fun b -> primitive_flag := b) }
-let interp_fields_evars evars env impls_env nots l =
+let typeclasses_strict = ref false
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "strict typeclass resolution";
+ optkey = ["Typeclasses";"Strict";"Resolution"];
+ optread = (fun () -> !typeclasses_strict);
+ optwrite = (fun b -> typeclasses_strict := b); }
+
+let typeclasses_unique = ref false
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "unique typeclass instances";
+ optkey = ["Typeclasses";"Unique";"Instances"];
+ optread = (fun () -> !typeclasses_unique);
+ optwrite = (fun b -> typeclasses_unique := b); }
+
+let interp_fields_evars env evars impls_env nots l =
List.fold_left2
(fun (env, uimpls, params, impls) no ((loc, i), b, t) ->
- let impl, t' = interp_evars evars env impls Pretyping.IsType t in
- let b' = Option.map (fun x -> snd (interp_evars evars env impls (Pretyping.OfType (Some t')) x)) b in
+ let t', impl = interp_type_evars_impls env evars ~impls t in
+ let b' = Option.map (fun x -> fst (interp_casted_constr_evars_impls env evars ~impls x t')) b in
let impls =
match i with
| Anonymous -> impls
- | Name id -> Idmap.add id (compute_internalization_data env Constrintern.Method t' impl) impls
+ | Name id -> Id.Map.add id (compute_internalization_data env Constrintern.Method t' impl) impls
in
let d = (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) ->
+ let univ =
+ if b = None then
+ let s = Retyping.get_sort_of env evars t in
+ Univ.sup (univ_of_sort s) univ
+ else univ
+ in (push_rel d env, univ))
+ l (env, Univ.type0m_univ)
+
let binder_of_decl = function
| Vernacexpr.AssumExpr(n,t) -> (n,None,t)
- | Vernacexpr.DefExpr(n,c,t) -> (n,Some c, match t with Some c -> c | None -> CHole (fst n, None))
+ | Vernacexpr.DefExpr(n,c,t) -> (n,Some c, match t with Some c -> c | None -> CHole (fst n, None, Misctypes.IntroAnonymous, None))
let binders_of_decls = List.map binder_of_decl
-let typecheck_params_and_fields id t ps nots fs =
+let typecheck_params_and_fields def id t ps nots fs =
let env0 = Global.env () in
- let evars = ref Evd.empty in
+ let evars = ref (Evd.from_env env0) in
let _ =
let error bk (loc, name) =
- match bk with
- | Default _ ->
- if name = Anonymous then
- user_err_loc (loc, "record", str "Record parameters must be named")
+ match bk, name with
+ | Default _, Anonymous ->
+ user_err_loc (loc, "record", str "Record parameters must be named")
| _ -> ()
in
List.iter
(function LocalRawDef (b, _) -> error default_binder_kind b
| LocalRawAssum (ls, bk, ce) -> List.iter (error bk) ls) ps
in
- let impls_env, ((env1,newps), imps) = interp_context_evars evars env0 ps in
- let fullarity = it_mkProd_or_LetIn (Option.cata (fun x -> x) (Termops.new_Type ()) t) newps 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 s = interp_type_evars env evars ~impls:empty_internalization_env t in
+ let sred = Reductionops.whd_betadeltaiota 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)
+ | _ -> 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
+ 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 env2,impls,newfs,data =
- interp_fields_evars evars env_ar impls_env nots (binders_of_decls fs)
+ interp_fields_evars env_ar evars impls_env nots (binders_of_decls fs)
in
- let evars = Evarconv.consider_remaining_unif_problems env_ar !evars in
- let evars = Typeclasses.resolve_typeclasses env_ar evars in
- let sigma = evars in
- let newps = Evarutil.nf_rel_context_evar sigma newps in
- let newfs = Evarutil.nf_rel_context_evar sigma newfs in
+ let sigma =
+ Pretyping.solve_remaining_evars Pretyping.all_and_fail_flags env_ar !evars (Evd.empty,!evars) in
+ let evars, nf = Evarutil.nf_evars_and_universes sigma in
+ let arity = nf t' in
+ let evars =
+ let _, univ = compute_constructor_level evars env_ar newfs in
+ let ctx, aritysort = Reduction.dest_arity env0 arity in
+ assert(List.is_empty ctx); (* Ensured by above analysis *)
+ if Sorts.is_prop aritysort ||
+ (Sorts.is_set aritysort && engagement env0 = Some ImpredicativeSet) then
+ evars
+ else Evd.set_leq_sort env_ar evars (Type univ) aritysort
+ 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);
- imps, newps, impls, newfs
+ Evd.universe_context evars, nf arity, template, imps, newps, impls, newfs
let degenerate_decl (na,b,t) =
let id = match na with
| Name id -> id
- | Anonymous -> anomaly "Unnamed record variable" in
+ | Anonymous -> anomaly (Pp.str "Unnamed record variable") in
match b with
| None -> (id, Entries.LocalAssum t)
| Some b -> (id, Entries.LocalDef b)
type record_error =
- | MissingProj of identifier * identifier list
- | BadTypedProj of identifier * env * Type_errors.type_error
+ | MissingProj of Id.t * Id.t list
+ | BadTypedProj of Id.t * env * Type_errors.type_error
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(string_of_id fi) ++
+ (str(Id.to_string 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.")
@@ -121,10 +183,10 @@ 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 ppnl (hov 0 (str"Warning: " ++ st))
+ Flags.if_verbose msg_warning (hov 0 st)
type field_status =
- | NoProjection of name
+ | NoProjection of Name.t
| Projection of constr
exception NotDefinable of record_error
@@ -148,7 +210,7 @@ let subst_projection fid l c =
| NoProjection (Name id) -> bad_projs := id :: !bad_projs; mkRel k
| NoProjection Anonymous ->
errorlabstrm "" (str "Field " ++ pr_id fid ++
- str " depends on the " ++ str (ordinal (k-depth-1)) ++ str
+ str " depends on the " ++ pr_nth (k-depth-1) ++ str
" field which has no name.")
else
mkRel (k-lv)
@@ -156,88 +218,121 @@ let subst_projection fid l c =
in
let c' = lift 1 c in (* to get [c] defined in ctxt [[params;fields;x:ind]] *)
let c'' = substrec 0 c' in
- if !bad_projs <> [] then
+ if not (List.is_empty !bad_projs) then
raise (NotDefinable (MissingProj (fid,List.rev !bad_projs)));
c''
-let instantiate_possibly_recursive_type indsp paramdecls fields =
- let subst = list_map_i (fun i _ -> mkRel i) 1 paramdecls in
- Termops.substl_rel_context (subst@[mkInd indsp]) fields
+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
(* We build projections *)
-let declare_projections indsp ?(kind=StructureComponent) ?name coers fieldimpls fields =
+let declare_projections indsp ?(kind=StructureComponent) binder_name coers fieldimpls fields =
let env = Global.env() in
let (mib,mip) = Global.lookup_inductive indsp in
- let paramdecls = mib.mind_params_ctxt in
- let r = mkInd indsp in
+ let u = Declareops.inductive_instance mib in
+ let paramdecls = Inductive.inductive_paramdecls (mib, u) in
+ let poly = mib.mind_polymorphic and 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 x = match name with Some n -> Name n | None -> Namegen.named_hd (Global.env()) r Anonymous in
- let fields = instantiate_possibly_recursive_type indsp paramdecls fields in
+ let x = Name binder_name in
+ let fields = instantiate_possibly_recursive_type indu paramdecls fields in
let lifted_fields = Termops.lift_rel_context 1 fields in
- let (_,kinds,sp_projs,_) =
- list_fold_left3
- (fun (nfi,kinds,sp_projs,subst) coe (fi,optci,ti) impls ->
- let (sp_projs,subst) =
+ let primitive =
+ if !primitive_flag then
+ let is_primitive =
+ match mib.mind_record with
+ | Some (Some _) -> true
+ | Some None | None -> false
+ in
+ if not is_primitive then
+ Flags.if_verbose msg_warning
+ (hov 0 (str "The record " ++ Printer.pr_inductive env indsp ++
+ str" could not be defined as a primitive record"));
+ is_primitive
+ else false
+ in
+ let (_,_,kinds,sp_projs,_) =
+ List.fold_left3
+ (fun (nfi,i,kinds,sp_projs,subst) coe (fi,optci,ti) impls ->
+ let (sp_projs,i,subst) =
match fi with
| Anonymous ->
- (None::sp_projs,NoProjection fi::subst)
- | Name fid ->
- try
- let ccl = subst_projection fid subst ti in
- let body = match optci with
- | Some ci -> subst_projection fid subst ci
- | None ->
- (* [ccl] is defined in context [params;x:rp] *)
- (* [ccl'] is defined in context [params;x:rp;x:rp] *)
- let ccl' = liftn 1 2 ccl in
- let p = mkLambda (x, lift 1 rp, ccl') in
- let branch = it_mkLambda_or_LetIn (mkRel nfi) lifted_fields in
- let ci = Inductiveops.make_case_info env indsp LetStyle in
- mkCase (ci, p, mkRel 1, [|branch|]) in
- let proj =
- it_mkLambda_or_LetIn (mkLambda (x,rp,body)) paramdecls in
- let projtyp =
- it_mkProd_or_LetIn (mkProd (x,rp,ccl)) paramdecls in
- let kn =
+ (None::sp_projs,i,NoProjection fi::subst)
+ | Name fid -> try
+ let kn, term =
+ if optci = None && 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 ->
+ (* [ccl] is defined in context [params;x:rp] *)
+ (* [ccl'] is defined in context [params;x:rp;x:rp] *)
+ let ccl' = liftn 1 2 ccl in
+ let p = mkLambda (x, lift 1 rp, ccl') in
+ let branch = it_mkLambda_or_LetIn (mkRel nfi) lifted_fields in
+ let ci = Inductiveops.make_case_info env indsp LetStyle in
+ mkCase (ci, p, mkRel 1, [|branch|])
+ in
+ let proj =
+ it_mkLambda_or_LetIn (mkLambda (x,rp,body)) paramdecls in
+ let projtyp =
+ it_mkProd_or_LetIn (mkProd (x,rp,ccl)) paramdecls in
try
- let cie = {
- const_entry_body = proj;
- const_entry_secctx = None;
- const_entry_type = Some projtyp;
- const_entry_opaque = false } in
- let k = (DefinitionEntry cie,IsDefinition kind) in
+ let entry = {
+ const_entry_body =
+ Future.from_val (Term_typing.mk_pure_proof proj);
+ const_entry_secctx = None;
+ const_entry_type = Some projtyp;
+ const_entry_polymorphic = poly;
+ const_entry_universes = ctx;
+ const_entry_opaque = false;
+ const_entry_inline_code = false;
+ const_entry_feedback = None } in
+ let k = (DefinitionEntry entry,IsDefinition kind) in
let kn = declare_constant ~internal:KernelSilent fid k in
- Flags.if_verbose message (string_of_id fid ^" is defined");
- kn
+ let constr_fip =
+ let proj_args = (*Rel 1 refers to "x"*) paramargs@[mkRel 1] in
+ applist (mkConstU (kn,u),proj_args)
+ in
+ Declare.definition_message fid;
+ kn, constr_fip
with Type_errors.TypeError (ctx,te) ->
- raise (NotDefinable (BadTypedProj (fid,ctx,te))) in
- let refi = ConstRef kn in
- let constr_fi = mkConst kn in
- Impargs.maybe_declare_manual_implicits false refi impls;
- if coe then begin
- let cl = Class.class_of_global (IndRef indsp) in
- Class.try_add_new_coercion_with_source refi Global ~source:cl
- end;
- let proj_args = (*Rel 1 refers to "x"*) paramargs@[mkRel 1] in
- let constr_fip = applist (constr_fi,proj_args) in
- (Some kn::sp_projs, Projection constr_fip::subst)
+ raise (NotDefinable (BadTypedProj (fid,ctx,te)))
+ in
+ let refi = ConstRef kn in
+ Impargs.maybe_declare_manual_implicits false refi impls;
+ if coe then begin
+ let cl = Class.class_of_global (IndRef indsp) in
+ Class.try_add_new_coercion_with_source refi ~local:false poly ~source:cl
+ end;
+ let i = if Option.is_empty optci 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,NoProjection fi::subst) in
- (nfi-1,(fi, optci=None)::kinds,sp_projs,subst))
- (List.length fields,[],[],[]) coers (List.rev fields) (List.rev fieldimpls)
+ (None::sp_projs,i,NoProjection fi::subst) in
+ (nfi-1,i,(fi, Option.is_empty optci)::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)] -> Evd.add evm (Evarutil.new_untyped_evar())
- (Evd.make_evar Environ.empty_named_context_val typ)
+ | [(_,_,typ)] ->
+ let env = Environ.empty_named_context_val in
+ let (evm, _) = Evarutil.new_pure_evar env evm typ in
+ evm
| (_,_,typ)::tl ->
- let ev = Evarutil.new_untyped_evar() in
- let evm = Evd.add evm ev (Evd.make_evar Environ.empty_named_context_val typ) in
- let new_tl = Util.list_map_i
+ let env = Environ.empty_named_context_val in
+ let (evm, ev) = Evarutil.new_pure_evar env evm typ 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
deps_to_evar evm new_tl in
@@ -245,45 +340,43 @@ let structure_signature ctx =
open Typeclasses
-let declare_structure finite infer id idbuild paramimpls params arity fieldimpls fields
- ?(kind=StructureComponent) ?name is_coe coers sign =
+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 ind = applist (mkRel (1+nparams+nfields), args) in
let type_constructor = it_mkProd_or_LetIn ind fields in
+ let binder_name =
+ match name with
+ | None -> Id.of_string (Unicode.lowercase_first_char (Id.to_string id))
+ | Some n -> n
+ in
let mie_ind =
{ mind_entry_typename = id;
mind_entry_arity = arity;
+ mind_entry_template = not poly && template;
mind_entry_consnames = [idbuild];
mind_entry_lc = [type_constructor] }
in
- (* spiwack: raises an error if the structure is supposed to be non-recursive,
- but isn't *)
- (* there is probably a way to push this to "declare_mutual" *)
- begin match finite with
- | BiFinite ->
- if Termops.dependent (mkRel (nparams+1)) (it_mkProd_or_LetIn mkProp fields) then
- error "Records declared with the keyword Record or Structure cannot be recursive. You can, however, define recursive records using the Inductive or CoInductive command."
- | _ -> ()
- end;
let mie =
{ mind_entry_params = List.map degenerate_decl params;
- mind_entry_record = true;
- mind_entry_finite = recursivity_flag_of_kind finite;
- mind_entry_inds = [mie_ind] } in
- let kn = Command.declare_mutual_inductive_with_eliminations KernelVerbose mie [(paramimpls,[])] in
+ mind_entry_record = Some (if !primitive_flag then Some binder_name else None);
+ mind_entry_finite = finite;
+ mind_entry_inds = [mie_ind];
+ mind_entry_polymorphic = poly;
+ mind_entry_private = None;
+ mind_entry_universes = ctx } in
+ let kn = Command.declare_mutual_inductive_with_eliminations mie [(paramimpls,[])] in
let rsp = (kn,0) in (* This is ind path of idstruc *)
let cstr = (rsp,1) in
- let kinds,sp_projs = declare_projections rsp ~kind ?name coers fieldimpls fields in
+ let kinds,sp_projs = declare_projections rsp ~kind binder_name coers fieldimpls fields in
let build = ConstructRef cstr in
- if is_coe then Class.try_add_new_coercion build Global;
+ let () = if is_coe then Class.try_add_new_coercion build ~local:false poly in
Recordops.declare_structure(rsp,cstr,List.rev kinds,List.rev sp_projs);
- if infer then
- Evd.fold (fun ev evi () -> Recordops.declare_method (ConstructRef cstr) ev sign) sign ();
rsp
let implicits_of_context ctx =
- list_map_i (fun i name ->
+ List.map_i (fun i name ->
let explname =
match name with
| Name n -> Some n
@@ -291,44 +384,32 @@ let implicits_of_context ctx =
in ExplByPos (i, explname), (true, true, true))
1 (List.rev (Anonymous :: (List.map pi1 ctx)))
-let declare_instance_cst glob con pri =
- let instance = Typeops.type_of_constant (Global.env ()) con in
- let _, r = decompose_prod_assum instance in
- match class_of_constr r with
- | Some (_, (tc, _)) -> add_instance (new_instance tc pri glob (ConstRef con))
- | None -> errorlabstrm "" (Pp.strbrk "Constant does not build instances of a declared type class.")
-
-let declare_class finite def infer id idbuild paramimpls params arity fieldimpls fields
- ?(kind=StructureComponent) ?name is_coe coers priorities sign =
+let declare_class finite def poly ctx id idbuild paramimpls params arity
+ template fieldimpls fields ?(kind=StructureComponent) is_coe coers priorities sign =
let fieldimpls =
- (* Make the class and all params implicits in the projections *)
- let ctx_impls = implicits_of_context params in
- let len = succ (List.length params) in
- List.map (fun x -> ctx_impls @ Impargs.lift_implicits len x) fieldimpls
+ (* Make the class implicit in the projections, and the params if applicable. *)
+ let len = List.length params in
+ let impls = implicits_of_context params in
+ List.map (fun x -> impls @ Impargs.lift_implicits (succ len) x) fieldimpls
in
+ let binder_name = Namegen.next_ident_away (snd id) (Termops.ids_of_context (Global.env())) in
let impl, projs =
match fields with
| [(Name proj_name, _, field)] when def ->
let class_body = it_mkLambda_or_LetIn field params in
- let class_type = Option.map (fun ar -> it_mkProd_or_LetIn ar params) arity in
- let class_entry =
- { const_entry_body = class_body;
- const_entry_secctx = None;
- const_entry_type = class_type;
- const_entry_opaque = false }
- in
+ let _class_type = it_mkProd_or_LetIn arity params in
+ let class_entry =
+ Declare.definition_entry (* ?types:class_type *) ~poly ~univs:ctx class_body in
let cst = Declare.declare_constant (snd id)
(DefinitionEntry class_entry, IsDefinition Definition)
in
- let inst_type = appvectc (mkConst cst) (Termops.rel_vect 0 (List.length params)) in
- let proj_type = it_mkProd_or_LetIn (mkProd(Name (snd id), inst_type, lift 1 field)) params in
- let proj_body = it_mkLambda_or_LetIn (mkLambda (Name (snd id), inst_type, mkRel 1)) params in
- let proj_entry =
- { const_entry_body = proj_body;
- const_entry_secctx = None;
- const_entry_type = Some proj_type;
- const_entry_opaque = false }
- in
+ let cstu = (cst, if poly then Univ.UContext.instance ctx else Univ.Instance.empty) in
+ let inst_type = appvectc (mkConstU cstu) (Termops.rel_vect 0 (List.length params)) in
+ let proj_type =
+ it_mkProd_or_LetIn (mkProd(Name binder_name, inst_type, lift 1 field)) params in
+ let proj_body =
+ it_mkLambda_or_LetIn (mkLambda (Name binder_name, inst_type, mkRel 1)) params in
+ let proj_entry = Declare.definition_entry ~types:proj_type ~poly ~univs:ctx proj_body in
let proj_cst = Declare.declare_constant proj_name
(DefinitionEntry proj_entry, IsDefinition Definition)
in
@@ -336,54 +417,87 @@ let declare_class finite def infer id idbuild paramimpls params arity fieldimpls
Impargs.declare_manual_implicits false cref [paramimpls];
Impargs.declare_manual_implicits false (ConstRef proj_cst) [List.hd fieldimpls];
Classes.set_typeclass_transparency (EvalConstRef cst) false false;
- if infer then Evd.fold (fun ev evi _ -> Recordops.declare_method (ConstRef cst) ev sign) sign ();
- let sub = match List.hd coers with Some b -> Some ((if b then Backward else Forward), List.hd priorities) | None -> None in
+ let sub = match List.hd coers with
+ | Some b -> Some ((if b then Backward else Forward), List.hd priorities)
+ | None -> None
+ in
cref, [Name proj_name, sub, Some proj_cst]
| _ ->
- let idarg = Namegen.next_ident_away (snd id) (Termops.ids_of_context (Global.env())) in
- let ind = declare_structure BiFinite infer (snd id) idbuild paramimpls
- params (Option.default (Termops.new_Type ()) arity) fieldimpls fields
- ~kind:Method ~name:idarg false (List.map (fun _ -> false) fields) sign
+ let ind = declare_structure BiFinite poly ctx (snd id) idbuild paramimpls
+ params arity template fieldimpls fields
+ ~kind:Method ~name:binder_name false (List.map (fun _ -> false) fields) sign
in
let coers = List.map2 (fun coe pri ->
- Option.map (fun b -> if b then Backward, pri else Forward, pri) coe)
- coers priorities
+ Option.map (fun b ->
+ if b then Backward, pri else Forward, pri) coe)
+ coers priorities
in
- IndRef ind, (list_map3 (fun (id, _, _) b y -> (id, b, y))
+ IndRef ind, (List.map3 (fun (id, _, _) b y -> (id, b, y))
(List.rev fields) coers (Recordops.lookup_projections ind))
in
let ctx_context =
List.map (fun (na, b, t) ->
match Typeclasses.class_of_constr t with
- | Some (_, (cl, _)) -> Some (cl.cl_impl, true) (*List.exists (fun (_, n) -> n = na) supnames)*)
+ | Some (_, ((cl,_), _)) -> Some (cl.cl_impl, true)
| None -> None)
params, params
in
let k =
{ cl_impl = impl;
+ cl_strict = !typeclasses_strict;
+ cl_unique = !typeclasses_unique;
cl_context = ctx_context;
cl_props = fields;
cl_projs = projs }
in
-(* list_iter3 (fun p sub pri -> *)
-(* if sub then match p with (_, _, Some p) -> declare_instance_cst true p pri | _ -> ()) *)
-(* k.cl_projs coers priorities; *)
- add_class k; impl
+ add_class k; impl
+
+
+let add_constant_class cst =
+ let ty = Universes.unsafe_type_of_global (ConstRef cst) in
+ let ctx, arity = decompose_prod_assum ty in
+ let tc =
+ { cl_impl = ConstRef cst;
+ cl_context = (List.map (const None) ctx, ctx);
+ cl_props = [(Anonymous, None, arity)];
+ cl_projs = [];
+ cl_strict = !typeclasses_strict;
+ cl_unique = !typeclasses_unique
+ }
+ in add_class tc;
+ set_typeclass_transparency (EvalConstRef cst) false false
+
+let add_inductive_class ind =
+ let mind, oneind = Global.lookup_inductive ind in
+ let k =
+ let ctx = oneind.mind_arity_ctxt in
+ let inst = Univ.UContext.instance mind.mind_universes in
+ let ty = Inductive.type_of_inductive_knowing_parameters
+ (push_rel_context ctx (Global.env ()))
+ ((mind,oneind),inst)
+ (Array.map (fun x -> lazy x) (Termops.extended_rel_vect 0 ctx))
+ in
+ { cl_impl = IndRef ind;
+ cl_context = List.map (const None) ctx, ctx;
+ cl_props = [Anonymous, None, ty];
+ cl_projs = [];
+ cl_strict = !typeclasses_strict;
+ cl_unique = !typeclasses_unique }
+ in add_class k
-let interp_and_check_sort sort =
- Option.map (fun sort ->
- let env = Global.env() and sigma = Evd.empty in
- let s = interp_constr sigma env sort in
- if isSort (Reductionops.whd_betadeltaiota env sigma s) then s
- else user_err_loc (constr_loc sort,"", str"Sort expected.")) sort
+let declare_existing_class g =
+ match g with
+ | ConstRef x -> add_constant_class x
+ | IndRef x -> add_inductive_class x
+ | _ -> user_err_loc (Loc.dummy_loc, "declare_existing_class",
+ Pp.str"Unsupported class type, only constants and inductives are allowed")
open Vernacexpr
-open Autoinstance
(* [fs] corresponds to fields and [ps] to parameters; [coers] is a
list telling if the corresponding fields must me declared as coercions
or subinstances *)
-let definition_structure (kind,finite,infer,(is_coe,(loc,idstruc)),ps,cfs,idbuild,s) =
+let definition_structure (kind,poly,finite,(is_coe,(loc,idstruc)),ps,cfs,idbuild,s) =
let cfs,notations = List.split cfs in
let cfs,priorities = List.split cfs in
let coers,fs = List.split cfs in
@@ -392,26 +506,26 @@ let definition_structure (kind,finite,infer,(is_coe,(loc,idstruc)),ps,cfs,idbuil
| Vernacexpr.DefExpr ((_,Name id),_,_) -> id::acc
| _ -> acc in
let allnames = idstruc::(List.fold_left extract_name [] fs) in
- if not (list_distinct allnames) then error "Two objects have the same name";
- if not (kind = Class false) && List.exists ((<>) None) priorities then
+ if not (List.distinct_f Id.compare allnames)
+ then error "Two objects have the same name";
+ let isnot_class = match kind with Class false -> false | _ -> true in
+ if isnot_class && List.exists (fun opt -> not (Option.is_empty opt)) priorities then
error "Priorities only allowed for type class substructures";
(* Now, younger decl in params and fields is on top *)
- let sc = interp_and_check_sort s in
- let implpars, params, implfs, fields =
+ let ctx, arity, template, implpars, params, implfs, fields =
States.with_state_protection (fun () ->
- typecheck_params_and_fields idstruc sc ps notations fs) () in
+ typecheck_params_and_fields (kind = Class true) idstruc s ps notations fs) () in
let sign = structure_signature (fields@params) in
match kind with
| Class def ->
- let gr = declare_class finite def infer (loc,idstruc) idbuild
- implpars params sc implfs fields is_coe coers priorities sign in
- if infer then search_record declare_class_instance gr sign;
+ 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 arity = Option.default (Termops.new_Type ()) sc in
let implfs = List.map
- (fun impls -> implpars @ Impargs.lift_implicits (succ (List.length params)) impls) implfs in
- let ind = declare_structure finite infer idstruc idbuild implpars params arity implfs
- fields is_coe (List.map (fun coe -> coe <> None) coers) sign in
- if infer then search_record declare_record_instance (ConstructRef (ind,1)) sign;
+ (fun impls -> implpars @ Impargs.lift_implicits
+ (succ (List.length params)) impls) implfs in
+ let ind = declare_structure finite poly ctx idstruc
+ idbuild implpars params arity template implfs
+ fields is_coe (List.map (fun coe -> not (Option.is_empty coe)) coers) sign in
IndRef ind
diff --git a/toplevel/record.mli b/toplevel/record.mli
index 721d0d97..91dccb96 100644
--- a/toplevel/record.mli
+++ b/toplevel/record.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,32 +8,38 @@
open Names
open Term
-open Sign
+open Context
open Vernacexpr
-open Topconstr
+open Constrexpr
open Impargs
-open Libnames
+open Globnames
+
+val primitive_flag : bool ref
(** [declare_projections ref name coers params fields] declare projections of
record [ref] (if allowed) using the given [name] as argument, and put them
as coercions accordingly to [coers]; it returns the absolute names of projections *)
val declare_projections :
- inductive -> ?kind:Decl_kinds.definition_object_kind -> ?name:identifier ->
+ inductive -> ?kind:Decl_kinds.definition_object_kind -> Id.t ->
coercion_flag list -> manual_explicitation list list -> rel_context ->
- (name * bool) list * constant option list
+ (Name.t * bool) list * constant option list
val declare_structure : Decl_kinds.recursivity_kind ->
- bool (**infer?*) -> identifier -> identifier ->
+ bool (** polymorphic?*) -> Univ.universe_context ->
+ Id.t -> Id.t ->
manual_explicitation list -> rel_context -> (** params *) constr -> (** arity *)
+ bool (** template arity ? *) ->
Impargs.manual_explicitation list list -> rel_context -> (** fields *)
- ?kind:Decl_kinds.definition_object_kind -> ?name:identifier ->
+ ?kind:Decl_kinds.definition_object_kind -> ?name:Id.t ->
bool -> (** coercion? *)
bool list -> (** field coercions *)
Evd.evar_map ->
inductive
val definition_structure :
- inductive_kind * Decl_kinds.recursivity_kind * bool(**infer?*)* lident with_coercion * local_binder list *
+ inductive_kind * Decl_kinds.polymorphic * Decl_kinds.recursivity_kind * lident with_coercion * local_binder list *
(local_decl_expr with_instance with_priority with_notation) list *
- identifier * constr_expr option -> global_reference
+ Id.t * constr_expr option -> global_reference
+
+val declare_existing_class : global_reference -> unit
diff --git a/toplevel/search.ml b/toplevel/search.ml
index 28c14a77..59283edf 100644
--- a/toplevel/search.ml
+++ b/toplevel/search.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -9,19 +9,23 @@
open Pp
open Util
open Names
-open Nameops
open Term
-open Glob_term
open Declarations
open Libobject
-open Declare
open Environ
open Pattern
-open Matching
open Printer
open Libnames
+open Globnames
open Nametab
+type filter_function = global_reference -> env -> constr -> bool
+type display_function = global_reference -> env -> constr -> unit
+
+type glob_search_about_item =
+ | GlobSearchSubPattern of constr_pattern
+ | GlobSearchString of string
+
module SearchBlacklist =
Goptions.MakeStringTable
(struct
@@ -32,215 +36,302 @@ module SearchBlacklist =
let synchronous = true
end)
-(* The functions print_constructors and crible implement the behavior needed
- for the Coq searching commands.
+(* The functions iter_constructors and iter_declarations implement the behavior
+ needed for the Coq searching commands.
These functions take as first argument the procedure
that will be called to treat each entry. This procedure receives the name
of the object, the assumptions that will make it possible to print its type,
and the constr term that represent its type. *)
-let print_constructors indsp fn env nconstr =
+let iter_constructors indsp u fn env nconstr =
for i = 1 to nconstr do
- fn (ConstructRef (indsp,i)) env (Inductiveops.type_of_constructor env (indsp,i))
+ let typ = Inductiveops.type_of_constructor env ((indsp, i), u) in
+ fn (ConstructRef (indsp, i)) env typ
done
-let rec head_const c = match kind_of_term c with
- | Prod (_,_,d) -> head_const d
- | LetIn (_,_,_,d) -> head_const d
- | App (f,_) -> head_const f
- | Cast (d,_,_) -> head_const d
- | _ -> c
+let iter_named_context_name_type f = List.iter (fun (nme,_,typ) -> f nme typ)
-(* General search, restricted to head constant if [only_head] *)
+(* General search over hypothesis of a goal *)
+let iter_hypothesis glnum (fn : global_reference -> env -> constr -> unit) =
+ let env = Global.env () in
+ let iter_hyp idh typ = fn (VarRef idh) env typ in
+ let evmap,e = Pfedit.get_goal_context glnum in
+ let pfctxt = named_context e in
+ iter_named_context_name_type iter_hyp pfctxt
-let gen_crible refopt (fn : global_reference -> env -> constr -> unit) =
+(* General search over declarations *)
+let iter_declarations (fn : global_reference -> env -> constr -> unit) =
let env = Global.env () in
- let crible_rec (sp,kn) lobj = match object_tag lobj with
- | "VARIABLE" ->
- (try
- let (id,_,typ) = Global.lookup_named (basename sp) in
- if refopt = None
- || head_const typ = constr_of_global (Option.get refopt)
- then
- fn (VarRef id) env typ
- with Not_found -> (* we are in a section *) ())
- | "CONSTANT" ->
- let cst = Global.constant_of_delta_kn kn in
- let typ = Typeops.type_of_constant env cst in
- if refopt = None
- || head_const typ = constr_of_global (Option.get refopt)
- then
- fn (ConstRef cst) env typ
- | "INDUCTIVE" ->
- let mind = Global.mind_of_delta_kn kn in
- let mib = Global.lookup_mind mind in
- (match refopt with
- | Some (IndRef ((kn',tyi) as ind)) when eq_mind mind kn' ->
- print_constructors ind fn env
- (Array.length (mib.mind_packets.(tyi).mind_user_lc))
- | Some _ -> ()
- | _ ->
- Array.iteri
- (fun i mip -> print_constructors (mind,i) fn env
- (Array.length mip.mind_user_lc)) mib.mind_packets)
- | _ -> ()
+ 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
+ with Not_found -> (* we are in a section *) () end
+ | "CONSTANT" ->
+ let cst = Global.constant_of_delta_kn kn in
+ let gr = ConstRef cst in
+ let typ = Global.type_of_global_unsafe gr in
+ fn gr env typ
+ | "INDUCTIVE" ->
+ let mind = Global.mind_of_delta_kn kn in
+ let mib = Global.lookup_mind mind in
+ let iter_packet i mip =
+ let ind = (mind, i) in
+ let u = Declareops.inductive_instance mib in
+ let i = (ind, u) in
+ let typ = Inductiveops.type_of_inductive env i in
+ let () = fn (IndRef ind) env typ in
+ let len = Array.length mip.mind_user_lc in
+ iter_constructors ind u fn env len
+ in
+ Array.iteri iter_packet mib.mind_packets
+ | _ -> ()
in
- try
- Declaremods.iter_all_segments crible_rec
- with Not_found ->
- ()
-
-let crible ref = gen_crible (Some ref)
-
-(* Fine Search. By Yves Bertot. *)
+ try Declaremods.iter_all_segments iter_obj
+ with Not_found -> ()
-exception No_full_path
+let generic_search glnumopt fn =
+ (match glnumopt with
+ | None -> ()
+ | Some glnum -> iter_hypothesis glnum fn);
+ iter_declarations fn
-let rec head c =
- let c = strip_outer_cast c in
- match kind_of_term c with
- | Prod (_,_,c) -> head c
- | LetIn (_,_,_,c) -> head c
- | _ -> c
+(** Standard display *)
-let xor a b = (a or b) & (not (a & b))
-
-let plain_display ref a c =
- let pc = pr_lconstr_env a c in
+let plain_display accu ref env c =
+ let pc = pr_lconstr_env env Evd.empty c in
let pr = pr_global ref in
- msg (hov 2 (pr ++ str":" ++ spc () ++ pc) ++ fnl ())
-
-let filter_by_module (module_list:dir_path list) (accept:bool)
- (ref:global_reference) _ _ =
- try
- let sp = path_of_global ref in
- let sl = dirpath sp in
- let rec filter_aux = function
- | m :: tl -> (not (is_dirpath_prefix_of m sl)) && (filter_aux tl)
- | [] -> true
- in
- xor accept (filter_aux module_list)
- with No_full_path ->
- false
-
-let ref_eq = Libnames.encode_mind Coqlib.logic_module (id_of_string "eq"), 0
-let c_eq = mkInd ref_eq
-let gref_eq = IndRef ref_eq
-
-let mk_rewrite_pattern1 eq pattern =
- PApp (PRef eq, [| PMeta None; pattern; PMeta None |])
-
-let mk_rewrite_pattern2 eq pattern =
- PApp (PRef eq, [| PMeta None; PMeta None; pattern |])
-
-let pattern_filter pat _ a c =
- try
- try
- is_matching pat (head c)
- with e when Errors.noncritical e ->
- is_matching
- pat (head (Typing.type_of (Global.env()) Evd.empty c))
- with UserError _ ->
- false
-
-let filtered_search filter_function display_function ref =
- crible ref
- (fun s a c -> if filter_function s a c then display_function s a c)
-
-let rec id_from_pattern = function
- | PRef gr -> gr
-(* should be appear as a PRef (VarRef sp) !!
- | PVar id -> Nametab.locate (make_qualid [] (string_of_id id))
- *)
- | PApp (p,_) -> id_from_pattern p
- | _ -> error "The pattern is not simple enough."
-
-let raw_pattern_search extra_filter display_function pat =
- let name = id_from_pattern pat in
- filtered_search
- (fun s a c -> (pattern_filter pat s a c) & extra_filter s a c)
- display_function name
-
-let raw_search_rewrite extra_filter display_function pattern =
- filtered_search
- (fun s a c ->
- ((pattern_filter (mk_rewrite_pattern1 gref_eq pattern) s a c) ||
- (pattern_filter (mk_rewrite_pattern2 gref_eq pattern) s a c))
- && extra_filter s a c)
- display_function gref_eq
-
-let raw_search_by_head extra_filter display_function pattern =
- Util.todo "raw_search_by_head"
-
-let name_of_reference ref = string_of_id (basename_of_global ref)
+ 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. *)
+(** FIXME: this is quite dummy, we may find a more efficient algorithm. *)
+let rec pattern_filter pat ref env typ =
+ let typ = strip_outer_cast typ in
+ if Constr_matching.is_matching env Evd.empty pat typ then true
+ else match kind_of_term typ with
+ | Prod (_, _, typ)
+ | LetIn (_, _, _, typ) -> pattern_filter pat ref env typ
+ | _ -> false
+
+let rec head_filter pat ref env typ =
+ let typ = strip_outer_cast typ in
+ if Constr_matching.is_matching_head env Evd.empty pat typ then true
+ else match kind_of_term typ with
+ | Prod (_, _, typ)
+ | LetIn (_, _, _, typ) -> head_filter pat ref env typ
+ | _ -> false
let full_name_of_reference ref =
let (dir,id) = repr_path (path_of_global ref) in
- string_of_dirpath dir ^ "." ^ string_of_id id
-
-(*
- * functions to use the new Libtypes facility
- *)
+ DirPath.to_string dir ^ "." ^ Id.to_string id
-let raw_search search_function extra_filter display_function pat =
- let env = Global.env() in
- List.iter
- (fun (gr,_,_) ->
- let typ = Global.type_of_global gr in
- if extra_filter gr env typ then
- display_function gr env typ
- ) (search_function pat)
-
-let text_pattern_search extra_filter =
- raw_search Libtypes.search_concl extra_filter plain_display
-
-let text_search_rewrite extra_filter =
- raw_search (Libtypes.search_eq_concl c_eq) extra_filter plain_display
-
-let text_search_by_head extra_filter =
- raw_search Libtypes.search_head_concl extra_filter plain_display
-
-let filter_by_module_from_list = function
- | [], _ -> (fun _ _ _ -> true)
- | l, outside -> filter_by_module l (not outside)
-
-let filter_blacklist gr _ _ =
- let name = full_name_of_reference gr in
+(** Whether a reference is blacklisted *)
+let blacklist_filter ref env typ =
let l = SearchBlacklist.elements () in
- List.for_all (fun str -> not (string_string_contains ~where:name ~what:str)) l
+ let name = full_name_of_reference ref in
+ let is_not_bl str = not (String.string_contains ~where:name ~what:str) in
+ List.for_all is_not_bl l
+
+let module_filter (mods, outside) ref env typ =
+ let sp = path_of_global ref in
+ let sl = dirpath sp in
+ let is_outside md = not (is_dirpath_prefix_of md sl) in
+ let is_inside md = is_dirpath_prefix_of md sl in
+ if outside then List.for_all is_outside mods
+ else List.is_empty mods || List.exists is_inside mods
+
+let name_of_reference ref = Id.to_string (basename_of_global ref)
+
+let search_about_filter query gr env typ = match query with
+| GlobSearchSubPattern pat ->
+ Constr_matching.is_matching_appsubterm ~closed:false env Evd.empty pat typ
+| GlobSearchString s ->
+ String.string_contains ~where:(name_of_reference gr) ~what:s
+
+
+(** SearchPattern *)
+
+let search_pattern gopt pat mods =
+ let ans = ref [] 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
+ in
+ let iter ref env typ =
+ if filter ref env typ then plain_display ans ref env typ
+ in
+ let () = generic_search gopt iter in
+ format_display !ans
-let (&&&&&) f g x y z = f x y z && g x y z
+(** SearchRewrite *)
-let search_by_head pat inout =
- text_search_by_head (filter_by_module_from_list inout &&&&& filter_blacklist) pat
+let eq = Coqlib.glob_eq
-let search_rewrite pat inout =
- text_search_rewrite (filter_by_module_from_list inout &&&&& filter_blacklist) pat
+let rewrite_pat1 pat =
+ PApp (PRef eq, [| PMeta None; pat; PMeta None |])
-let search_pattern pat inout =
- text_pattern_search (filter_by_module_from_list inout &&&&& filter_blacklist) pat
+let rewrite_pat2 pat =
+ PApp (PRef eq, [| PMeta None; PMeta None; pat |])
-let gen_filtered_search filter_function display_function =
- gen_crible None
- (fun s a c -> if filter_function s a c then display_function s a c)
+let search_rewrite gopt pat mods =
+ let pat1 = rewrite_pat1 pat in
+ let pat2 = rewrite_pat2 pat in
+ let ans = ref [] 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
+ in
+ let iter ref env typ =
+ if filter ref env typ then plain_display ans ref env typ
+ in
+ let () = generic_search gopt iter in
+ format_display !ans
+
+(** Search *)
+
+let search_by_head gopt pat mods =
+ let ans = ref [] 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
+ in
+ let iter ref env typ =
+ if filter ref env typ then plain_display ans ref env typ
+ in
+ let () = generic_search gopt iter in
+ format_display !ans
(** SearchAbout *)
-type glob_search_about_item =
- | GlobSearchSubPattern of constr_pattern
- | GlobSearchString of string
-
-let search_about_item (itemref,typ) = function
- | GlobSearchSubPattern pat -> is_matching_appsubterm ~closed:false pat typ
- | GlobSearchString s -> string_string_contains ~where:(name_of_reference itemref) ~what:s
-
-let raw_search_about filter_modules display_function l =
- let filter ref' env typ =
- filter_modules ref' env typ &&
- List.for_all (fun (b,i) -> b = search_about_item (ref',typ) i) l &&
- filter_blacklist ref' () ()
+let search_about gopt items mods =
+ let ans = ref [] 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
in
- gen_filtered_search filter display_function
-
-let search_about ref inout =
- raw_search_about (filter_by_module_from_list inout) plain_display ref
+ let iter ref env typ =
+ if filter ref env typ then plain_display ans ref env typ
+ in
+ let () = generic_search gopt iter in
+ format_display !ans
+
+type search_constraint =
+ | Name_Pattern of string
+ | Type_Pattern of string
+ | SubType_Pattern of string
+ | In_Module of string list
+ | Include_Blacklist
+
+type 'a coq_object = {
+ coq_object_prefix : string list;
+ coq_object_qualid : string list;
+ coq_object_object : 'a;
+}
+
+let interface_search flags =
+ let env = Global.env () in
+ 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.error ("Invalid regexp: " ^ s)
+ in
+ 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
+ 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
+ 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.error ("Module " ^ path ^ " not found.")
+ in
+ extract_flags name tpe subtpe ((id, b) :: mods) blacklist l
+ | (Include_Blacklist, b) :: l ->
+ extract_flags name tpe subtpe mods b l
+ in
+ let (name, tpe, subtpe, mods, blacklist) =
+ extract_flags [] [] [] [] false flags
+ in
+ let filter_function ref env constr =
+ let id = Names.Id.to_string (Nametab.basename_of_global ref) in
+ let path = Libnames.dirpath (Nametab.path_of_global ref) in
+ let toggle x b = if x then b else not b in
+ let match_name (regexp, flag) =
+ toggle (Str.string_match regexp id 0) flag
+ in
+ let match_type (pat, flag) =
+ toggle (Constr_matching.is_matching env Evd.empty pat constr) flag
+ in
+ let match_subtype (pat, flag) =
+ toggle
+ (Constr_matching.is_matching_appsubterm ~closed:false
+ env Evd.empty pat constr) flag
+ in
+ let match_module (mdl, flag) =
+ toggle (Libnames.is_dirpath_prefix_of mdl path) flag
+ in
+ 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
+ in
+ let ans = ref [] in
+ let print_function ref env constr =
+ let fullpath = DirPath.repr (Nametab.dirpath_of_global ref) in
+ let qualid = Nametab.shortest_qualid_of_global Id.Set.empty ref in
+ let (shortpath, basename) = Libnames.repr_qualid qualid in
+ let shortpath = DirPath.repr shortpath in
+ (* [shortpath] is a suffix of [fullpath] and we're looking for the missing
+ prefix *)
+ let rec prefix full short accu = match full, short with
+ | _, [] ->
+ let full = List.rev_map Id.to_string full in
+ (full, accu)
+ | _ :: full, m :: short ->
+ prefix full short (Id.to_string m :: accu)
+ | _ -> assert false
+ in
+ let (prefix, qualid) = prefix fullpath shortpath [Id.to_string basename] in
+ let answer = {
+ coq_object_prefix = prefix;
+ coq_object_qualid = qualid;
+ coq_object_object = string_of_ppcmds (pr_lconstr_env env Evd.empty constr);
+ } in
+ ans := answer :: !ans;
+ in
+ let iter ref env typ =
+ if filter_function ref env typ then print_function ref env typ
+ in
+ let () = generic_search None iter in (* TODO: chose a goal number? *)
+ !ans
diff --git a/toplevel/search.mli b/toplevel/search.mli
index b6cca47d..f69489c3 100644
--- a/toplevel/search.mli
+++ b/toplevel/search.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -11,8 +11,7 @@ open Names
open Term
open Environ
open Pattern
-open Libnames
-open Nametab
+open Globnames
(** {6 Search facilities. } *)
@@ -20,33 +19,55 @@ type glob_search_about_item =
| GlobSearchSubPattern of constr_pattern
| GlobSearchString of string
-val search_by_head : constr -> dir_path list * bool -> unit
-val search_rewrite : constr -> dir_path list * bool -> unit
-val search_pattern : constr -> dir_path list * bool -> unit
-val search_about :
- (bool * glob_search_about_item) list -> dir_path list * bool -> unit
-
-(** The filtering function that is by standard search facilities.
- It can be passed as argument to the raw search functions.
- It is used in pcoq. *)
-
-val filter_by_module_from_list :
- dir_path list * bool -> global_reference -> env -> 'a -> bool
-
-val filter_blacklist : global_reference -> env -> constr -> bool
-
-(** raw search functions can be used for various extensions.
- They are also used for pcoq. *)
-val gen_filtered_search : (global_reference -> env -> constr -> bool) ->
- (global_reference -> env -> constr -> unit) -> unit
-val filtered_search : (global_reference -> env -> constr -> bool) ->
- (global_reference -> env -> constr -> unit) -> global_reference -> unit
-val raw_pattern_search : (global_reference -> env -> constr -> bool) ->
- (global_reference -> env -> constr -> unit) -> constr_pattern -> unit
-val raw_search_rewrite : (global_reference -> env -> constr -> bool) ->
- (global_reference -> env -> constr -> unit) -> constr_pattern -> unit
-val raw_search_about : (global_reference -> env -> constr -> bool) ->
- (global_reference -> env -> constr -> unit) ->
- (bool * glob_search_about_item) list -> unit
-val raw_search_by_head : (global_reference -> env -> constr -> bool) ->
- (global_reference -> env -> constr -> unit) -> constr_pattern -> unit
+type filter_function = global_reference -> env -> constr -> bool
+type display_function = global_reference -> env -> constr -> unit
+
+(** {6 Generic filter functions} *)
+
+val blacklist_filter : filter_function
+(** Check whether a reference is blacklisted. *)
+
+val module_filter : DirPath.t list * bool -> filter_function
+(** Check whether a reference pertains or not to a set of modules *)
+
+val search_about_filter : glob_search_about_item -> filter_function
+(** Check whether a reference matches a SearchAbout query. *)
+
+(** {6 Specialized search functions}
+
+[search_xxx gl pattern modinout] searches the hypothesis of the [gl]th
+goal and the global environment for things matching [pattern] and
+satisfying module exclude/include clauses of [modinout]. *)
+
+val search_by_head : int option -> constr_pattern -> DirPath.t list * bool -> 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_about : int option -> (bool * glob_search_about_item) list
+ -> DirPath.t list * bool -> std_ppcmds
+
+type search_constraint =
+ (** Whether the name satisfies a regexp (uses Ocaml Str syntax) *)
+ | Name_Pattern of string
+ (** Whether the object type satisfies a pattern *)
+ | Type_Pattern of string
+ (** Whether some subtype of object type satisfies a pattern *)
+ | SubType_Pattern of string
+ (** Whether the object pertains to a module *)
+ | In_Module of string list
+ (** Bypass the Search blacklist *)
+ | Include_Blacklist
+
+type 'a coq_object = {
+ coq_object_prefix : string list;
+ coq_object_qualid : string list;
+ coq_object_object : 'a;
+}
+
+val interface_search : (search_constraint * bool) list ->
+ string coq_object list
+
+(** {6 Generic search function} *)
+
+val generic_search : int option -> display_function -> unit
+(** This function iterates over all hypothesis of the goal numbered
+ [glnum] (if present) and all known declarations. *)
diff --git a/toplevel/toplevel.mllib b/toplevel/toplevel.mllib
index e1e349a6..d22524e5 100644
--- a/toplevel/toplevel.mllib
+++ b/toplevel/toplevel.mllib
@@ -1,27 +1,21 @@
Himsg
Cerrors
Class
-Vernacexpr
+Locality
Metasyntax
Auto_ind_decl
-Libtypes
Search
-Autoinstance
-Lemmas
Indschemes
+Obligations
Command
Classes
Record
-Ppvernac
-Backtrack
Vernacinterp
Mltop
Vernacentries
Whelp
Vernac
-Ide_intf
-Ide_slave
-Toplevel
Usage
+Coqloop
Coqinit
Coqtop
diff --git a/toplevel/usage.ml b/toplevel/usage.ml
index 52bd2d33..d4d44569 100644
--- a/toplevel/usage.ml
+++ b/toplevel/usage.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -14,23 +14,25 @@ let version ret =
(* print the usage of coqtop (or coqc) on channel co *)
+let extra_usage = ref []
+let add_to_usage name text = extra_usage := (name,text) :: !extra_usage
+
let print_usage_channel co command =
output_string co command;
output_string co "Coq options are:\n";
output_string co
-" -I dir -as coqdir map physical dir to logical coqdir\
-\n -I dir map directory dir to the empty logical path\
+" -I dir look for ML files in dir\
\n -include dir (idem)\
+\n -I dir -as coqdir implicitly map physical dir to logical coqdir\
\n -R dir -as coqdir recursively map physical dir to logical coqdir\
\n -R dir coqdir (idem)\
+\n -Q dir coqdir map physical dir to logical coqdir\
\n -top coqdir set the toplevel name to be coqdir instead of Top\
\n -notop set the toplevel name to be the empty logical path\
\n -exclude-dir f exclude subdirectories named f for option -R\
\n\
-\n -inputstate f read state from file f.coq\
-\n -is f (idem)\
-\n -nois start with an empty state\
-\n -outputstate f write state in file f.coq\
+\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\
@@ -46,47 +48,54 @@ let print_usage_channel co command =
\n -compile f compile Coq file f.v (implies -batch)\
\n -compile-verbose f verbosely compile Coq file f.v (implies -batch)\
\n\
-\n -opt run the native-code version of Coq\
-\n -byte run the bytecode version of Coq\
-\n\
\n -where print Coq's standard library location and exit\
\n -config print Coq's configuration information and exit\
\n -v print Coq version and exit\
+\n -list-tags print highlight color tags known by Coq and exit\
\n\
\n -q skip loading of rcfile\
\n -init-file f set the rcfile to f\
\n -batch batch mode (exits just after arguments parsing)\
\n -boot boot mode (implies -q and -batch)\
+\n -bt print backtraces (requires configure debug flag)\
+\n -debug debug mode (implies -bt)\
\n -emacs tells Coq it is executed under Emacs\
+\n -color (on|off|auto) configure color output (only active through coqtop)\
\n -noglob do not dump globalizations\
\n -dump-glob f dump globalizations in file f (to be used by coqdoc)\
\n -with-geoproof (yes|no) to (de)activate special functions for Geoproof within Coqide (default is yes)\
\n -impredicative-set set sort Set impredicative\
-\n -force-load-proofs load opaque proofs in memory initially\
-
-\n -lazy-load-proofs load opaque proofs in memory by necessity (default)\
-\n -dont-load-proofs see opaque proofs as axioms instead of loading them\
-\n -xml export XML files either to the hierarchy rooted in\
-\n the directory $COQ_XML_LIBRARY_ROOT (if set) or to\
-\n stdout (if unset)\
-\n -quality improve the legibility of the proof terms produced by\
-\n some tactics\
-\n -h, --help print this list of options\
-\n"
+\n -indices-matter levels of indices (and nonuniform parameters) contribute to the level of inductives\
+\n -time display the time taken by each command\
+\n -no-native-compiler disable the native_compute reduction machinery\
+\n -h, -help print this list of options\
+\n";
+ List.iter (fun (name, text) ->
+ output_string co
+ ("\nWith the flag '-toploop "^name^
+ "' these extra option are also available:\n"^
+ text^"\n"))
+ !extra_usage
(* print the usage on standard error *)
let print_usage = print_usage_channel stderr
let print_usage_coqtop () =
- print_usage "Usage: coqtop <options>\n\n"
+ print_usage "Usage: coqtop <options>\n\n";
+ flush stderr ;
+ exit 1
let print_usage_coqc () =
print_usage "Usage: coqc <options> <Coq options> file...\n\
\noptions are:\
\n -verbose compile verbosely\
\n -image f specify an alternative executable for Coq\
-\n -t keep temporary files\n\n"
+\n -opt run the native-code version of Coq\
+\n -byte run the bytecode version of Coq\
+\n -t keep temporary files\n\n";
+ flush stderr ;
+ exit 1
(* Print the configuration information *)
@@ -101,6 +110,8 @@ let print_config () =
Printf.printf "CAMLBIN=%s/\n" (Envars.camlbin ());
Printf.printf "CAMLLIB=%s/\n" (Envars.camllib ());
Printf.printf "CAMLP4=%s\n" Coq_config.camlp4;
+ Printf.printf "CAMLP4O=%s\n" Coq_config.camlp4o;
Printf.printf "CAMLP4BIN=%s/\n" (Envars.camlp4bin ());
Printf.printf "CAMLP4LIB=%s\n" (Envars.camlp4lib ());
+ Printf.printf "CAMLP4OPTIONS=%s\n" Coq_config.camlp4compat;
Printf.printf "HASNATDYNLINK=%s\n" (if Coq_config.has_natdynlink then "true" else "false")
diff --git a/toplevel/usage.mli b/toplevel/usage.mli
index c3533b95..ed0cd477 100644
--- a/toplevel/usage.mli
+++ b/toplevel/usage.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -13,6 +13,9 @@ val version : int -> 'a
(** {6 Prints the usage on the error output, preceeded by a user-provided message. } *)
val print_usage : string -> unit
+(** {6 Enable toploop plugins to insert some text in the usage message. } *)
+val add_to_usage : string -> string -> unit
+
(** {6 Prints the usage on the error output. } *)
val print_usage_coqtop : unit -> unit
val print_usage_coqc : unit -> unit
diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml
index cb90909e..176a6c33 100644
--- a/toplevel/vernac.ml
+++ b/toplevel/vernac.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -9,22 +9,38 @@
(* Parsing of vernacular. *)
open Pp
-open Lexer
+open Errors
open Util
open Flags
open System
open Vernacexpr
-open Vernacinterp
-open Ppvernac
-open Compat
(* 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). *)
-exception DuringCommandInterp of Util.loc * exn
+(* The navigation vernac commands will be handled separately *)
-exception HasNotFailed
+let rec is_navigation_vernac = function
+ | VernacResetInitial
+ | VernacResetName _
+ | VernacBacktrack _
+ | VernacBackTo _
+ | VernacBack _ -> true
+ | VernacTime l ->
+ List.exists
+ (fun (_,c) -> is_navigation_vernac c) l (* Time Back* is harmless *)
+ | c -> is_deep_navigation_vernac c
+
+and is_deep_navigation_vernac = function
+ | VernacTimeout (_,c) | VernacFail c -> is_navigation_vernac c
+ | _ -> false
+
+(* NB: Reset is now allowed again as asked by A. Chlipala *)
+
+let is_reset = function
+ | VernacResetInitial | VernacResetName _ -> true
+ | _ -> false
(* When doing Load on a file, two behaviors are possible:
@@ -50,85 +66,35 @@ let _ =
Goptions.optread = (fun () -> !atomic_load);
Goptions.optwrite = ((:=) atomic_load) }
-(* Specifies which file is read. The intermediate file names are
- discarded here. The Drop exception becomes an error. We forget
- if the error ocurred during interpretation or not *)
-
-let raise_with_file file exc =
- let (cmdloc,re) =
- match exc with
- | DuringCommandInterp(loc,e) -> (loc,e)
- | e -> (dummy_loc,e)
- in
- let (inner,inex) =
- match re with
- | Error_in_file (_, (b,f,loc), e) when loc <> dummy_loc ->
- ((b, f, loc), e)
- | Loc.Exc_located (loc, e) when loc <> dummy_loc ->
- ((false,file, loc), e)
- | Loc.Exc_located (_, e) | e -> ((false,file,cmdloc), e)
- in
- raise (Error_in_file (file, inner, disable_drop inex))
-
-let real_error = function
- | Loc.Exc_located (_, e) -> e
- | Error_in_file (_, _, e) -> e
- | e -> e
-
-let user_error loc s = Util.user_err_loc (loc,"_",str s)
+(* 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. *)
-(** Timeout handling *)
+type location_files = { outer : string; inner : string }
-(** A global default timeout, controled by option "Set Default Timeout n".
- Use "Unset Default Timeout" to deactivate it (or set it to 0). *)
+let files_of_exn : location_files Exninfo.t = Exninfo.make ()
-let default_timeout = ref None
+let get_exn_files e = Exninfo.get e files_of_exn
-let _ =
- Goptions.declare_int_option
- { Goptions.optsync = true;
- Goptions.optdepr = false;
- Goptions.optname = "the default timeout";
- Goptions.optkey = ["Default";"Timeout"];
- Goptions.optread = (fun () -> !default_timeout);
- Goptions.optwrite = ((:=) default_timeout) }
-
-(** When interpreting a command, the current timeout is initially
- the default one, but may be modified locally by a Timeout command. *)
-
-let current_timeout = ref None
+let add_exn_files e f = Exninfo.add e files_of_exn f
-(** Installing and de-installing a timer.
- Note: according to ocaml documentation, Unix.alarm isn't available
- for native win32. *)
+let raise_with_file f (e, info) =
+ let inner_f = match get_exn_files info with None -> f | Some ff -> ff.inner in
+ iraise (e, add_exn_files info { outer = f; inner = inner_f })
-let timeout_handler _ = raise Timeout
-
-let set_timeout n =
- let psh =
- Sys.signal Sys.sigalrm (Sys.Signal_handle timeout_handler) in
- ignore (Unix.alarm n);
- Some psh
-
-let default_set_timeout () =
- match !current_timeout with
- | Some n -> set_timeout n
- | None -> None
-
-let restore_timeout = function
- | None -> ()
- | Some psh ->
- (* stop alarm *)
- ignore(Unix.alarm 0);
- (* restore handler *)
- Sys.set_signal Sys.sigalrm psh
+let disable_drop = function
+ | Drop -> Errors.error "Drop is forbidden."
+ | e -> e
+let user_error loc s = Errors.user_err_loc (loc,"_",str s)
(* Open an utf-8 encoded file and skip the byte-order mark if any *)
let open_utf8_file_in fname =
let is_bom s =
- Char.code s.[0] = 0xEF && Char.code s.[1] = 0xBB && Char.code s.[2] = 0xBF
+ Int.equal (Char.code s.[0]) 0xEF &&
+ Int.equal (Char.code s.[1]) 0xBB &&
+ Int.equal (Char.code s.[2]) 0xBF
in
let in_chan = open_in fname in
let s = " " in
@@ -141,7 +107,7 @@ let open_utf8_file_in fname =
the file we parse seems a bit risky to me. B.B. *)
let open_file_twice_if verbosely fname =
- let paths = Library.get_load_paths () in
+ let paths = Loadpath.get_paths () in
let _,longfname =
find_file_in_path ~warn:(Flags.is_verbose()) paths fname in
let in_chan = open_utf8_file_in longfname in
@@ -159,23 +125,24 @@ let close_input in_chan (_,verb) =
with e when Errors.noncritical e -> ()
let verbose_phrase verbch loc =
- let loc = unloc loc in
+ let loc = Loc.unloc loc in
match verbch with
| Some ch ->
let len = snd loc - fst loc in
let s = String.create len in
seek_in ch (fst loc);
really_input ch s 0 len;
- message s;
+ ppnl (str s);
pp_flush()
- | _ -> ()
+ | None -> ()
exception End_of_input
-let parse_sentence (po, verbch) =
+let parse_sentence = Flags.with_option Flags.we_are_parsing
+ (fun (po, verbch) ->
match Pcoq.Gram.entry_parse Pcoq.main_entry po with
| Some (loc,_ as com) -> verbose_phrase verbch loc; com
- | None -> raise End_of_input
+ | None -> raise End_of_input)
(* vernac parses the given stream, executes interpfun on the syntax tree it
* parses, and is verbose on "primitives" commands if verbosely is true *)
@@ -191,120 +158,106 @@ let set_formatter_translator() =
Format.set_max_boxes max_int
let pr_new_syntax loc ocom =
- let loc = unloc loc in
+ let loc = Loc.unloc loc in
if !beautify_file then set_formatter_translator();
- let fs = States.freeze () in
+ let fs = States.freeze ~marshallable:`No in
let com = match ocom with
| Some VernacNop -> mt()
- | Some com -> pr_vernac com
+ | Some com -> Ppvernac.pr_vernac com
| None -> mt() in
if !beautify_file then
msg (hov 0 (comment (fst loc) ++ com ++ comment (snd loc)))
else
- msgnl (hov 4 (str"New Syntax:" ++ fnl() ++ (hov 0 com)));
+ msg_info (hov 4 (str"New Syntax:" ++ fnl() ++ (hov 0 com)));
States.unfreeze fs;
Format.set_formatter_out_channel stdout
-let rec vernac_com interpfun checknav (loc,com) =
- let rec interp = function
- | VernacLoad (verbosely, fname) ->
- let fname = expand_path_macros fname in
- (* translator state *)
- let ch = !chan_beautify in
- let cs = Lexer.com_state() in
- let cl = !Pp.comments in
- (* end translator state *)
- (* coqdoc state *)
- let cds = Dumpglob.coqdoc_freeze() in
- if !Flags.beautify_file then
- begin
- let _,f = find_file_in_path ~warn:(Flags.is_verbose())
- (Library.get_load_paths ())
- (make_suffix fname ".v") in
- chan_beautify := open_out (f^beautify_suffix);
- Pp.comments := []
- end;
- begin
- try
- read_vernac_file verbosely (make_suffix fname ".v");
- if !Flags.beautify_file then close_out !chan_beautify;
- chan_beautify := ch;
- Lexer.restore_com_state cs;
- Pp.comments := cl;
- Dumpglob.coqdoc_unfreeze cds
- with reraise ->
- if !Flags.beautify_file then close_out !chan_beautify;
- chan_beautify := ch;
- Lexer.restore_com_state cs;
- Pp.comments := cl;
- Dumpglob.coqdoc_unfreeze cds;
- raise reraise
- end
-
- | VernacList l -> List.iter (fun (_,v) -> interp v) l
-
- | v when !just_parsing -> ()
+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
+
+(* For coqtop -time, we display the position in the file,
+ and a glimpse of the executed command *)
+
+let display_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
+ match s.[i] with
+ | ' ' | '\n' | '\t' | '\r' -> s.[i] <- '~'
+ | _ -> ()
+ done;
+ s
+ in
+ let (start,stop) = Loc.unloc loc in
+ let safe_pr_vernac x =
+ 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 (" ["^cmd^"] "));
+ Pp.flush_all ()
- | VernacFail v ->
- begin try
- (* If the command actually works, ignore its effects on the state *)
- States.with_state_protection
- (fun v -> interp v; raise HasNotFailed) v
- with e when Errors.noncritical e -> match real_error e with
- | HasNotFailed ->
- errorlabstrm "Fail" (str "The command has not failed !")
- | e ->
- (* Anomalies are re-raised by the next line *)
- let msg = Errors.print_no_anomaly e in
- if_verbose msgnl
- (str "The command has indeed failed with message:" ++
- fnl () ++ str "=> " ++ hov 0 msg)
+let rec vernac_com verbosely 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 st = save_translator_coqdoc () in
+ if !Flags.beautify_file then
+ begin
+ let paths = Loadpath.get_paths () in
+ let _,f = find_file_in_path ~warn:(Flags.is_verbose())
+ paths
+ (CUnix.make_suffix fname ".v") in
+ chan_beautify := open_out (f^beautify_suffix);
+ Pp.comments := []
+ end;
+ begin
+ try
+ read_vernac_file verbosely (CUnix.make_suffix fname ".v");
+ restore_translator_coqdoc st;
+ with reraise ->
+ let reraise = Errors.push reraise in
+ restore_translator_coqdoc st;
+ iraise reraise
end
- | VernacTime v ->
- let tstart = System.get_time() in
- interp v;
- let tend = System.get_time() in
- msgnl (str"Finished transaction in " ++
- System.fmt_time_difference tstart tend)
-
- | VernacTimeout(n,v) ->
- current_timeout := Some n;
- interp v
+ | v when !just_parsing -> ()
- | v ->
- let psh = default_set_timeout () in
- try
- States.with_heavy_rollback interpfun
- Cerrors.process_vernac_interp_error v;
- restore_timeout psh
- with reraise -> restore_timeout psh; raise reraise
+ | v -> Stm.interp verbosely (loc,v)
in
try
checknav loc com;
- current_timeout := !default_timeout;
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
interp com
- with any ->
+ with reraise ->
+ let (reraise, info) = Errors.push reraise in
Format.set_formatter_out_channel stdout;
- raise (DuringCommandInterp (loc, any))
+ 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 =
Flags.make_warn verbosely;
- let interpfun =
- if verbosely then Vernacentries.interp
- else Flags.silently Vernacentries.interp
- in
let checknav loc cmd =
if is_navigation_vernac cmd && not (is_reset cmd) then
user_error loc "Navigation commands forbidden in files"
in
- let end_inner_command cmd =
- if !atomic_load || is_reset cmd then
- Lib.mark_end_of_command () (* for Reset in coqc or coqtop -l *)
- else
- Backtrack.mark_command cmd; (* for Show Script, cf bug #2820 *)
- in
let (in_chan, fname, input) =
open_file_twice_if verbosely s in
try
@@ -312,19 +265,20 @@ and read_vernac_file verbosely s =
* 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 interpfun checknav loc_ast;
- end_inner_command (snd loc_ast);
+ vernac_com verbosely checknav loc_ast;
pp_flush ()
done
- with reraise -> (* whatever the exception *)
+ with any -> (* whatever the exception *)
+ let (e, info) = Errors.push any in
Format.set_formatter_out_channel stdout;
close_input in_chan input; (* we must close the file first *)
- match real_error reraise with
+ match e with
| End_of_input ->
- if do_beautify () then pr_new_syntax (make_loc (max_int,max_int)) None
- | _ -> raise_with_file fname reraise
+ if do_beautify () then
+ pr_new_syntax (Loc.make_loc (max_int,max_int)) None
+ | _ -> raise_with_file fname (disable_drop e, info)
-(** [eval_expr : ?preserving:bool -> Pp.loc * Vernacexpr.vernac_expr -> unit]
+(** [eval_expr : ?preserving:bool -> Loc.t * Vernacexpr.vernac_expr -> unit]
It executes one vernacular command. By default the command is
considered as non-state-preserving, in which case we add it to the
Backtrack stack (triggering a save of a frozen state and the generation
@@ -335,43 +289,64 @@ let checknav loc ast =
if is_deep_navigation_vernac ast then
user_error loc "Navigation commands forbidden in nested commands"
-let eval_expr ?(preserving=false) loc_ast =
- vernac_com Vernacentries.interp checknav loc_ast;
- if not preserving && not (is_navigation_vernac (snd loc_ast)) then
- Backtrack.mark_command (snd loc_ast)
-
-(* raw_do_vernac : Pcoq.Gram.parsable -> unit
- * vernac_step . parse_sentence *)
-let raw_do_vernac po = eval_expr (parse_sentence (po,None))
-
-(* XML output hooks *)
-let xml_start_library = ref (fun _ -> ())
-let xml_end_library = ref (fun _ -> ())
-
-let set_xml_start_library f = xml_start_library := f
-let set_xml_end_library f = xml_end_library := f
+let eval_expr loc_ast = vernac_com (Flags.is_verbose()) checknav loc_ast
(* 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
- Lib.mark_end_of_command (); (* in case we're still in coqtop init *)
read_vernac_file verb file;
if !Flags.beautify_file then close_out !chan_beautify;
- with reraise ->
+ with any ->
+ let (e, info) = Errors.push any in
if !Flags.beautify_file then close_out !chan_beautify;
- raise_with_file file reraise
+ raise_with_file file (disable_drop e, info)
(* Compile a vernac file (f is assumed without .v suffix) *)
let compile verbosely f =
- let ldir,long_f_dot_v = Flags.verbosely Library.start_library f in
- Dumpglob.start_dump_glob long_f_dot_v;
- Dumpglob.dump_string ("F" ^ Names.string_of_dirpath ldir ^ "\n");
- if !Flags.xml_export then !xml_start_library ();
- let _ = load_vernac verbosely long_f_dot_v in
- if Pfedit.get_all_proof_names () <> [] then
- (message "Error: There are pending proofs"; exit 1);
- if !Flags.xml_export then !xml_end_library ();
- Dumpglob.end_dump_glob ();
- Library.save_library_to ldir (long_f_dot_v ^ "o")
+ 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
+ match !Flags.compilation_mode with
+ | BuildVo ->
+ let ldir,long_f_dot_v = Flags.verbosely Library.start_library f 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;
+ Dumpglob.dump_string ("F" ^ Names.DirPath.to_string ldir ^ "\n");
+ let wall_clock1 = Unix.gettimeofday () in
+ let _ = load_vernac verbosely long_f_dot_v in
+ Stm.join ();
+ let wall_clock2 = Unix.gettimeofday () in
+ check_pending_proofs ();
+ Library.save_library_to ldir long_f_dot_v (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 ();
+ Dumpglob.end_dump_glob ()
+ | BuildVio ->
+ let ldir, long_f_dot_v = Flags.verbosely Library.start_library f in
+ Dumpglob.noglob ();
+ Stm.set_compilation_hints long_f_dot_v;
+ let _ = load_vernac verbosely long_f_dot_v in
+ Stm.finish ();
+ check_pending_proofs ();
+ Stm.snapshot_vio ldir long_f_dot_v;
+ Stm.reset_task_queue ()
+ | Vio2Vo ->
+ let open Filename in
+ let open Library in
+ Dumpglob.noglob ();
+ let f = if check_suffix f ".vio" then chop_extension f else f in
+ let lfdv, lib, univs, disch, tasks, proofs = load_library_todo f in
+ Stm.set_compilation_hints lfdv;
+ let univs, proofs = Stm.finish_tasks lfdv univs disch proofs tasks in
+ Library.save_library_raw lfdv lib univs proofs
+
+let compile v f =
+ ignore(CoqworkmgrApi.get 1);
+ compile v f;
+ CoqworkmgrApi.giveback 1
+
diff --git a/toplevel/vernac.mli b/toplevel/vernac.mli
index c1ac2154..affc2171 100644
--- a/toplevel/vernac.mli
+++ b/toplevel/vernac.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -12,28 +12,16 @@
Raises [End_of_file] if EOF (or Ctrl-D) is reached. *)
val parse_sentence : Pcoq.Gram.parsable * in_channel option ->
- Util.loc * Vernacexpr.vernac_expr
+ Loc.t * Vernacexpr.vernac_expr
(** Reads and executes vernac commands from a stream.
The boolean [just_parsing] disables interpretation of commands. *)
-exception DuringCommandInterp of Util.loc * exn
exception End_of_input
val just_parsing : bool ref
-(** [eval_expr] executes one vernacular command. By default the command is
- considered as non-state-preserving, in which case we add it to the
- Backtrack stack (triggering a save of a frozen state and the generation
- of a new state label). An example of state-preserving command is one coming
- from the query panel of Coqide. *)
-
-val eval_expr : ?preserving:bool -> Util.loc * Vernacexpr.vernac_expr -> unit
-val raw_do_vernac : Pcoq.Gram.parsable -> unit
-
-(** Set XML hooks *)
-val set_xml_start_library : (unit -> unit) -> unit
-val set_xml_end_library : (unit -> unit) -> unit
+val eval_expr : Loc.t * Vernacexpr.vernac_expr -> unit
(** Load a vernac file, verbosely or not. Errors are annotated with file
and location *)
@@ -44,3 +32,11 @@ val load_vernac : bool -> string -> unit
(** Compile a vernac file, verbosely or not (f is assumed without .v suffix) *)
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 394b58bd..fb12edfb 100644
--- a/toplevel/vernacentries.ml
+++ b/toplevel/vernacentries.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -9,10 +9,10 @@
(* Concrete syntax of the mathematical vernacular MV V2.6 *)
open Pp
+open Errors
open Util
open Flags
open Names
-open Entries
open Nameops
open Term
open Pfedit
@@ -24,39 +24,28 @@ open Tacinterp
open Command
open Goptions
open Libnames
-open Nametab
+open Globnames
open Vernacexpr
open Decl_kinds
-open Topconstr
-open Pretyping
+open Constrexpr
open Redexpr
-open Syntax_def
open Lemmas
-open Declaremods
-
-(* Pcoq hooks *)
-
-type pcoq_hook = {
- start_proof : unit -> unit;
- solve : int -> unit;
- abort : string -> unit;
- search : searchable -> dir_path list * bool -> unit;
- print_name : reference Genarg.or_by_notation -> unit;
- print_check : Environ.env -> Environ.unsafe_judgment -> unit;
- print_eval : Reductionops.reduction_function -> Environ.env -> Evd.evar_map -> constr_expr ->
- Environ.unsafe_judgment -> unit;
- show_goal : goal_reference -> unit
-}
+open Misctypes
+open Locality
-let pcoq = ref None
-let set_pcoq_hook f = pcoq := Some f
+let debug = false
+let prerr_endline =
+ if debug then prerr_endline else fun _ -> ()
(* Misc *)
let cl_of_qualid = function
| FunClass -> Classops.CL_FUN
| SortClass -> Classops.CL_SORT
- | RefClass r -> Class.class_of_global (Smartlocate.smart_global r)
+ | 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)
(*******************)
(* "Show" commands *)
@@ -65,77 +54,31 @@ 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
- msgnl (Util.prlist_with_sep Pp.fnl Printer.pr_constr pprf)
+ msg_notice (Pp.prlist_with_sep Pp.fnl Printer.pr_constr pprf)
let show_node () =
(* spiwack: I'm have little clue what this function used to do. I deactivated it,
could, possibly, be cleaned away. (Feb. 2010) *)
()
-(* indentation code for Show Script, initially contributed
- by D. de Rauglaudre *)
-
-let indent_script_item ((ng1,ngl1),nl,beginend,ppl) (cmd,ng) =
- (* ng1 : number of goals remaining at the current level (before cmd)
- ngl1 : stack of previous levels with their remaining goals
- ng : number of goals after the execution of cmd
- beginend : special indentation stack for { } *)
- let ngprev = List.fold_left (+) ng1 ngl1 in
- let new_ngl =
- if ng > ngprev then
- (* We've branched *)
- (ng - ngprev + 1, ng1 - 1 :: ngl1)
- else if ng < ngprev then
- (* A subgoal have been solved. Let's compute the new current level
- by discarding all levels with 0 remaining goals. *)
- let _ = assert (ng = ngprev - 1) in
- let rec loop = function
- | (0, ng2::ngl2) -> loop (ng2,ngl2)
- | p -> p
- in loop (ng1-1, ngl1)
- else
- (* Standard case, same goal number as before *)
- (ng1, ngl1)
- in
- (* When a subgoal have been solved, separate this block by an empty line *)
- let new_nl = (ng < ngprev)
- in
- (* Indentation depth *)
- let ind = List.length ngl1
- in
- (* Some special handling of bullets and { }, to get a nicer display *)
- let pred n = max 0 (n-1) in
- let ind, nl, new_beginend = match cmd with
- | VernacSubproof _ -> pred ind, nl, (pred ind)::beginend
- | VernacEndSubproof -> List.hd beginend, false, List.tl beginend
- | VernacBullet _ -> pred ind, nl, beginend
- | _ -> ind, nl, beginend
- in
- let pp =
- (if nl then fnl () else mt ()) ++
- (hov (ind+1) (str (String.make ind ' ') ++ Ppvernac.pr_vernac cmd))
- in
- (new_ngl, new_nl, new_beginend, pp :: ppl)
-
-let show_script () =
- let prf = Pfedit.get_current_proof_name () in
- let cmds = Backtrack.get_script prf in
- let _,_,_,indented_cmds =
- List.fold_left indent_script_item ((1,[]),false,[],[]) cmds
- in
- let indented_cmds = List.rev (indented_cmds) in
- msgnl (v 0 (Util.prlist_with_sep Pp.fnl (fun x -> x) indented_cmds))
-
let show_thesis () =
- msgnl (anomaly "TODO" )
+ 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 (pr_evars_int 1 (Evarutil.non_instantiated sigma))
-
+ 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
+ let cstrs = Univ.merge_constraints (Univ.ContextSet.constraints ctx) Univ.empty_universes in
+ msg_notice (Evd.pr_evar_universe_context (Evd.evar_universe_context sigma));
+ msg_notice (str"Normalized constraints: " ++ Univ.pr_universes (Evd.pr_evd_level sigma) cstrs)
let show_prooftree () =
(* Spiwack: proof tree is currently not working *)
@@ -145,7 +88,9 @@ let enable_goal_printing = ref true
let print_subgoals () =
if !enable_goal_printing && is_verbose ()
- then msg (pr_open_subgoals ())
+ then begin
+ msg_notice (pr_open_subgoals ())
+ end
let try_print_subgoals () =
Pp.flush_all();
@@ -156,18 +101,18 @@ let try_print_subgoals () =
let show_intro all =
let pf = get_pftreestate() in
- let {Evd.it=gls ; sigma=sigma} = Proof.V82.subgoals pf in
- let gl = {Evd.it=List.hd gls ; sigma = sigma} in
+ let {Evd.it=gls ; sigma=sigma; } = Proof.V82.subgoals pf in
+ let gl = {Evd.it=List.hd gls ; sigma = sigma; } in
let l,_= decompose_prod_assum (strip_outer_cast (pf_concl gl)) in
if all
then
let lid = Tactics.find_intro_names l gl in
- msgnl (hov 0 (prlist_with_sep spc pr_id lid))
+ msg_notice (hov 0 (prlist_with_sep spc pr_id lid))
else
try
- let n = list_last l in
- msgnl (pr_id (List.hd (Tactics.find_intro_names [n] gl)))
- with Failure "list_last" -> message ""
+ let n = List.last l in
+ msg_notice (pr_id (List.hd (Tactics.find_intro_names [n] gl)))
+ with Failure "List.last" -> ()
(** Prepare a "match" template for a given inductive type.
For each branch of the match, we list the constructor name
@@ -179,21 +124,21 @@ let make_cases s =
let qualified_name = Libnames.qualid_of_string s in
let glob_ref = Nametab.locate qualified_name in
match glob_ref with
- | Libnames.IndRef i ->
+ | Globnames.IndRef i ->
let {Declarations.mind_nparams = np}
, {Declarations.mind_consnames = carr ; Declarations.mind_nf_lc = tarr }
= Global.lookup_inductive i in
- Util.array_fold_right2
+ Util.Array.fold_right2
(fun consname typ l ->
- let al = List.rev (fst (Term.decompose_prod typ)) in
- let al = Util.list_skipn np al in
+ let al = List.rev (fst (decompose_prod typ)) in
+ let al = Util.List.skipn np al in
let rec rename avoid = function
| [] -> []
| (n,_)::l ->
- let n' = Namegen.next_name_away_in_cases_pattern n avoid in
- string_of_id n' :: rename (n'::avoid) l in
+ let n' = Namegen.next_name_away_in_cases_pattern ([],mkMeta 0) n avoid in
+ Id.to_string n' :: rename (n'::avoid) l in
let al' = rename [] al in
- (string_of_id consname :: al') :: l)
+ (Id.to_string consname :: al') :: l)
carr tarr []
| _ -> raise Not_found
@@ -201,36 +146,41 @@ let make_cases s =
let show_match id =
let patterns =
- try make_cases (string_of_id (snd id))
+ try make_cases (Id.to_string (snd 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 (v 1 (str "match # with" ++ fnl () ++
+ 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 (s,l) =
- (str (string_of_dirpath l) ++ str " " ++ tbrk (0,0) ++ str s)
+let print_path_entry p =
+ let dir = str (DirPath.to_string (Loadpath.logical p)) in
+ let path = str (Loadpath.physical p) in
+ (dir ++ str " " ++ tbrk (0, 0) ++ path)
let print_loadpath dir =
- let l = Library.get_full_load_paths () in
+ let l = Loadpath.get_load_paths () in
let l = match dir with
- | None -> l
- | Some dir -> List.filter (fun (s,l) -> is_dirpath_prefix_of dir l) l in
- msgnl (Pp.t (str "Logical Path: " ++
- tab () ++ str "Physical path:" ++ fnl () ++
- prlist_with_sep pr_fnl print_path_entry l))
+ | None -> l
+ | Some dir ->
+ let filter p = is_dirpath_prefix_of dir (Loadpath.logical p) in
+ List.filter filter l
+ in
+ Pp.t (str "Logical Path: " ++
+ tab () ++ str "Physical path:" ++ fnl () ++
+ prlist_with_sep fnl print_path_entry l)
let print_modules () =
let opened = Library.opened_libraries ()
and loaded = Library.loaded_libraries () in
(* we intersect over opened to preserve the order of opened since *)
(* non-commutative operations (e.g. visibility) are done at import time *)
- let loaded_opened = list_intersect opened loaded
- and only_loaded = list_subtract loaded opened in
+ let loaded_opened = List.intersect DirPath.equal opened loaded
+ and only_loaded = List.subtract DirPath.equal loaded opened in
str"Loaded and imported library files: " ++
pr_vertical_list pr_dirpath loaded_opened ++ fnl () ++
str"Loaded and not imported library files: " ++
@@ -243,23 +193,131 @@ let print_module r =
let globdir = Nametab.locate_dir qid in
match globdir with
DirModule (dirpath,(mp,_)) ->
- msgnl (Printmod.print_module (Printmod.printable_body dirpath) mp)
+ msg_notice (Printmod.print_module (Printmod.printable_body dirpath) mp)
| _ -> raise Not_found
with
- Not_found -> msgnl (str"Unknown Module " ++ pr_qualid qid)
+ Not_found -> 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
- msgnl (Printmod.print_modtype kn)
+ 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
- msgnl (Printmod.print_module false mp)
+ msg_notice (Printmod.print_module false mp)
with Not_found ->
- msgnl (str"Unknown Module Type or Module " ++ pr_qualid qid)
+ msg_error (str"Unknown Module Type or Module " ++ pr_qualid qid)
+
+let print_namespace ns =
+ let ns = List.rev (Names.DirPath.repr ns) in
+ (* [match_dirpath], [match_modulpath] are helpers for [matches]
+ which checks whether a constant is in the namespace [ns]. *)
+ let rec match_dirpath ns = function
+ | [] -> Some ns
+ | id::dir ->
+ begin match match_dirpath ns dir with
+ | Some [] as y -> y
+ | Some (a::ns') ->
+ if Names.Id.equal a id then Some ns'
+ else None
+ | None -> None
+ end
+ in
+ let rec match_modulepath ns = function
+ | MPbound _ -> None (* Not a proper namespace. *)
+ | MPfile dir -> match_dirpath ns (Names.DirPath.repr dir)
+ | MPdot (mp,lbl) ->
+ let id = Names.Label.to_id lbl in
+ begin match match_modulepath ns mp with
+ | Some [] as y -> y
+ | Some (a::ns') ->
+ if Names.Id.equal a id then Some ns'
+ else None
+ | None -> None
+ end
+ in
+ (* [qualified_minus n mp] returns a list of qualifiers representing
+ [mp] except the [n] first (in the concrete syntax order). The
+ idea is that if [mp] matches [ns], then [qualified_minus mp
+ (length ns)] will be the correct representation of [mp] assuming
+ [ns] is imported. *)
+ (* precondition: [mp] matches some namespace of length [n] *)
+ let qualified_minus n mp =
+ let rec list_of_modulepath = function
+ | MPbound _ -> assert false (* MPbound never matches *)
+ | MPfile dir -> Names.DirPath.repr dir
+ | MPdot (mp,lbl) -> (Names.Label.to_id lbl)::(list_of_modulepath mp)
+ in
+ snd (Util.List.chop n (List.rev (list_of_modulepath mp)))
+ in
+ let print_list pr l = prlist_with_sep (fun () -> str".") pr l in
+ let print_kn kn =
+ (* spiwack: I'm ignoring the dirpath, is that bad? *)
+ let (mp,_,lbl) = Names.repr_kn kn in
+ let qn = (qualified_minus (List.length ns) mp)@[Names.Label.to_id lbl] in
+ print_list pr_id qn
+ in
+ let print_constant k body =
+ (* FIXME: universes *)
+ let t = Typeops.type_of_constant_type (Global.env ()) body.Declarations.const_type in
+ print_kn k ++ str":" ++ spc() ++ Printer.pr_type t
+ in
+ let matches mp = match match_modulepath ns mp with
+ | Some [] -> true
+ | _ -> false in
+ let constants = (Environ.pre_env (Global.env ())).Pre_env.env_globals.Pre_env.env_constants in
+ let constants_in_namespace =
+ Cmap_env.fold (fun c (body,_) acc ->
+ let kn = user_con c in
+ if matches (modpath kn) then
+ acc++fnl()++hov 2 (print_constant kn body)
+ else
+ acc
+ ) constants (str"")
+ in
+ msg_notice ((print_list pr_id ns)++str":"++fnl()++constants_in_namespace)
+
+let print_strategy r =
+ let open Conv_oracle in
+ let pr_level = function
+ | Expand -> str "expand"
+ | Level 0 -> str "transparent"
+ | Level n -> str "level" ++ spc() ++ int n
+ | Opaque -> str "opaque"
+ in
+ let pr_strategy (ref, lvl) = pr_global ref ++ str " : " ++ pr_level lvl in
+ let oracle = Environ.oracle (Global.env ()) in
+ match r with
+ | None ->
+ let fold key lvl (vacc, cacc) = match key with
+ | VarKey id -> ((VarRef id, lvl) :: vacc, cacc)
+ | ConstKey cst -> (vacc, (ConstRef cst, lvl) :: cacc)
+ | RelKey _ -> (vacc, cacc)
+ in
+ let var_lvl, cst_lvl = fold_strategy fold oracle ([], []) in
+ let var_msg =
+ if List.is_empty var_lvl then mt ()
+ else str "Variable strategies" ++ fnl () ++
+ hov 0 (prlist_with_sep fnl pr_strategy var_lvl) ++ fnl ()
+ in
+ let cst_msg =
+ if List.is_empty cst_lvl then mt ()
+ else str "Constant strategies" ++ fnl () ++
+ hov 0 (prlist_with_sep fnl pr_strategy cst_lvl)
+ in
+ msg_notice (var_msg ++ cst_msg)
+ | Some r ->
+ let r = Smartlocate.smart_global r in
+ let key = match r with
+ | VarRef id -> VarKey id
+ | ConstRef cst -> ConstKey cst
+ | IndRef _ | ConstructRef _ -> error "The reference is not unfoldable"
+ in
+ let lvl = get_strategy oracle key in
+ msg_notice (pr_strategy (r, lvl))
let dump_universes_gen g s =
let output = open_out s in
@@ -293,9 +351,11 @@ let dump_universes_gen g s =
try
Univ.dump_universes output_constraint g;
close ();
- msgnl (str ("Universes written to file \""^s^"\"."))
- with
- reraise -> close (); raise reraise
+ msg_info (str ("Universes written to file \""^s^"\"."))
+ with reraise ->
+ let reraise = Errors.push reraise in
+ close ();
+ iraise reraise
let dump_universes sorted s =
let g = Global.universes () in
@@ -306,111 +366,109 @@ let dump_universes sorted s =
(* "Locate" commands *)
let locate_file f =
- let _,file = System.find_file_in_path ~warn:false (Library.get_load_paths ()) f in
- msgnl (str file)
+ let paths = Loadpath.get_paths () in
+ let _, file = System.find_file_in_path ~warn:false paths f in
+ str file
let msg_found_library = function
| Library.LibLoaded, fulldir, file ->
- msgnl (hov 0
+ msg_info (hov 0
(pr_dirpath fulldir ++ strbrk " has been loaded from file " ++
str file))
| Library.LibInPath, fulldir, file ->
- msgnl (hov 0
+ msg_info (hov 0
(pr_dirpath fulldir ++ strbrk " is bound to file " ++ str file))
-let msg_notfound_library loc qid = function
- | Library.LibUnmappedDir ->
- let dir = fst (repr_qualid qid) in
- user_err_loc (loc,"locate_library",
- strbrk "Cannot find a physical path bound to logical path " ++
- pr_dirpath dir ++ str".")
- | Library.LibNotFound ->
- msgnl (hov 0
- (strbrk "Unable to locate library " ++ pr_qualid qid ++ str"."))
- | e -> assert false
-let print_located_library r =
- let (loc,qid) = qualid_of_reference r in
- try msg_found_library (Library.locate_qualified_library false qid)
- with e when Errors.noncritical e -> msg_notfound_library loc qid e
+let err_unmapped_library loc qid =
+ let dir = fst (repr_qualid qid) in
+ user_err_loc
+ (loc,"locate_library",
+ strbrk "Cannot find a physical path bound to logical path " ++
+ pr_dirpath dir ++ str".")
-let print_located_module r =
- let (loc,qid) = qualid_of_reference r in
- let msg =
- try
- let dir = Nametab.full_name_module qid in
- str "Module " ++ pr_dirpath dir
- with Not_found ->
- (if fst (repr_qualid qid) = empty_dirpath then
- str "No module is referred to by basename "
- else
- str "No module is referred to by name ") ++ pr_qualid qid
- in msgnl msg
+let err_notfound_library loc qid =
+ msg_error
+ (hov 0 (strbrk "Unable to locate library " ++ pr_qualid qid ++ str"."))
-let print_located_tactic r =
+let print_located_library r =
let (loc,qid) = qualid_of_reference r in
- msgnl
- (try
- str "Ltac " ++
- pr_path (Nametab.path_of_tactic (Nametab.locate_tactic qid))
- with Not_found ->
- str "No Ltac definition is referred to by " ++ pr_qualid qid)
+ try msg_found_library (Library.locate_qualified_library false qid)
+ with
+ | Library.LibUnmappedDir -> err_unmapped_library loc qid
+ | Library.LibNotFound -> err_notfound_library loc qid
let smart_global r =
let gr = Smartlocate.smart_global r in
- Dumpglob.add_glob (Genarg.loc_of_or_by_notation loc_of_reference r) gr;
+ Dumpglob.add_glob (Constrarg.loc_of_or_by_notation loc_of_reference r) gr;
gr
let dump_global r =
try
let gr = Smartlocate.smart_global r in
- Dumpglob.add_glob (Genarg.loc_of_or_by_notation loc_of_reference r) gr
+ Dumpglob.add_glob (Constrarg.loc_of_or_by_notation loc_of_reference r) gr
with e when Errors.noncritical e -> ()
(**********)
(* Syntax *)
-let vernac_syntax_extension = Metasyntax.add_syntax_extension
+let vernac_syntax_extension locality local =
+ let local = enforce_module_locality locality local in
+ Metasyntax.add_syntax_extension local
let vernac_delimiters = Metasyntax.add_delimiters
let vernac_bind_scope sc cll =
- List.iter (fun cl -> Metasyntax.add_class_scope sc (cl_of_qualid cl)) cll
+ Metasyntax.add_class_scope sc (List.map scope_class_of_qualid cll)
-let vernac_open_close_scope = Notation.open_close_scope
+let vernac_open_close_scope locality local (b,s) =
+ let local = enforce_section_locality locality local in
+ Notation.open_close_scope (local,b,s)
-let vernac_arguments_scope local r scl =
+let vernac_arguments_scope locality r scl =
+ let local = make_section_locality locality in
Notation.declare_arguments_scope local (smart_global r) scl
-let vernac_infix = Metasyntax.add_infix
+let vernac_infix locality local =
+ let local = enforce_module_locality locality local in
+ Metasyntax.add_infix local
-let vernac_notation = Metasyntax.add_notation
+let vernac_notation locality local =
+ let local = enforce_module_locality locality local in
+ Metasyntax.add_notation local
(***********)
(* Gallina *)
-let start_proof_and_print k l hook =
- check_locality (); (* early check, cf #2975 *)
- start_proof_com k l hook;
- print_subgoals ();
- if !pcoq <> None then (Option.get !pcoq).start_proof ()
+let start_proof_and_print k l hook = start_proof_com k l hook
+
+let no_hook = Lemmas.mk_hook (fun _ _ -> ())
-let vernac_definition (local,k) (loc,id as lid) def hook =
- if local = Local then Dumpglob.dump_definition lid true "var"
- else Dumpglob.dump_definition lid false "def";
+let vernac_definition_hook p = function
+| Coercion -> Class.add_coercion_hook p
+| CanonicalStructure ->
+ Lemmas.mk_hook (fun _ -> Recordops.declare_canonical_structure)
+| SubClass -> Class.add_subclass_hook p
+| _ -> no_hook
+
+let vernac_definition locality p (local,k) (loc,id as lid) def =
+ let local = enforce_locality_exp locality local in
+ let hook = vernac_definition_hook p k in
+ let () = match local with
+ | Discharge -> Dumpglob.dump_definition lid true "var"
+ | Local | Global -> Dumpglob.dump_definition lid false "def"
+ in
(match def with
| ProveBody (bl,t) -> (* local binders, typ *)
- let hook _ _ = () in
- start_proof_and_print (local,DefinitionBody Definition)
- [Some lid, (bl,t,None)] hook
+ start_proof_and_print (local,p,DefinitionBody Definition)
+ [Some lid, (bl,t,None)] no_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
- let ce,imps = interp_definition bl red_option c typ_opt in
- declare_definition id (local,k) ce imps hook)
+ do_definition id (local,p,k) bl red_option c typ_opt hook)
-let vernac_start_proof kind l lettop hook =
+let vernac_start_proof p kind l lettop =
if Dumpglob.dump () then
List.iter (fun (id, _) ->
match id with
@@ -420,23 +478,16 @@ let vernac_start_proof kind l lettop hook =
if lettop then
errorlabstrm "Vernacentries.StartProof"
(str "Let declarations can only be used in proof editing mode.");
- start_proof_and_print (Global, Proof kind) l hook
+ start_proof_and_print (Global, p, Proof kind) l no_hook
let qed_display_script = ref true
-let vernac_end_proof = function
- | Admitted ->
- Backtrack.mark_unreachable [Pfedit.get_current_proof_name ()];
- admit ()
- | Proved (is_opaque,idopt) ->
- let prf = Pfedit.get_current_proof_name () in
- if is_verbose () && !qed_display_script then (show_script (); msg (fnl()));
- begin match idopt with
- | None -> save_named is_opaque
- | Some ((_,id),None) -> save_anonymous is_opaque id
- | Some ((_,id),Some kind) -> save_anonymous_with_strength kind is_opaque id
- end;
- Backtrack.mark_unreachable [prf]
+let vernac_end_proof ?proof = function
+ | Admitted -> save_proof ?proof Admitted
+ | Proved (_,_) as e ->
+ if is_verbose () && !qed_display_script && !Flags.coqtop_ui then
+ Stm.show_script ?proof ();
+ save_proof ?proof e
(* A stupid macro that should be replaced by ``Exact c. Save.'' all along
the theories [??] *)
@@ -444,22 +495,23 @@ let vernac_end_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 prf = Pfedit.get_current_proof_name () in
- by (Tactics.exact_proof c);
- save_named true;
- Backtrack.mark_unreachable [prf]
-
-let vernac_assumption kind l nl=
- let global = fst kind = Global in
- List.iter (fun (is_coe,(idl,c)) ->
- if Dumpglob.dump () then
- List.iter (fun lid ->
- if global then Dumpglob.dump_definition lid false "ax"
- else Dumpglob.dump_definition lid true "var") idl;
- let t,imps = interp_assumption [] c in
- declare_assumptions idl is_coe kind t imps false nl) l
-
-let vernac_record k finite infer struc binders sort nameopt cfs =
+ let status = by (Tactics.New.exact_proof c) in
+ save_proof (Vernacexpr.Proved(true,None));
+ if not status then Pp.feedback Feedback.AddedAxiom
+
+let vernac_assumption locality poly (local, kind) l nl =
+ let local = enforce_locality_exp locality local in
+ let global = local == Global in
+ let kind = local, poly, kind in
+ List.iter (fun (is_coe,(idl,c)) ->
+ if Dumpglob.dump () then
+ List.iter (fun lid ->
+ if global then Dumpglob.dump_definition lid false "ax"
+ else Dumpglob.dump_definition lid true "var") idl) l;
+ let status = do_assumptions kind nl l in
+ if not status then Pp.feedback Feedback.AddedAxiom
+
+let vernac_record k poly finite struc binders sort nameopt cfs =
let const = match nameopt with
| None -> add_prefix "Build_" (snd (snd struc))
| Some (_,id as lid) ->
@@ -470,9 +522,9 @@ let vernac_record k finite infer struc binders sort nameopt cfs =
match x with
| Vernacexpr.AssumExpr ((loc, Name id), _) -> Dumpglob.dump_definition (loc,id) false "proj"
| _ -> ()) cfs);
- ignore(Record.definition_structure (k,finite,infer,struc,binders,cfs,const,sort))
+ ignore(Record.definition_structure (k,poly,finite,struc,binders,cfs,const,sort))
-let vernac_inductive finite infer indl =
+let vernac_inductive poly lo finite indl =
if Dumpglob.dump () then
List.iter (fun (((coe,lid), _, _, _, cstrs), _) ->
match cstrs with
@@ -483,37 +535,45 @@ let vernac_inductive finite infer indl =
| _ -> () (* dumping is done by vernac_record (called below) *) )
indl;
match indl with
+ | [ ( _ , _ , _ ,Record, Constructors _ ),_ ] ->
+ Errors.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."
| [ ( id , bl , c , b, RecordDecl (oc,fs) ), [] ] ->
vernac_record (match b with Class true -> Class false | _ -> b)
- finite infer id bl c oc fs
+ poly finite id bl c oc fs
| [ ( id , bl , c , Class true, 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) finite infer id bl c None [f]
+ in vernac_record (Class true) poly finite id bl c None [f]
| [ ( id , bl , c , Class true, _), _ ] ->
- Util.error "Definitional classes must have a single method"
+ Errors.error "Definitional classes must have a single method"
| [ ( id , bl , c , Class false, Constructors _), _ ] ->
- Util.error "Inductive classes not supported"
+ Errors.error "Inductive classes not supported"
| [ ( _ , _ , _ , _, RecordDecl _ ) , _ ] ->
- Util.error "where clause not supported for (co)inductive records"
+ Errors.error "where clause not supported for (co)inductive records"
| _ -> let unpack = function
- | ( (_, id) , bl , c , _ , Constructors l ) , ntn -> ( id , bl , c , l ) , ntn
- | _ -> Util.error "Cannot handle mutually (co)inductive records."
+ | ( (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."
in
let indl = List.map unpack indl in
- do_mutual_inductive indl (recursivity_flag_of_kind finite)
+ do_mutual_inductive indl poly lo finite
-let vernac_fixpoint l =
+let vernac_fixpoint locality poly local l =
+ let local = enforce_locality_exp locality local in
if Dumpglob.dump () then
List.iter (fun ((lid, _, _, _, _), _) -> Dumpglob.dump_definition lid false "def") l;
- do_fixpoint l
+ do_fixpoint local poly l
-let vernac_cofixpoint l =
+let vernac_cofixpoint locality poly local l =
+ let local = enforce_locality_exp locality local in
if Dumpglob.dump () then
List.iter (fun ((lid, _, _, _), _) -> Dumpglob.dump_definition lid false "def") l;
- do_cofixpoint l
+ do_cofixpoint local poly l
let vernac_scheme l =
if Dumpglob.dump () then
@@ -528,9 +588,12 @@ let vernac_scheme l =
let vernac_combined_scheme lid l =
if Dumpglob.dump () then
(Dumpglob.dump_definition lid false "def";
- List.iter (fun lid -> dump_global (Genarg.AN (Ident lid))) l);
+ List.iter (fun lid -> dump_global (Misctypes.AN (Ident lid))) l);
Indschemes.do_combined_scheme lid l
+let vernac_universe l = do_universe l
+let vernac_constraint l = do_constraint l
+
(**********************)
(* Modules *)
@@ -548,19 +611,18 @@ let vernac_declare_module export (loc, id) binders_ast mty_ast =
error "Modules and Module Types are not allowed inside sections.";
let binders_ast = List.map
(fun (export,idl,ty) ->
- if export <> None then
+ if not (Option.is_empty export) then
error ("Arguments of a functor declaration cannot be exported. " ^
"Remove the \"Export\" and \"Import\" keywords from every functor " ^
"argument.")
else (idl,ty)) binders_ast in
- let mp = Declaremods.declare_module
- Modintern.interp_modtype Modintern.interp_modexpr
- Modintern.interp_modexpr_or_modtype
- id binders_ast (Enforce mty_ast) []
+ let mp =
+ Declaremods.declare_module Modintern.interp_module_ast
+ id binders_ast (Enforce mty_ast) []
in
- Dumpglob.dump_moddef loc mp "mod";
- if_verbose message ("Module "^ string_of_id id ^" is declared");
- Option.iter (fun export -> vernac_import export [Ident (dummy_loc,id)]) export
+ Dumpglob.dump_moddef loc mp "mod";
+ if_verbose msg_info (str ("Module "^ Id.to_string id ^" is declared"));
+ Option.iter (fun export -> vernac_import export [Ident (Loc.ghost,id)]) export
let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l =
(* We check the state of the system (in section, in module type)
@@ -575,40 +637,40 @@ let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l =
(fun (export,idl,ty) (args,argsexport) ->
(idl,ty)::args, (List.map (fun (_,i) -> export,i)idl)@argsexport) binders_ast
([],[]) in
- let mp = Declaremods.start_module Modintern.interp_modtype export
- id binders_ast mty_ast_o
+ let mp =
+ Declaremods.start_module Modintern.interp_module_ast
+ export id binders_ast mty_ast_o
in
- Dumpglob.dump_moddef loc mp "mod";
- if_verbose message
- ("Interactive Module "^ string_of_id id ^" started") ;
- List.iter
- (fun (export,id) ->
- Option.iter
- (fun export -> vernac_import export [Ident (dummy_loc,id)]) export
- ) argsexport
+ Dumpglob.dump_moddef loc mp "mod";
+ if_verbose msg_info
+ (str ("Interactive Module "^ Id.to_string id ^" started"));
+ List.iter
+ (fun (export,id) ->
+ Option.iter
+ (fun export -> vernac_import export [Ident (Loc.ghost,id)]) export
+ ) argsexport
| _::_ ->
let binders_ast = List.map
(fun (export,idl,ty) ->
- if export <> None then
+ if not (Option.is_empty export) then
error ("Arguments of a functor definition can be imported only if" ^
" the definition is interactive. Remove the \"Export\" and " ^
"\"Import\" keywords from every functor argument.")
else (idl,ty)) binders_ast in
- let mp = Declaremods.declare_module
- Modintern.interp_modtype Modintern.interp_modexpr
- Modintern.interp_modexpr_or_modtype
- id binders_ast mty_ast_o mexpr_ast_l
+ let mp =
+ Declaremods.declare_module Modintern.interp_module_ast
+ id binders_ast mty_ast_o mexpr_ast_l
in
- Dumpglob.dump_moddef loc mp "mod";
- if_verbose message
- ("Module "^ string_of_id id ^" is defined");
- Option.iter (fun export -> vernac_import export [Ident (dummy_loc,id)])
- export
+ Dumpglob.dump_moddef loc mp "mod";
+ if_verbose msg_info
+ (str ("Module "^ Id.to_string id ^" is defined"));
+ Option.iter (fun export -> vernac_import export [Ident (Loc.ghost,id)])
+ export
let vernac_end_module export (loc,id as lid) =
let mp = Declaremods.end_module () in
Dumpglob.dump_modref loc mp "mod";
- if_verbose message ("Module "^ string_of_id id ^" is defined") ;
+ if_verbose msg_info (str ("Module "^ Id.to_string id ^" 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 =
@@ -618,45 +680,48 @@ let vernac_declare_module_type (loc,id) binders_ast mty_sign mty_ast_l =
match mty_ast_l with
| [] ->
check_no_pending_proofs ();
- let binders_ast,argsexport =
+ let binders_ast,argsexport =
List.fold_right
(fun (export,idl,ty) (args,argsexport) ->
(idl,ty)::args, (List.map (fun (_,i) -> export,i)idl)@argsexport) binders_ast
([],[]) in
-
- let mp = Declaremods.start_modtype
- Modintern.interp_modtype id binders_ast mty_sign in
- Dumpglob.dump_moddef loc mp "modtype";
- if_verbose message
- ("Interactive Module Type "^ string_of_id id ^" started");
- List.iter
+
+ let mp =
+ Declaremods.start_modtype Modintern.interp_module_ast
+ id binders_ast mty_sign
+ in
+ Dumpglob.dump_moddef loc mp "modtype";
+ if_verbose msg_info
+ (str ("Interactive Module Type "^ Id.to_string id ^" started"));
+ List.iter
(fun (export,id) ->
Option.iter
- (fun export -> vernac_import export [Ident (dummy_loc,id)]) export
+ (fun export -> vernac_import export [Ident (Loc.ghost,id)]) export
) argsexport
| _ :: _ ->
let binders_ast = List.map
(fun (export,idl,ty) ->
- if export <> None then
+ if not (Option.is_empty export) then
error ("Arguments of a functor definition can be imported only if" ^
" the definition is interactive. Remove the \"Export\" " ^
"and \"Import\" keywords from every functor argument.")
else (idl,ty)) binders_ast in
- let mp = Declaremods.declare_modtype Modintern.interp_modtype
- Modintern.interp_modexpr_or_modtype
- id binders_ast mty_sign mty_ast_l in
- Dumpglob.dump_moddef loc mp "modtype";
- if_verbose message
- ("Module Type "^ string_of_id id ^" is defined")
+ let mp =
+ Declaremods.declare_modtype Modintern.interp_module_ast
+ id binders_ast mty_sign mty_ast_l
+ in
+ Dumpglob.dump_moddef loc mp "modtype";
+ if_verbose msg_info
+ (str ("Module Type "^ Id.to_string id ^" is defined"))
let vernac_end_modtype (loc,id) =
let mp = Declaremods.end_modtype () in
Dumpglob.dump_modref loc mp "modtype";
- if_verbose message ("Module Type "^ string_of_id id ^" is defined")
+ if_verbose msg_info (str ("Module Type "^ Id.to_string id ^" is defined"))
let vernac_include l =
- Declaremods.declare_include Modintern.interp_modexpr_or_modtype l
+ Declaremods.declare_include Modintern.interp_module_ast l
(**********************)
(* Gallina extensions *)
@@ -670,9 +735,11 @@ let vernac_begin_section (_, id as lid) =
let vernac_end_section (loc,_) =
Dumpglob.dump_reference loc
- (string_of_dirpath (Lib.current_dirpath true)) "<>" "sec";
+ (DirPath.to_string (Lib.current_dirpath true)) "<>" "sec";
Lib.close_section ()
+let vernac_name_sec_hyp (_,id) set = Proof_using.name_set id set
+
(* Dispatcher of the "End" command *)
let vernac_end_segment (_,id as lid) =
@@ -685,9 +752,9 @@ let vernac_end_segment (_,id as lid) =
(* Libraries *)
-let vernac_require import _ qidl =
+let vernac_require import qidl =
let qidl = List.map qualid_of_reference qidl in
- let modrefl = Flags.silently (List.map Library.try_locate_qualified_library) qidl in
+ let modrefl = List.map Library.try_locate_qualified_library qidl in
if Dumpglob.dump () then
List.iter2 (fun (loc, _) dp -> Dumpglob.dump_libref loc dp "lib") qidl (List.map fst modrefl);
Library.require_library_from_dirpath modrefl import
@@ -697,32 +764,36 @@ let vernac_require import _ qidl =
let vernac_canonical r =
Recordops.declare_canonical_structure (smart_global r)
-let vernac_coercion stre ref qids qidt =
+let vernac_coercion locality poly local ref qids qidt =
+ let local = enforce_locality locality local in
let target = cl_of_qualid qidt in
let source = cl_of_qualid qids in
let ref' = smart_global ref in
- Class.try_add_new_coercion_with_target ref' stre ~source ~target;
- if_verbose msgnl (pr_global ref' ++ str " is now a coercion")
+ Class.try_add_new_coercion_with_target ref' ~local poly ~source ~target;
+ if_verbose msg_info (pr_global ref' ++ str " is now a coercion")
-let vernac_identity_coercion stre id qids qidt =
+let vernac_identity_coercion locality poly local id qids qidt =
+ let local = enforce_locality locality local in
let target = cl_of_qualid qidt in
let source = cl_of_qualid qids in
- Class.try_add_new_identity_coercion id stre ~source ~target
+ Class.try_add_new_identity_coercion id ~local poly ~source ~target
(* Type classes *)
-let vernac_instance abst glob sup inst props pri =
+let vernac_instance abst locality poly sup inst props pri =
+ let global = not (make_section_locality locality) in
Dumpglob.dump_constraint inst false "inst";
- ignore(Classes.new_instance ~abstract:abst ~global:glob sup inst props pri)
+ ignore(Classes.new_instance ~abstract:abst ~global poly sup inst props pri)
-let vernac_context l =
- Classes.context l
+let vernac_context poly l =
+ if not (Classes.context poly l) then Pp.feedback Feedback.AddedAxiom
-let vernac_declare_instances glob ids =
- List.iter (fun (id) -> Classes.existing_instance glob id) ids
+let vernac_declare_instances locality ids pri =
+ let glob = not (make_section_locality locality) in
+ List.iter (fun id -> Classes.existing_instance glob id pri) ids
let vernac_declare_class id =
- Classes.declare_class id
+ Record.declare_existing_class (Nametab.global id)
(***********)
(* Solving *)
@@ -731,18 +802,32 @@ let command_focus = Proof.new_focus_kind ()
let focus_command_cond = Proof.no_cond command_focus
-let vernac_solve n tcom b =
+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 p = Proof_global.give_me_the_proof () in
- Proof.transaction p begin fun () ->
- solve_nth n (Tacinterp.hide_interp tcom None) ~with_end_tac:b;
+ 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 *)
- Proof_global.maximal_unfocus command_focus p;
- print_subgoals();
- if !pcoq <> None then (Option.get !pcoq).solve n
- end
+ 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
@@ -756,54 +841,69 @@ let vernac_solve_existential = instantiate_nth_evar_com
let vernac_set_end_tac tac =
if not (refining ()) then
error "Unknown command of the non proof-editing mode.";
- if tac <> (Tacexpr.TacId []) then set_end_tac (Tacinterp.interp tac) else ()
+ match tac with
+ | Tacexpr.TacId [] -> ()
+ | _ -> set_end_tac tac
(* TO DO verifier s'il faut pas mettre exist s | TacId s ici*)
-let vernac_set_used_variables l =
- let l = List.map snd l in
- if not (list_distinct l) then error "Used variables list contains duplicates";
- let vars = Environ.named_context (Global.env ()) in
+let vernac_set_used_variables e =
+ let env = Global.env () in
+ let tys =
+ List.map snd (Proof.initial_goals (Proof_global.give_me_the_proof ())) in
+ let l = Proof_using.process_expr env e tys in
+ let vars = Environ.named_context env in
List.iter (fun id ->
- if not (List.exists (fun (id',_,_) -> id = id') vars) then
- error ("Unknown variable: " ^ string_of_id id))
+ if not (List.exists (fun (id',_,_) -> Id.equal id id') vars) then
+ error ("Unknown variable: " ^ Id.to_string id))
l;
- set_used_variables l
+ let closure_l = List.map pi1 (set_used_variables l) in
+ let closure_l = List.fold_right Id.Set.add closure_l Id.Set.empty in
+ let vars_of = Environ.global_vars_set in
+ let aux env entry (all_safe,rest as orig) =
+ match entry with
+ | (x,None,_) ->
+ if Id.Set.mem x all_safe then orig else (all_safe, (Loc.ghost,x)::rest)
+ | (x,Some bo, ty) ->
+ let vars = Id.Set.union (vars_of env bo) (vars_of env ty) in
+ if Id.Set.subset vars all_safe then (Id.Set.add x all_safe, rest)
+ else (all_safe, (Loc.ghost,x) :: rest) in
+ let _,to_clear = Environ.fold_named_context aux env ~init:(closure_l,[]) in
+ vernac_solve
+ SelectAll None Tacexpr.(TacAtom (Loc.ghost,TacClear(false,to_clear))) false
+
(*****************************)
(* Auxiliary file management *)
-let vernac_require_from export spec filename =
- Library.require_library_from_file None
- (System.expand_path_macros filename)
- export
+let expand filename =
+ Envars.expand_path_macros ~warn:(fun x -> msg_warning (str x)) filename
let vernac_add_loadpath isrec pdir ldiropt =
- let pdir = System.expand_path_macros pdir in
- let alias = match ldiropt with
- | None -> Nameops.default_root_prefix
- | Some ldir -> ldir in
- (if isrec then Mltop.add_rec_path else Mltop.add_path) ~unix_path:pdir ~coq_root:alias
+ let pdir = expand pdir in
+ let alias = Option.default Nameops.default_root_prefix ldiropt in
+ (if isrec then Mltop.add_rec_path else Mltop.add_path)
+ ~unix_path:pdir ~coq_root:alias ~implicit:true
let vernac_remove_loadpath path =
- Library.remove_load_path (System.expand_path_macros path)
+ Loadpath.remove_load_path (expand path)
(* Coq syntax for ML or system commands *)
let vernac_add_ml_path isrec path =
- (if isrec then Mltop.add_rec_ml_dir else Mltop.add_ml_dir)
- (System.expand_path_macros path)
+ (if isrec then Mltop.add_rec_ml_dir else Mltop.add_ml_dir) (expand path)
-let vernac_declare_ml_module local l =
- Mltop.declare_ml_modules local (List.map System.expand_path_macros l)
+let vernac_declare_ml_module locality l =
+ let local = make_locality locality in
+ Mltop.declare_ml_modules local (List.map expand l)
let vernac_chdir = function
- | None -> message (Sys.getcwd())
+ | None -> msg_notice (str (Sys.getcwd()))
| Some path ->
begin
- try Sys.chdir (System.expand_path_macros path)
- with Sys_error str -> msg_warn ("Cd failed: " ^ str)
+ try Sys.chdir (expand path)
+ with Sys_error err -> msg_warning (str ("Cd failed: " ^ err))
end;
- if_verbose message (Sys.getcwd())
+ if_verbose msg_info (str (Sys.getcwd()))
(********************)
@@ -820,42 +920,118 @@ let vernac_restore_state file =
(************)
(* Commands *)
-let vernac_declare_tactic_definition (local,x,def) =
- Tacinterp.add_tacdef local x def
+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 ltacrecvars =
+ let fold accu (op, _) = match op with
+ | UpdateTac _ -> accu
+ | NewTac id -> Id.Map.add id (Lib.make_kn id) accu
+ in
+ if isrec then List.fold_left fold Id.Map.empty rfun
+ else Id.Map.empty
+ in
+ let ist = { (Tacintern.make_empty_glob_sign ()) with Genintern.ltacrecvars; } 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 = List.map map rfun 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_create_hintdb local id b =
- Auto.create_hint_db local id full_transparent_state b
+let vernac_declare_tactic_definition locality (x,def) =
+ let local = make_module_locality locality in
+ register_ltac local x def
-let vernac_remove_hints local dbs ids =
- Auto.remove_hints local dbs (List.map Smartlocate.global_with_alias ids)
+let vernac_create_hintdb locality id b =
+ let local = make_module_locality locality in
+ Hints.create_hint_db local id full_transparent_state b
-let vernac_hints local lb h =
- Auto.add_hints local lb (Auto.interp_hints h)
+let vernac_remove_hints locality dbs ids =
+ let local = make_module_locality locality in
+ Hints.remove_hints local dbs (List.map Smartlocate.global_with_alias ids)
-let vernac_syntactic_definition lid =
+let vernac_hints locality poly local lb h =
+ let local = enforce_module_locality locality local in
+ Hints.add_hints local lb (Hints.interp_hints poly h)
+
+let vernac_syntactic_definition locality lid x local y =
Dumpglob.dump_definition lid false "syndef";
- Metasyntax.add_syntactic_definition (snd lid)
+ let local = enforce_module_locality locality local in
+ Metasyntax.add_syntactic_definition (snd lid) x local y
-let vernac_declare_implicits local r = function
+let vernac_declare_implicits locality r l =
+ let local = make_section_locality locality in
+ match l with
| [] ->
Impargs.declare_implicits local (smart_global r)
| _::_ as imps ->
Impargs.declare_manual_implicits local (smart_global r) ~enriching:false
(List.map (List.map (fun (ex,b,f) -> ex, (b,true,f))) imps)
-let vernac_declare_arguments local r l nargs flags =
+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 ((<>) names) rest then
+ 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 (Util.list_distinct (List.filter ((<>) Anonymous) names)) then
- error "Arguments names must be distinct.";
+ 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 inf_names =
- Impargs.compute_implicits_names (Global.env()) (Global.type_of_global sr) in
- let string_of_name = function Anonymous -> "_" | Name id -> string_of_id id in
+ let ty = Global.type_of_global_unsafe sr in
+ Impargs.compute_implicits_names (Global.env ()) ty in
+ let string_of_name = function Anonymous -> "_" | Name id -> Id.to_string id in
let rec check li ld ls = match li, ld, ls with
| [], [], [] -> ()
| [], Anonymous::ld, (Some _)::ls when extra_scope_flag -> check li ld ls
@@ -866,18 +1042,23 @@ let vernac_declare_arguments local r l nargs flags =
(String.concat ", " (List.map string_of_name l)) ^ ".")
| _::li, _::ld, _::ls -> check li ld ls
| _ -> assert false in
- if l <> [[]] then
- List.iter2 (fun l -> check inf_names l) (names :: rest) scopes;
+ let () = match l with
+ | [[]] -> ()
+ | _ ->
+ List.iter2 (fun l -> check inf_names l) (names :: rest) scopes
+ 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
+ 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
+ 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 = if l = [[]] then l else
+ 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
@@ -885,81 +1066,103 @@ let vernac_declare_arguments local r l nargs flags =
let names_decl = List.map (List.map (fun (id, _,_,_,_) -> id)) l in
let renamed_arg = ref None in
let set_renamed a b =
- if !renamed_arg = None && a <> b then renamed_arg := Some(b,a) in
+ if Option.is_empty !renamed_arg && not (Id.equal a b) then renamed_arg := Some(b,a) in
let pr_renamed_arg () = match !renamed_arg with None -> ""
| Some (o,n) ->
"\nArgument "^string_of_id o ^" renamed to "^string_of_id n^"." in
let some_renaming_specified =
- try Arguments_renaming.arguments_names sr <> names_decl
+ 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 =
- if l = [[]] then false, [[]] else
- Util.list_fold_map (fun sr il ->
- let sr', impl = Util.list_fold_map (fun b -> function
+ match l with
+ | [[]] -> false, [[]]
+ | _ ->
+ 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 ->
- error ("Argument "^string_of_id x^" cannot be declared implicit.")
+ error ("Argument "^Id.to_string x^" cannot be declared implicit.")
| (Name iid, _,_, true, max), Name id ->
set_renamed iid id;
- b || iid <> id, Some (ExplByName id, max, false)
+ b || not (Id.equal iid id), Some (ExplByName id, max, false)
| (Name iid, _,_, _, _), Name id ->
set_renamed iid id;
- b || iid <> id, None
+ b || not (Id.equal iid id), None
| _ -> b, None)
sr (List.combine il inf_names) in
- sr || sr', Util.list_map_filter (fun x -> x) impl)
+ 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
error ("To rename arguments the \"rename\" flag must be specified."
^ pr_renamed_arg ())
- else Arguments_renaming.rename_arguments local sr names_decl;
+ 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 = implicits <> [[]] in
+ let some_implicits_specified = match implicits with
+ | [[]] -> false | _ -> true in
let scopes = List.map (function
| None -> None
- | Some (o, k) ->
- try Some(ignore(Notation.find_scope k); k)
- with e when Errors.noncritical e ->
+ | Some (o, k) ->
+ try ignore (Notation.find_scope k); Some k
+ with UserError _ ->
Some (Notation.find_delimiters_scope o k)) scopes
in
- let some_scopes_specified = List.exists ((<>) None) scopes in
+ let some_scopes_specified = List.exists ((!=) None) scopes in
let rargs =
- Util.list_map_filter (function (n, true) -> Some n | _ -> None)
- (Util.list_map_i (fun i (_, b, _,_,_) -> i, b) 0 l) in
+ 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 local r scopes;
+ vernac_arguments_scope locality r scopes;
if not some_implicits_specified && List.mem `DefaultImplicits flags then
- vernac_declare_implicits local r []
+ vernac_declare_implicits locality r []
else if some_implicits_specified || List.mem `ClearImplicits flags then
- vernac_declare_implicits local r implicits;
+ 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 rec narrow = function
- | #Tacred.simpl_flag as x :: tl -> x :: narrow tl
+ | #Reductionops.ReductionBehaviour.flag as x :: tl -> x :: narrow tl
| [] -> [] | _ :: tl -> narrow tl in
let flags = narrow flags in
- if rargs <> [] || nargs >= 0 || flags <> [] then
+ 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
match sr with
| ConstRef _ as c ->
- Tacred.set_simpl_behaviour local c (rargs, nargs, flags)
+ 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.")
+ end;
+ if not (some_renaming_specified ||
+ some_implicits_specified ||
+ some_scopes_specified ||
+ some_simpl_flags_specified) &&
+ List.length flags = 0 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'")
+
+
+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 =
let sb_decl = (fun (idl,c) ->
- let t = Constrintern.interp_type Evd.empty (Global.env()) c in
- let t = Detyping.detype false [] [] t in
- let t = aconstr_of_glob_constr [] [] t in
+ let env = Global.env() in
+ let t,ctx = Constrintern.interp_type env Evd.empty c in
+ let t = Detyping.detype false [] env Evd.empty t in
+ let t = Notation_ops.notation_constr_of_glob_constr (default_env ()) t in
Reserve.declare_reserved_type idl t)
in List.iter sb_decl bl
-let vernac_generalizable = Implicit_quantifiers.declare_generalizable
-
-let make_silent_if_not_pcoq b =
- if !pcoq <> None then
- error "Turning on/off silent flag is not supported in Pcoq mode."
- else make_silent b
+let vernac_generalizable locality =
+ let local = make_non_locality locality in
+ Implicit_quantifiers.declare_generalizable local
let _ =
declare_bool_option
@@ -968,7 +1171,7 @@ let _ =
optname = "silent";
optkey = ["Silent"];
optread = is_silent;
- optwrite = make_silent_if_not_pcoq }
+ optwrite = make_silent }
let _ =
declare_bool_option
@@ -1048,8 +1251,8 @@ let _ =
optdepr = false;
optname = "printing of existential variable instances";
optkey = ["Printing";"Existential";"Instances"];
- optread = (fun () -> !Constrextern.print_evar_arguments);
- optwrite = (:=) Constrextern.print_evar_arguments }
+ optread = (fun () -> !Detyping.print_evar_arguments);
+ optwrite = (:=) Detyping.print_evar_arguments }
let _ =
declare_bool_option
@@ -1109,6 +1312,24 @@ let _ =
declare_bool_option
{ optsync = true;
optdepr = false;
+ optname = "use of the program extension";
+ optkey = ["Program";"Mode"];
+ optread = (fun () -> !Flags.program_mode);
+ optwrite = (fun b -> Flags.program_mode:=b) }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "universe polymorphism";
+ optkey = ["Universe"; "Polymorphism"];
+ optread = Flags.is_universe_polymorphism;
+ optwrite = Flags.make_universe_polymorphism }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
optname = "use of virtual machine inside the kernel";
optkey = ["Virtual";"Machine"];
optread = (fun () -> Vconv.use_vm ());
@@ -1129,6 +1350,15 @@ let _ =
declare_bool_option
{ optsync = true;
optdepr = false;
+ optname = "kernel term sharing";
+ optkey = ["Kernel"; "Term"; "Sharing"];
+ optread = (fun () -> !Closure.share);
+ optwrite = (fun b -> Closure.share := b) }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
optname = "use of boxed values";
optkey = ["Boxed";"Values"];
optread = (fun _ -> not (Vm.transp_values ()));
@@ -1191,7 +1421,7 @@ let _ =
optdepr = false;
optname = "Ltac debug";
optkey = ["Ltac";"Debug"];
- optread = (fun () -> get_debug () <> Tactic_debug.DebugOff);
+ optread = (fun () -> get_debug () != Tactic_debug.DebugOff);
optwrite = vernac_debug }
let _ =
@@ -1203,15 +1433,27 @@ let _ =
optread = (fun () -> !Constrintern.parsing_explicit);
optwrite = (fun b -> Constrintern.parsing_explicit := b) }
-let vernac_set_opacity local str =
+let vernac_set_strategy locality l =
+ let local = make_locality locality in
+ let glob_ref r =
+ match smart_global r with
+ | ConstRef sp -> EvalConstRef sp
+ | VarRef id -> EvalVarRef id
+ | _ -> error
+ "cannot set an inductive type or a constructor as transparent" in
+ let l = List.map (fun (lev,ql) -> (lev,List.map glob_ref ql)) l in
+ Redexpr.set_strategy local l
+
+let vernac_set_opacity locality (v,l) =
+ let local = make_non_locality locality in
let glob_ref r =
match smart_global r with
| ConstRef sp -> EvalConstRef sp
| VarRef id -> EvalVarRef id
| _ -> error
"cannot set an inductive type or a constructor as transparent" in
- let str = List.map (fun (lev,ql) -> (lev,List.map glob_ref ql)) str in
- Redexpr.set_strategy local str
+ let l = List.map glob_ref l in
+ Redexpr.set_strategy local [v,l]
let vernac_set_option locality key = function
| StringValue s -> set_string_option_value_gen locality key s
@@ -1255,91 +1497,138 @@ let get_current_context_of_args = function
| None -> get_current_context ()
let vernac_check_may_eval redexp glopt rc =
- let module P = Pretype_errors in
let (sigma, env) = get_current_context_of_args glopt in
- let sigma', c = interp_open_constr sigma env rc in
+ let sigma', c = interp_open_constr env sigma rc in
let sigma' = Evarconv.consider_remaining_unif_problems env sigma' in
+ Evarconv.check_problems_are_solved env sigma';
+ let sigma',nf = Evarutil.nf_evars_and_universes sigma' in
+ let uctx = Evd.universe_context sigma' in
+ let env = Environ.push_context uctx env in
+ let c = nf c in
let j =
- try
- Evarutil.check_evars env sigma sigma' c;
- Arguments_renaming.rename_typing env c
- with P.PretypeError (_,_,P.UnsolvableImplicit _)
- | Compat.Loc.Exc_located (_,P.PretypeError (_,_,P.UnsolvableImplicit _)) ->
- Evarutil.j_nf_evar sigma' (Retyping.get_judgment_of env sigma' c) in
+ if Evarutil.has_undefined_evars sigma' c then
+ Evarutil.j_nf_evar sigma' (Retyping.get_judgment_of env sigma' c)
+ else
+ (* OK to call kernel which does not support evars *)
+ Arguments_renaming.rename_typing env c in
match redexp with
| None ->
- if !pcoq <> None then (Option.get !pcoq).print_check env j
- else msg (print_judgment env j)
+ 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 ++
+ (if l != Evar.Set.empty then
+ let l = Evar.Set.fold (fun ev -> Evar.Map.add ev (Evarutil.nf_evar_info sigma' (Evd.find sigma' ev))) l Evar.Map.empty in
+ (fnl () ++ str "where" ++ fnl () ++ pr_evars sigma' l)
+ else
+ mt ()) ++
+ Printer.pr_universe_ctx uctx)
| Some r ->
- Tacinterp.dump_glob_red_expr r;
+ Tacintern.dump_glob_red_expr r;
let (sigma',r_interp) = interp_redexp env sigma' r in
- let redfun = fst (reduction_of_red_expr r_interp) in
- if !pcoq <> None
- then (Option.get !pcoq).print_eval redfun env sigma' rc j
- else msg (print_eval redfun env sigma' rc j)
+ 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 vernac_declare_reduction locality s r =
- declare_red_expr locality s (snd (interp_redexp (Global.env()) Evd.empty r))
+ let local = make_locality locality in
+ declare_red_expr local s (snd (interp_redexp (Global.env()) Evd.empty r))
(* The same but avoiding the current goal context if any *)
let vernac_global_check c =
- let evmap = Evd.empty in
let env = Global.env() in
- let c = interp_constr evmap env c in
+ let sigma = Evd.from_env env in
+ let c,ctx = interp_constr env sigma c in
let senv = Global.safe_env() in
+ let cstrs = snd (Evd.evar_universe_context_set ctx) in
+ let senv = Safe_typing.add_constraints cstrs senv in
let j = Safe_typing.typing senv c in
- msg (print_safe_judgment env j)
+ let env = Safe_typing.env_of_safe_env senv in
+ msg_notice (print_safe_judgment env sigma j)
+
+let get_nth_goal n =
+ let pf = get_pftreestate() in
+ let {Evd.it=gls ; sigma=sigma; } = Proof.V82.subgoals pf in
+ let gl = {Evd.it=List.nth gls (n-1) ; sigma = sigma; } in
+ gl
+
+exception NoHyp
+(* Printing "About" information of a hypothesis of the current goal.
+ We only print the type and a small statement to this comes from the
+ goal. Precondition: there must be at least one current goal. *)
+let print_about_hyp_globs ref_or_by_not glnumopt =
+ try
+ let gl,id =
+ match glnumopt,ref_or_by_not with
+ | None,AN (Ident (_loc,id)) -> (* goal number not given, catch any failure *)
+ (try get_nth_goal 1,id with _ -> raise NoHyp)
+ | Some n,AN (Ident (_loc,id)) -> (* goal number given, catch if wong *)
+ (try get_nth_goal n,id
+ with
+ Failure _ -> Errors.error ("No such goal: "^string_of_int n^"."))
+ | _ , _ -> 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()
+ ++ str natureofid ++ str " of the goal context.")
+ with (* fallback to globals *)
+ | NoHyp | Not_found -> print_about ref_or_by_not
+
+
let vernac_print = function
- | PrintTables -> print_tables ()
- | PrintFullContext-> msg (print_full_context_typ ())
- | PrintSectionContext qid -> msg (print_sec_context_typ qid)
- | PrintInspect n -> msg (inspect n)
- | PrintGrammar ent -> Metasyntax.print_grammar ent
- | PrintLoadPath dir -> (* For compatibility ? *) print_loadpath dir
- | PrintModules -> msg (print_modules ())
+ | PrintTables -> msg_notice (print_tables ())
+ | PrintFullContext-> msg_notice (print_full_context_typ ())
+ | PrintSectionContext qid -> msg_notice (print_sec_context_typ qid)
+ | PrintInspect n -> msg_notice (inspect n)
+ | PrintGrammar ent -> msg_notice (Metasyntax.pr_grammar ent)
+ | PrintLoadPath dir -> (* For compatibility ? *) msg_notice (print_loadpath dir)
+ | PrintModules -> msg_notice (print_modules ())
| PrintModule qid -> print_module qid
| PrintModuleType qid -> print_modtype qid
- | PrintMLLoadPath -> Mltop.print_ml_path ()
- | PrintMLModules -> Mltop.print_ml_modules ()
- | PrintName qid ->
- if !pcoq <> None then (Option.get !pcoq).print_name qid
- else msg (print_name qid)
- | PrintGraph -> ppnl (Prettyp.print_graph())
- | PrintClasses -> ppnl (Prettyp.print_classes())
- | PrintTypeClasses -> ppnl (Prettyp.print_typeclasses())
- | PrintInstances c -> ppnl (Prettyp.print_instances (smart_global c))
- | PrintLtac qid -> ppnl (Tacinterp.print_ltac (snd (qualid_of_reference qid)))
- | PrintCoercions -> ppnl (Prettyp.print_coercions())
+ | PrintNamespace ns -> print_namespace ns
+ | PrintMLLoadPath -> msg_notice (Mltop.print_ml_path ())
+ | PrintMLModules -> msg_notice (Mltop.print_ml_modules ())
+ | PrintDebugGC -> msg_notice (Mltop.print_gc ())
+ | PrintName qid -> dump_global qid; msg_notice (print_name qid)
+ | PrintGraph -> msg_notice (Prettyp.print_graph())
+ | PrintClasses -> msg_notice (Prettyp.print_classes())
+ | PrintTypeClasses -> msg_notice (Prettyp.print_typeclasses())
+ | PrintInstances c -> msg_notice (Prettyp.print_instances (smart_global c))
+ | PrintLtac qid -> msg_notice (Tacintern.print_ltac (snd (qualid_of_reference qid)))
+ | PrintCoercions -> msg_notice (Prettyp.print_coercions())
| PrintCoercionPaths (cls,clt) ->
- ppnl (Prettyp.print_path_between (cl_of_qualid cls) (cl_of_qualid clt))
- | PrintCanonicalConversions -> ppnl (Prettyp.print_canonical_projections ())
+ msg_notice (Prettyp.print_path_between (cl_of_qualid cls) (cl_of_qualid clt))
+ | PrintCanonicalConversions -> msg_notice (Prettyp.print_canonical_projections ())
| PrintUniverses (b, None) ->
let univ = Global.universes () in
let univ = if b then Univ.sort_universes univ else univ in
- pp (Univ.pr_universes univ)
+ msg_notice (Univ.pr_universes Universes.pr_with_global_universes univ)
| PrintUniverses (b, Some s) -> dump_universes b s
- | PrintHint r -> Auto.print_hint_ref (smart_global r)
- | PrintHintGoal -> Auto.print_applicable_hint ()
- | PrintHintDbName s -> Auto.print_hint_db_by_name s
- | PrintRewriteHintDbName s -> Autorewrite.print_rewrite_hintdb s
- | PrintHintDb -> Auto.print_searchtable ()
+ | 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 ->
- pp (Notation.pr_scopes (Constrextern.without_symbols pr_lglob_constr))
+ msg_notice (Notation.pr_scopes (Constrextern.without_symbols pr_lglob_constr))
| PrintScope s ->
- pp (Notation.pr_scope (Constrextern.without_symbols pr_lglob_constr) s)
+ msg_notice (Notation.pr_scope (Constrextern.without_symbols pr_lglob_constr) s)
| PrintVisibility s ->
- pp (Notation.pr_visibility (Constrextern.without_symbols pr_lglob_constr) s)
- | PrintAbout qid ->
- msg (print_about qid)
- | PrintImplicit qid ->
- dump_global qid; msg (print_impargs qid)
- | PrintAssumptions (o,r) ->
+ msg_notice (Notation.pr_visibility (Constrextern.without_symbols pr_lglob_constr) s)
+ | PrintAbout (ref_or_by_not,glnumopt) ->
+ msg_notice (print_about_hyp_globs ref_or_by_not glnumopt)
+ | PrintImplicit qid ->
+ dump_global qid; msg_notice (print_impargs qid)
+ | PrintAssumptions (o,t,r) ->
(* Prints all the axioms and section variables used by a term *)
- let cstr = constr_of_global (smart_global r) in
- let st = Conv_oracle.get_transp_state () in
- let nassums = Assumptions.assumptions st ~add_opaque:o cstr in
- msg (Printer.pr_assumptionset (Global.env ()) nassums)
+ let cstr = printable_constr_of_global (smart_global r) in
+ let st = Conv_oracle.get_transp_state (Environ.oracle (Global.env())) in
+ let nassums =
+ Assumptions.assumptions st ~add_opaque:o ~add_transparent:t cstr in
+ msg_notice (Printer.pr_assumptionset (Global.env ()) nassums)
+ | PrintStrategy r -> print_strategy r
let global_module r =
let (loc,qid) = qualid_of_reference r in
@@ -1354,177 +1643,89 @@ let interp_search_restriction = function
open Search
-let is_ident s = try ignore (check_ident s); true with UserError _ -> false
-
-let interp_search_about_item = function
+let interp_search_about_item env =
+ function
| SearchSubPattern pat ->
- let _,pat = intern_constr_pattern Evd.empty (Global.env()) pat in
+ let _,pat = intern_constr_pattern env pat in
GlobSearchSubPattern pat
- | SearchString (s,None) when is_ident s ->
+ | SearchString (s,None) when Id.is_valid s ->
GlobSearchString s
| SearchString (s,sc) ->
try
let ref =
- Notation.interp_notation_as_global_reference dummy_loc
+ Notation.interp_notation_as_global_reference Loc.ghost
(fun _ -> true) s sc in
GlobSearchSubPattern (Pattern.PRef ref)
with UserError _ ->
error ("Unable to interp \""^s^"\" either as a reference or \
as an identifier component")
-let vernac_search s r =
+let vernac_search s gopt r =
let r = interp_search_restriction r in
- if !pcoq <> None then (Option.get !pcoq).search s r else
+ let env,gopt =
+ match gopt with | None ->
+ (* 1st goal by default if it exists, otherwise no goal at all *)
+ (try snd (Pfedit.get_goal_context 1) , Some 1
+ with _ -> Global.env (),None)
+ (* if goal selector is given and wrong, then let exceptions be raised. *)
+ | Some g -> snd (Pfedit.get_goal_context g) , Some g
+ in
+ let get_pattern c = snd (intern_constr_pattern env c) in
match s with
| SearchPattern c ->
- let (_,c) = interp_open_constr_patvar Evd.empty (Global.env()) c in
- Search.search_pattern c r
+ msg_notice (Search.search_pattern gopt (get_pattern c) r)
| SearchRewrite c ->
- let _,pat = interp_open_constr_patvar Evd.empty (Global.env()) c in
- Search.search_rewrite pat r
+ msg_notice (Search.search_rewrite gopt (get_pattern c) r)
| SearchHead c ->
- let _,pat = interp_open_constr_patvar Evd.empty (Global.env()) c in
- Search.search_by_head pat r
+ msg_notice (Search.search_by_head gopt (get_pattern c) r)
| SearchAbout sl ->
- Search.search_about (List.map (on_snd interp_search_about_item) sl) r
+ msg_notice (Search.search_about gopt (List.map (on_snd (interp_search_about_item env)) sl) r)
let vernac_locate = function
- | LocateTerm (Genarg.AN qid) -> msgnl (print_located_qualid qid)
- | LocateTerm (Genarg.ByNotation (_,ntn,sc)) ->
- ppnl
+ | LocateAny (AN qid) -> msg_notice (print_located_qualid qid)
+ | LocateTerm (AN qid) -> msg_notice (print_located_term qid)
+ | LocateAny (ByNotation (_, ntn, sc)) (** TODO : handle Ltac notations *)
+ | LocateTerm (ByNotation (_, ntn, sc)) ->
+ msg_notice
(Notation.locate_notation
(Constrextern.without_symbols pr_lglob_constr) ntn sc)
| LocateLibrary qid -> print_located_library qid
- | LocateModule qid -> print_located_module qid
- | LocateTactic qid -> print_located_tactic qid
- | LocateFile f -> locate_file f
-
-(****************)
-(* Backtracking *)
-
-(** NB: these commands are now forbidden in non-interactive use,
- e.g. inside VernacLoad, VernacList, ... *)
-
-let vernac_backto lbl =
- try
- let lbl' = Backtrack.backto lbl in
- if lbl <> lbl' then
- Pp.msg_warning
- (str "Actually back to state "++ Pp.int lbl' ++ str ".");
- try_print_subgoals ()
- with Backtrack.Invalid -> error "Invalid backtrack."
-
-let vernac_back n =
- try
- let extra = Backtrack.back n in
- if extra <> 0 then
- Pp.msg_warning
- (str "Actually back by " ++ Pp.int (extra+n) ++ str " steps.");
- try_print_subgoals ()
- with Backtrack.Invalid -> error "Invalid backtrack."
-
-let vernac_reset_name id =
- try
- let globalized =
- try
- let gr = Smartlocate.global_with_alias (Ident id) in
- Dumpglob.add_glob (fst id) gr;
- true
- with e when Errors.noncritical e -> false in
-
- if not globalized then begin
- try begin match Lib.find_opening_node (snd id) with
- | Lib.OpenedSection _ -> Dumpglob.dump_reference (fst id)
- (string_of_dirpath (Lib.current_dirpath true)) "<>" "sec";
- (* Might be nice to implement module cases, too.... *)
- | _ -> ()
- end with UserError _ -> ()
- end;
-
- if Backtrack.is_active () then
- (Backtrack.reset_name id; try_print_subgoals ())
- else
- (* When compiling files, Reset is now allowed again
- as asked by A. Chlipala. we emulate a simple reset,
- that discards all proofs. *)
- let lbl = Lib.label_before_name id in
- Pfedit.delete_all_proofs ();
- Pp.msg_warning (str "Reset command occurred in non-interactive mode.");
- Lib.reset_label lbl
- with Backtrack.Invalid | Not_found -> error "Invalid Reset."
-
-
-let vernac_reset_initial () =
- if Backtrack.is_active () then
- Backtrack.reset_initial ()
- else begin
- Pp.msg_warning (str "Reset command occurred in non-interactive mode.");
- Lib.reset_label Lib.first_command_label
- end
-
-(* For compatibility with ProofGeneral: *)
-
-let vernac_backtrack snum pnum naborts =
- Backtrack.backtrack snum pnum naborts;
- try_print_subgoals ()
-
+ | LocateModule qid -> msg_notice (print_located_module qid)
+ | LocateTactic qid -> msg_notice (print_located_tactic qid)
+ | LocateFile f -> msg_notice (locate_file f)
+
+let vernac_register id r =
+ if Pfedit.refining () then
+ error "Cannot register a primitive while in proof editing mode.";
+ let t = (Constrintern.global_reference (snd id)) in
+ if not (isConst t) then
+ error "Register inline: a constant is expected";
+ let kn = destConst t in
+ match r with
+ | RegisterInline -> Global.register_inline (Univ.out_punivs kn)
(********************)
(* Proof management *)
-let vernac_abort = function
- | None ->
- Backtrack.mark_unreachable [Pfedit.get_current_proof_name ()];
- delete_current_proof ();
- if_verbose message "Current goal aborted";
- if !pcoq <> None then (Option.get !pcoq).abort ""
- | Some id ->
- Backtrack.mark_unreachable [snd id];
- delete_proof id;
- let s = string_of_id (snd id) in
- if_verbose message ("Goal "^s^" aborted");
- if !pcoq <> None then (Option.get !pcoq).abort s
-
-let vernac_abort_all () =
- if refining() then begin
- Backtrack.mark_unreachable (Pfedit.get_all_proof_names ());
- delete_all_proofs ();
- message "Current goals aborted"
- end else
- error "No proof-editing in progress."
-
-let vernac_restart () =
- Backtrack.mark_unreachable [Pfedit.get_current_proof_name ()];
- restart_proof(); print_subgoals ()
-
-let vernac_undo n =
- let d = Pfedit.current_proof_depth () - n in
- Backtrack.mark_unreachable ~after:d [Pfedit.get_current_proof_name ()];
- Pfedit.undo n; print_subgoals ()
-
-let vernac_undoto n =
- Backtrack.mark_unreachable ~after:n [Pfedit.get_current_proof_name ()];
- Pfedit.undo_todepth n;
- print_subgoals ()
-
let vernac_focus gln =
- let p = Proof_global.give_me_the_proof () in
- let n = match gln with None -> 1 | Some n -> n in
- if n = 0 then
- Util.error "Invalid goal number: 0. Goal numbering starts with 1."
- else
- Proof.focus focus_command_cond () n p; print_subgoals ()
+ Proof_global.simple_with_current_proof (fun _ p ->
+ match gln with
+ | None -> Proof.focus focus_command_cond () 1 p
+ | Some 0 ->
+ Errors.error "Invalid goal number: 0. Goal numbering starts with 1."
+ | Some n ->
+ Proof.focus focus_command_cond () n p)
(* Unfocuses one step in the focus stack. *)
let vernac_unfocus () =
- let p = Proof_global.give_me_the_proof () in
- Proof.unfocus command_focus p; print_subgoals ()
+ Proof_global.simple_with_current_proof
+ (fun _ p -> Proof.unfocus command_focus p ())
(* Checks that a proof is fully unfocused. Raises an error if not. *)
let vernac_unfocused () =
let p = Proof_global.give_me_the_proof () in
if Proof.unfocused p then
- msg (str"The proof is indeed fully unfocused.")
+ msg_notice (str"The proof is indeed fully unfocused.")
else
error "The proof is not fully unfocused."
@@ -1539,44 +1740,39 @@ let subproof_kind = Proof.new_focus_kind ()
let subproof_cond = Proof.done_cond subproof_kind
let vernac_subproof gln =
- let p = Proof_global.give_me_the_proof () in
- begin match gln with
- | None -> Proof.focus subproof_cond () 1 p
- | Some n -> Proof.focus subproof_cond () n p
- end ;
- print_subgoals ()
+ Proof_global.simple_with_current_proof (fun _ p ->
+ match gln with
+ | None -> Proof.focus subproof_cond () 1 p
+ | Some n -> Proof.focus subproof_cond () n p)
let vernac_end_subproof () =
- let p = Proof_global.give_me_the_proof () in
- Proof.unfocus subproof_kind p ; print_subgoals ()
-
+ Proof_global.simple_with_current_proof (fun _ p ->
+ Proof.unfocus subproof_kind p ())
let vernac_bullet (bullet:Proof_global.Bullet.t) =
- let p = Proof_global.give_me_the_proof () in
- Proof.transaction p
- (fun () -> Proof_global.Bullet.put p bullet);
- (* Makes the focus visible in emacs by re-printing the goal. *)
- if !Flags.print_emacs then print_subgoals ()
-
+ Proof_global.simple_with_current_proof (fun _ p ->
+ Proof_global.Bullet.put p bullet)
let vernac_show = function
| ShowGoal goalref ->
- if !pcoq <> None then (Option.get !pcoq).show_goal goalref
- else msg (match goalref with
- | OpenSubgoals -> pr_open_subgoals ()
- | NthGoal n -> pr_nth_open_subgoal n
- | GoalId id -> pr_goal_by_id id)
+ let info = match goalref with
+ | OpenSubgoals -> pr_open_subgoals ()
+ | NthGoal n -> pr_nth_open_subgoal n
+ | GoalId id -> pr_goal_by_id id
+ in
+ msg_notice info
| ShowGoalImplicitly None ->
- Constrextern.with_implicits msg (pr_open_subgoals ())
+ Constrextern.with_implicits msg_notice (pr_open_subgoals ())
| ShowGoalImplicitly (Some n) ->
- Constrextern.with_implicits msg (pr_nth_open_subgoal n)
+ Constrextern.with_implicits msg_notice (pr_nth_open_subgoal n)
| ShowProof -> show_proof ()
| ShowNode -> show_node ()
- | ShowScript -> show_script ()
+ | ShowScript -> Stm.show_script ()
| ShowExistentials -> show_top_evars ()
+ | ShowUniverses -> show_universes ()
| ShowTree -> show_prooftree ()
| ShowProofNames ->
- msgnl (prlist_with_sep pr_spc pr_id (Pfedit.get_all_proof_names()))
+ msg_notice (pr_sequence pr_id (Pfedit.get_all_proof_names()))
| ShowIntros all -> show_intro all
| ShowMatch id -> show_match id
| ShowThesis -> show_thesis ()
@@ -1594,34 +1790,82 @@ let vernac_check_guard () =
with UserError(_,s) ->
(str ("Condition violated: ") ++s)
in
- msgnl message
-
-let interp c = match c with
- (* Control (done in vernac) *)
- | (VernacTime _|VernacList _|VernacLoad _|VernacTimeout _|VernacFail _) ->
- assert false
+ msg_notice message
+
+exception End_of_input
+
+let vernac_load interp fname =
+ 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 open_utf8_file_in fname =
+ let is_bom s =
+ Int.equal (Char.code s.[0]) 0xEF &&
+ Int.equal (Char.code s.[1]) 0xBB &&
+ Int.equal (Char.code s.[2]) 0xBF
+ in
+ let in_chan = open_in fname in
+ let s = " " in
+ if input in_chan s 0 3 < 3 || not (is_bom s) then seek_in in_chan 0;
+ in_chan in
+ let fname =
+ Envars.expand_path_macros ~warn:(fun x -> msg_warning (str x)) fname in
+ let fname = CUnix.make_suffix fname ".v" in
+ let input =
+ let paths = Loadpath.get_paths () in
+ let _,longfname =
+ System.find_file_in_path ~warn:(Flags.is_verbose()) paths fname in
+ let in_chan = open_utf8_file_in longfname in
+ Pcoq.Gram.parsable (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 *)
+let interp ?proof locality poly c =
+ prerr_endline ("interpreting: " ^ Pp.string_of_ppcmds (Ppvernac.pr_vernac c));
+ match c with
+ (* Done later in this file *)
+ | VernacLoad _ -> assert false
+ | VernacFail _ -> assert false
+ | VernacTime _ -> assert false
+ | VernacTimeout _ -> assert false
+ | VernacStm _ -> assert false
+
+ | VernacError e -> raise e
(* Syntax *)
- | VernacTacticNotation (n,r,e) -> Metasyntax.add_tactic_notation (n,r,e)
- | VernacSyntaxExtension (lcl,sl) -> vernac_syntax_extension lcl sl
+ | 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
| VernacBindScope (sc,rl) -> vernac_bind_scope sc rl
- | VernacOpenCloseScope sc -> vernac_open_close_scope sc
- | VernacArgumentsScope (lcl,qid,scl) -> vernac_arguments_scope lcl qid scl
- | VernacInfix (local,mv,qid,sc) -> vernac_infix local mv qid sc
- | VernacNotation (local,c,infpl,sc) -> vernac_notation local c infpl sc
+ | VernacOpenCloseScope (local, s) -> vernac_open_close_scope locality local s
+ | VernacArgumentsScope (qid,scl) -> vernac_arguments_scope locality qid scl
+ | VernacInfix (local,mv,qid,sc) -> vernac_infix locality local mv qid sc
+ | VernacNotation (local,c,infpl,sc) ->
+ vernac_notation locality local c infpl sc
+ | VernacNotationAddFormat(n,k,v) ->
+ Metasyntax.add_notation_extra_printing_rule n k v
(* Gallina *)
- | VernacDefinition (k,lid,d,f) -> vernac_definition k lid d f
- | VernacStartTheoremProof (k,l,top,f) -> vernac_start_proof k l top f
- | VernacEndProof e -> vernac_end_proof e
+ | VernacDefinition (k,lid,d) -> vernac_definition locality poly k lid d
+ | VernacStartTheoremProof (k,l,top) -> vernac_start_proof poly k l top
+ | VernacEndProof e -> vernac_end_proof ?proof e
| VernacExactProof c -> vernac_exact_proof c
- | VernacAssumption (stre,nl,l) -> vernac_assumption stre l nl
- | VernacInductive (finite,infer,l) -> vernac_inductive finite infer l
- | VernacFixpoint l -> vernac_fixpoint l
- | VernacCoFixpoint l -> vernac_cofixpoint l
+ | VernacAssumption (stre,nl,l) -> vernac_assumption locality poly stre l nl
+ | VernacInductive (priv,finite,l) -> vernac_inductive poly priv finite l
+ | VernacFixpoint (local, l) -> vernac_fixpoint locality poly local l
+ | VernacCoFixpoint (local, l) -> vernac_cofixpoint locality poly local l
| VernacScheme l -> vernac_scheme l
| VernacCombinedScheme (id, l) -> vernac_combined_scheme id l
+ | VernacUniverse l -> vernac_universe l
+ | VernacConstraint l -> vernac_constraint l
(* Modules *)
| VernacDeclareModule (export,lid,bl,mtyo) ->
@@ -1637,29 +1881,31 @@ let interp c = match c with
| VernacEndSegment lid -> vernac_end_segment lid
- | VernacRequire (export,spec,qidl) -> vernac_require export spec qidl
+ | VernacNameSectionHypSet (lid, set) -> vernac_name_sec_hyp lid set
+
+ | VernacRequire (export, qidl) -> vernac_require export qidl
| VernacImport (export,qidl) -> vernac_import export qidl
| VernacCanonical qid -> vernac_canonical qid
- | VernacCoercion (str,r,s,t) -> vernac_coercion str r s t
- | VernacIdentityCoercion (str,(_,id),s,t) -> vernac_identity_coercion str id s t
+ | VernacCoercion (local,r,s,t) -> vernac_coercion locality poly local r s t
+ | VernacIdentityCoercion (local,(_,id),s,t) ->
+ vernac_identity_coercion locality poly local id s t
(* Type classes *)
- | VernacInstance (abst, glob, sup, inst, props, pri) ->
- vernac_instance abst glob sup inst props pri
- | VernacContext sup -> vernac_context sup
- | VernacDeclareInstances (glob, ids) -> vernac_declare_instances glob ids
+ | VernacInstance (abst, sup, inst, props, pri) ->
+ vernac_instance abst locality poly sup inst props pri
+ | VernacContext sup -> vernac_context poly sup
+ | VernacDeclareInstances (ids, pri) -> vernac_declare_instances locality ids pri
| VernacDeclareClass id -> vernac_declare_class id
(* Solving *)
- | VernacSolve (n,tac,b) -> vernac_solve n tac b
+ | VernacSolve (n,info,tac,b) -> vernac_solve n info tac b
| VernacSolveExistential (n,c) -> vernac_solve_existential n c
(* Auxiliary file and library management *)
- | VernacRequireFrom (exp,spec,f) -> vernac_require_from exp spec f
| VernacAddLoadPath (isrec,s,alias) -> vernac_add_loadpath isrec s alias
| VernacRemoveLoadPath s -> vernac_remove_loadpath s
| VernacAddMLPath (isrec,s) -> vernac_add_ml_path isrec s
- | VernacDeclareMLModule (local, l) -> vernac_declare_ml_module local l
+ | VernacDeclareMLModule l -> vernac_declare_ml_module locality l
| VernacChdir s -> vernac_chdir s
(* State management *)
@@ -1667,45 +1913,52 @@ let interp c = match c with
| VernacRestoreState s -> vernac_restore_state s
(* Resetting *)
- | VernacResetName id -> vernac_reset_name id
- | VernacResetInitial -> vernac_reset_initial ()
- | VernacBack n -> vernac_back n
- | VernacBackTo n -> vernac_backto n
+ | VernacResetName _ -> anomaly (str "VernacResetName not handled by Stm")
+ | VernacResetInitial -> anomaly (str "VernacResetInitial not handled by Stm")
+ | VernacBack _ -> anomaly (str "VernacBack not handled by Stm")
+ | VernacBackTo _ -> anomaly (str "VernacBackTo not handled by Stm")
(* Commands *)
- | VernacDeclareTacticDefinition def -> vernac_declare_tactic_definition def
- | VernacCreateHintDb (local,dbname,b) -> vernac_create_hintdb local dbname b
- | VernacRemoveHints (local,dbnames,ids) -> vernac_remove_hints local dbnames ids
- | VernacHints (local,dbnames,hints) -> vernac_hints local dbnames hints
- | VernacSyntacticDefinition (id,c,l,b) ->vernac_syntactic_definition id c l b
- | VernacDeclareImplicits (local,qid,l) ->vernac_declare_implicits local qid l
- | VernacArguments (local, qid, l, narg, flags) -> vernac_declare_arguments local qid l narg flags
+ | 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) ->
+ vernac_hints locality poly local dbnames hints
+ | VernacSyntacticDefinition (id,c,local,b) ->
+ vernac_syntactic_definition locality id c local b
+ | VernacDeclareImplicits (qid,l) ->
+ vernac_declare_implicits locality qid l
+ | VernacArguments (qid, l, narg, flags) ->
+ vernac_declare_arguments locality qid l narg flags
| VernacReserve bl -> vernac_reserve bl
- | VernacGeneralizable (local,gen) -> vernac_generalizable local gen
- | VernacSetOpacity (local,qidl) -> vernac_set_opacity local qidl
- | VernacSetOption (locality,key,v) -> vernac_set_option locality key v
- | VernacUnsetOption (locality,key) -> vernac_unset_option locality key
+ | 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
+ | VernacUnsetOption key -> vernac_unset_option locality key
| VernacRemoveOption (key,v) -> vernac_remove_option key v
| VernacAddOption (key,v) -> vernac_add_option key v
| VernacMemOption (key,v) -> vernac_mem_option key v
| VernacPrintOption key -> vernac_print_option key
| VernacCheckMayEval (r,g,c) -> vernac_check_may_eval r g c
- | VernacDeclareReduction (b,s,r) -> vernac_declare_reduction b s r
+ | VernacDeclareReduction (s,r) -> vernac_declare_reduction locality s r
| VernacGlobalCheck c -> vernac_global_check c
| VernacPrint p -> vernac_print p
- | VernacSearch (s,r) -> vernac_search s r
+ | VernacSearch (s,g,r) -> vernac_search s g r
| VernacLocate l -> vernac_locate l
- | VernacComments l -> if_verbose message ("Comments ok\n")
+ | VernacRegister (id, r) -> vernac_register id r
+ | VernacComments l -> if_verbose msg_info (str "Comments ok\n")
| VernacNop -> ()
(* Proof management *)
- | VernacGoal t -> vernac_start_proof Theorem [None,([],t,None)] false (fun _ _->())
- | VernacAbort id -> vernac_abort id
- | VernacAbortAll -> vernac_abort_all ()
- | VernacRestart -> vernac_restart ()
- | VernacUndo n -> vernac_undo n
- | VernacUndoTo n -> vernac_undoto n
- | VernacBacktrack (snum,pnum,naborts) -> vernac_backtrack snum pnum naborts
+ | VernacGoal t -> vernac_start_proof poly Theorem [None,([],t,None)] false
+ | VernacAbort id -> anomaly (str "VernacAbort not handled by Stm")
+ | VernacAbortAll -> anomaly (str "VernacAbortAll not handled by Stm")
+ | VernacRestart -> anomaly (str "VernacRestart not handled by Stm")
+ | VernacUndo _ -> anomaly (str "VernacUndo not handled by Stm")
+ | VernacUndoTo _ -> anomaly (str "VernacUndoTo not handled by Stm")
+ | VernacBacktrack _ -> anomaly (str "VernacBacktrack not handled by Stm")
| VernacFocus n -> vernac_focus n
| VernacUnfocus -> vernac_unfocus ()
| VernacUnfocused -> vernac_unfocused ()
@@ -1714,17 +1967,182 @@ let interp c = match c with
| VernacEndSubproof -> vernac_end_subproof ()
| VernacShow s -> vernac_show s
| VernacCheckGuard -> vernac_check_guard ()
- | VernacProof (None, None) -> print_subgoals ()
- | VernacProof (Some tac, None) -> vernac_set_end_tac tac ; print_subgoals ()
- | VernacProof (None, Some l) -> vernac_set_used_variables l ; print_subgoals ()
+ | VernacProof (None, None) -> ()
+ | VernacProof (Some tac, None) -> vernac_set_end_tac tac
+ | VernacProof (None, Some l) -> vernac_set_used_variables l
| VernacProof (Some tac, Some l) ->
- vernac_set_end_tac tac; vernac_set_used_variables l ; print_subgoals ()
+ vernac_set_end_tac tac; vernac_set_used_variables l
| VernacProofMode mn -> Proof_global.set_proof_mode mn
(* Toplevel control *)
| VernacToplevelControl e -> raise e
(* Extensions *)
- | VernacExtend (opn,args) -> Vernacinterp.call (opn,args)
+ | VernacExtend (opn,args) -> Vernacinterp.call ?locality (opn,args)
+
+ (* Handled elsewhere *)
+ | VernacProgram _
+ | VernacPolymorphic _
+ | VernacLocal _ -> assert false
+
+(* Vernaculars that take a locality flag *)
+let check_vernac_supports_locality c l =
+ match l, c with
+ | None, _ -> ()
+ | Some _, (
+ VernacTacticNotation _
+ | VernacOpenCloseScope _
+ | VernacSyntaxExtension _ | VernacInfix _ | VernacNotation _
+ | VernacDefinition _ | VernacFixpoint _ | VernacCoFixpoint _
+ | VernacAssumption _
+ | VernacCoercion _ | VernacIdentityCoercion _
+ | VernacInstance _ | VernacDeclareInstances _
+ | VernacDeclareMLModule _
+ | VernacDeclareTacticDefinition _
+ | VernacCreateHintDb _ | VernacRemoveHints _ | VernacHints _
+ | VernacSyntacticDefinition _
+ | VernacArgumentsScope _ | VernacDeclareImplicits _ | VernacArguments _
+ | VernacGeneralizable _
+ | VernacSetOpacity _ | VernacSetStrategy _
+ | VernacSetOption _ | VernacUnsetOption _
+ | VernacDeclareReduction _
+ | VernacExtend _
+ | VernacInductive _) -> ()
+ | Some _, _ -> Errors.error "This command does not support Locality"
+
+(* Vernaculars that take a polymorphism flag *)
+let check_vernac_supports_polymorphism c p =
+ match p, c with
+ | None, _ -> ()
+ | Some _, (
+ VernacDefinition _ | VernacFixpoint _ | VernacCoFixpoint _
+ | VernacAssumption _ | VernacInductive _
+ | VernacStartTheoremProof _
+ | VernacCoercion _ | VernacIdentityCoercion _
+ | VernacInstance _ | VernacDeclareInstances _
+ | VernacHints _ | VernacContext _
+ | VernacExtend _ ) -> ()
+ | Some _, _ -> Errors.error "This command does not support Polymorphism"
+
+let enforce_polymorphism = function
+ | None -> Flags.is_universe_polymorphism ()
+ | Some b -> b
+
+(** A global default timeout, controled by option "Set Default Timeout n".
+ Use "Unset Default Timeout" to deactivate it (or set it to 0). *)
+
+let default_timeout = ref None
-let interp c = interp c ; check_locality ()
+let _ =
+ Goptions.declare_int_option
+ { Goptions.optsync = true;
+ Goptions.optdepr = false;
+ Goptions.optname = "the default timeout";
+ Goptions.optkey = ["Default";"Timeout"];
+ Goptions.optread = (fun () -> !default_timeout);
+ Goptions.optwrite = ((:=) default_timeout) }
+
+(** When interpreting a command, the current timeout is initially
+ the default one, but may be modified locally by a Timeout command. *)
+
+let current_timeout = ref None
+
+let vernac_timeout f =
+ match !current_timeout, !default_timeout with
+ | Some n, _ | None, Some n ->
+ let f () = f (); current_timeout := None in
+ Control.timeout n f Timeout
+ | None, None -> f ()
+
+let restore_timeout () = current_timeout := None
+
+let locate_if_not_already loc (e, info) =
+ match Loc.get_loc info with
+ | None -> (e, Loc.add_loc info loc)
+ | Some l -> if Loc.is_ghost l then (e, Loc.add_loc info loc) else (e, info)
+
+exception HasNotFailed
+exception HasFailed of string
+
+let with_fail b f =
+ if not b then f ()
+ else begin try
+ (* If the command actually works, ignore its effects on the state.
+ * Note that error has to be printed in the right state, hence
+ * within the purified function *)
+ Future.purify
+ (fun v ->
+ try f v; raise HasNotFailed
+ with
+ | HasNotFailed as e -> raise e
+ | e ->
+ let e = Errors.push e in
+ raise (HasFailed (Pp.string_of_ppcmds
+ (Errors.iprint (Cerrors.process_vernac_interp_error e)))))
+ ()
+ with e when Errors.noncritical e ->
+ let (e, _) = Errors.push e in
+ match e with
+ | HasNotFailed ->
+ errorlabstrm "Fail" (str "The command has not failed!")
+ | HasFailed msg ->
+ if is_verbose () || !Flags.ide_slave then msg_info
+ (str "The command has indeed failed with message:" ++
+ fnl () ++ str "=> " ++ hov 0 (str msg))
+ | _ -> assert false
+ end
+
+let interp ?(verbosely=true) ?proof (loc,c) =
+ let orig_program_mode = Flags.is_program_mode () in
+ let rec aux ?locality ?polymorphism isprogcmd = function
+ | VernacProgram c when not isprogcmd -> aux ?locality ?polymorphism true c
+ | VernacProgram _ -> Errors.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"
+ | VernacStm (Command c) -> aux ?locality ?polymorphism isprogcmd c
+ | VernacStm (PGLast c) -> aux ?locality ?polymorphism isprogcmd c
+ | VernacStm _ -> assert false (* Done by Stm *)
+ | VernacFail v ->
+ with_fail true (fun () -> aux ?locality ?polymorphism isprogcmd v)
+ | VernacTimeout (n,v) ->
+ current_timeout := Some n;
+ aux ?locality ?polymorphism isprogcmd v
+ | VernacTime v ->
+ System.with_time !Flags.time
+ (aux_list ?locality ?polymorphism isprogcmd) v;
+ | VernacLoad (_,fname) -> vernac_load (aux false) fname
+ | c ->
+ check_vernac_supports_locality c locality;
+ check_vernac_supports_polymorphism c polymorphism;
+ let poly = enforce_polymorphism polymorphism in
+ Obligations.set_program_mode isprogcmd;
+ try
+ vernac_timeout begin fun () ->
+ if verbosely then Flags.verbosely (interp ?proof locality poly) c
+ else Flags.silently (interp ?proof locality poly) c;
+ if orig_program_mode || not !Flags.program_mode || isprogcmd then
+ Flags.program_mode := orig_program_mode
+ end
+ with
+ | reraise when
+ (match reraise with
+ | Timeout -> true
+ | e -> Errors.noncritical e)
+ ->
+ let e = Errors.push reraise in
+ let e = locate_if_not_already loc e in
+ let () = restore_timeout () in
+ Flags.program_mode := orig_program_mode;
+ 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
+let () = Hook.set Stm.interp_hook interp
+let () = Hook.set Stm.process_error_hook Cerrors.process_vernac_interp_error
+let () = Hook.set Stm.with_fail_hook with_fail
diff --git a/toplevel/vernacentries.mli b/toplevel/vernacentries.mli
index 9a400fe6..4b1cd7a0 100644
--- a/toplevel/vernacentries.mli
+++ b/toplevel/vernacentries.mli
@@ -1,52 +1,36 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Names
-open Term
-open Vernacinterp
-open Vernacexpr
-open Topconstr
+open Misctypes
-val dump_global : Libnames.reference Genarg.or_by_notation -> unit
+val dump_global : Libnames.reference or_by_notation -> unit
(** Vernacular entries *)
-val show_script : unit -> unit
val show_prooftree : unit -> unit
val show_node : unit -> unit
(** This function can be used by any command that want to observe terms
- in the context of the current goal, as for instance in pcoq *)
+ in the context of the current goal *)
val get_current_context_of_args : int option -> Evd.evar_map * Environ.env
-type pcoq_hook = {
- start_proof : unit -> unit;
- solve : int -> unit;
- abort : string -> unit;
- search : searchable -> dir_path list * bool -> unit;
- print_name : Libnames.reference Genarg.or_by_notation -> unit;
- print_check : Environ.env -> Environ.unsafe_judgment -> unit;
- print_eval : Reductionops.reduction_function -> Environ.env -> Evd.evar_map -> constr_expr ->
- Environ.unsafe_judgment -> unit;
- show_goal : goal_reference -> unit
-}
-
-val set_pcoq_hook : pcoq_hook -> unit
-
(** The main interpretation function of vernacular expressions *)
-
-val interp : Vernacexpr.vernac_expr -> unit
+val interp :
+ ?verbosely:bool ->
+ ?proof:Proof_global.closed_proof ->
+ Loc.t * Vernacexpr.vernac_expr -> unit
(** Print subgoals when the verbose flag is on.
Meant to be used inside vernac commands from plugins. *)
val print_subgoals : unit -> unit
+val try_print_subgoals : unit -> unit
(** The printing of goals via [print_subgoals] or during
[interp] can be controlled by the following flag.
@@ -67,3 +51,8 @@ val qed_display_script : bool ref
a known inductive type. *)
val make_cases : string -> string list list
+
+val vernac_end_proof :
+ ?proof:Proof_global.closed_proof -> Vernacexpr.proof_end -> unit
+
+val with_fail : bool -> (unit -> unit) -> unit
diff --git a/toplevel/vernacexpr.ml b/toplevel/vernacexpr.ml
deleted file mode 100644
index 02e3eec1..00000000
--- a/toplevel/vernacexpr.ml
+++ /dev/null
@@ -1,508 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-open Compat
-open Util
-open Names
-open Tacexpr
-open Extend
-open Genarg
-open Topconstr
-open Decl_kinds
-open Ppextend
-open Declaremods
-
-(* Toplevel control exceptions *)
-exception Drop
-exception Quit
-
-open Libnames
-open Nametab
-
-type lident = identifier located
-type lname = name located
-type lstring = string located
-type lreference = reference
-
-type class_rawexpr = FunClass | SortClass | RefClass of reference or_by_notation
-
-type goal_identifier = string
-
-type goal_reference =
- | OpenSubgoals
- | NthGoal of int
- | GoalId of goal_identifier
-
-type printable =
- | PrintTables
- | PrintFullContext
- | PrintSectionContext of reference
- | PrintInspect of int
- | PrintGrammar of string
- | PrintLoadPath of dir_path option
- | PrintModules
- | PrintModule of reference
- | PrintModuleType of reference
- | PrintMLLoadPath
- | PrintMLModules
- | PrintName of reference or_by_notation
- | PrintGraph
- | PrintClasses
- | PrintTypeClasses
- | PrintInstances of reference or_by_notation
- | PrintLtac of reference
- | PrintCoercions
- | PrintCoercionPaths of class_rawexpr * class_rawexpr
- | PrintCanonicalConversions
- | PrintUniverses of bool * string option
- | PrintHint of reference or_by_notation
- | PrintHintGoal
- | PrintHintDbName of string
- | PrintRewriteHintDbName of string
- | PrintHintDb
- | PrintScopes
- | PrintScope of string
- | PrintVisibility of string option
- | PrintAbout of reference or_by_notation
- | PrintImplicit of reference or_by_notation
- | PrintAssumptions of bool * reference or_by_notation
-
-type search_about_item =
- | SearchSubPattern of constr_pattern_expr
- | SearchString of string * scope_name option
-
-type searchable =
- | SearchPattern of constr_pattern_expr
- | SearchRewrite of constr_pattern_expr
- | SearchHead of constr_pattern_expr
- | SearchAbout of (bool * search_about_item) list
-
-type locatable =
- | LocateTerm of reference or_by_notation
- | LocateLibrary of reference
- | LocateModule of reference
- | LocateTactic of reference
- | LocateFile of string
-
-type showable =
- | ShowGoal of goal_reference
- | ShowGoalImplicitly of int option
- | ShowProof
- | ShowNode
- | ShowScript
- | ShowExistentials
- | ShowTree
- | ShowProofNames
- | ShowIntros of bool
- | ShowMatch of lident
- | ShowThesis
-
-type comment =
- | CommentConstr of constr_expr
- | CommentString of string
- | CommentInt of int
-
-type hints_expr =
- | HintsResolve of (int option * bool * constr_expr) list
- | HintsImmediate of constr_expr list
- | HintsUnfold of reference list
- | HintsTransparency of reference list * bool
- | HintsConstructors of reference list
- | HintsExtern of int * constr_expr option * raw_tactic_expr
-
-type search_restriction =
- | SearchInside of reference list
- | SearchOutside of reference list
-
-type rec_flag = bool (* true = Rec; false = NoRec *)
-type verbose_flag = bool (* true = Verbose; false = Silent *)
-type opacity_flag = bool (* true = Opaque; false = Transparent *)
-type locality_flag = bool (* true = Local; false = Global *)
-type coercion_flag = bool (* true = AddCoercion false = NoCoercion *)
-type instance_flag = bool option (* Some true = Backward instance; Some false = Forward instance, None = NoInstance *)
-type export_flag = bool (* true = Export; false = Import *)
-type specif_flag = bool (* true = Specification; false = Implementation *)
-type inductive_flag = Decl_kinds.recursivity_kind
-type infer_flag = bool (* true = try to Infer record; false = nothing *)
-type full_locality_flag = bool option (* true = Local; false = Global *)
-type onlyparsing_flag = Flags.compat_version option
- (* Some v = Parse only; None = Print also.
- If v<>Current, it contains the name of the coq version
- which this notation is trying to be compatible with *)
-
-type option_value = Goptionstyp.option_value =
- | BoolValue of bool
- | IntValue of int option
- | StringValue of string
-
-type option_ref_value =
- | StringRefValue of string
- | QualidRefValue of reference
-
-type sort_expr = Glob_term.glob_sort
-
-type definition_expr =
- | ProveBody of local_binder list * constr_expr
- | DefineBody of local_binder list * raw_red_expr option * constr_expr
- * constr_expr option
-
-type fixpoint_expr =
- identifier located * (identifier located option * recursion_order_expr) * local_binder list * constr_expr * constr_expr option
-
-type cofixpoint_expr =
- identifier located * local_binder list * constr_expr * constr_expr option
-
-type local_decl_expr =
- | AssumExpr of lname * constr_expr
- | DefExpr of lname * constr_expr * constr_expr option
-
-type inductive_kind = Inductive_kw | CoInductive | Record | Structure | Class of bool (* true = definitional, false = inductive *)
-type decl_notation = lstring * constr_expr * scope_name option
-type simple_binder = lident list * constr_expr
-type class_binder = lident * constr_expr list
-type 'a with_coercion = coercion_flag * 'a
-type 'a with_instance = instance_flag * 'a
-type 'a with_notation = 'a * decl_notation list
-type 'a with_priority = 'a * int option
-type constructor_expr = (lident * constr_expr) with_coercion
-type constructor_list_or_record_decl_expr =
- | Constructors of constructor_expr list
- | RecordDecl of lident option * local_decl_expr with_instance with_priority with_notation list
-type inductive_expr =
- lident with_coercion * local_binder list * constr_expr option * inductive_kind *
- constructor_list_or_record_decl_expr
-
-type one_inductive_expr =
- lident * local_binder list * constr_expr option * constructor_expr list
-
-type module_ast_inl = module_ast annotated
-type module_binder = bool option * lident list * module_ast_inl
-
-type grammar_tactic_prod_item_expr =
- | TacTerm of string
- | TacNonTerm of loc * string * (Names.identifier * string) option
-
-type syntax_modifier =
- | SetItemLevel of string list * production_level
- | SetLevel of int
- | SetAssoc of gram_assoc
- | SetEntryType of string * simple_constr_prod_entry_key
- | SetOnlyParsing of Flags.compat_version
- | SetFormat of string located
-
-type proof_end =
- | Admitted
- | Proved of opacity_flag * (lident * theorem_kind option) option
-
-type scheme =
- | InductionScheme of bool * reference or_by_notation * sort_expr
- | CaseScheme of bool * reference or_by_notation * sort_expr
- | EqualityScheme of reference or_by_notation
-
-type inline = int option (* inlining level, none for no inlining *)
-
-type bullet =
- | Dash
- | Star
- | Plus
-
-type vernac_expr =
- (* Control *)
- | VernacList of located_vernac_expr list
- | VernacLoad of verbose_flag * string
- | VernacTime of vernac_expr
- | VernacTimeout of int * vernac_expr
- | VernacFail of vernac_expr
-
- (* Syntax *)
- | VernacTacticNotation of int * grammar_tactic_prod_item_expr list * raw_tactic_expr
- | VernacSyntaxExtension of locality_flag * (lstring * syntax_modifier list)
- | VernacOpenCloseScope of (locality_flag * bool * scope_name)
- | VernacDelimiters of scope_name * string
- | VernacBindScope of scope_name * class_rawexpr list
- | VernacInfix of locality_flag * (lstring * syntax_modifier list) *
- constr_expr * scope_name option
- | VernacNotation of
- locality_flag * constr_expr * (lstring * syntax_modifier list) *
- scope_name option
-
- (* Gallina *)
- | VernacDefinition of definition_kind * lident * definition_expr *
- declaration_hook
- | VernacStartTheoremProof of theorem_kind *
- (lident option * (local_binder list * constr_expr * (lident option * recursion_order_expr) option)) list *
- bool * declaration_hook
- | VernacEndProof of proof_end
- | VernacExactProof of constr_expr
- | VernacAssumption of assumption_kind * inline * simple_binder with_coercion list
- | VernacInductive of inductive_flag * infer_flag * (inductive_expr * decl_notation list) list
- | VernacFixpoint of (fixpoint_expr * decl_notation list) list
- | VernacCoFixpoint of (cofixpoint_expr * decl_notation list) list
- | VernacScheme of (lident option * scheme) list
- | VernacCombinedScheme of lident * lident list
-
- (* Gallina extensions *)
- | VernacBeginSection of lident
- | VernacEndSegment of lident
- | VernacRequire of
- export_flag option * specif_flag option * lreference list
- | VernacImport of export_flag * lreference list
- | VernacCanonical of reference or_by_notation
- | VernacCoercion of locality * reference or_by_notation *
- class_rawexpr * class_rawexpr
- | VernacIdentityCoercion of locality * lident *
- class_rawexpr * class_rawexpr
-
- (* Type classes *)
- | VernacInstance of
- bool * (* abstract instance *)
- bool * (* global *)
- local_binder list * (* super *)
- typeclass_constraint * (* instance name, class name, params *)
- constr_expr option * (* props *)
- int option (* Priority *)
-
- | VernacContext of local_binder list
-
- | VernacDeclareInstances of
- bool (* global *) * reference list (* instance names *)
-
- | VernacDeclareClass of reference (* inductive or definition name *)
-
- (* Modules and Module Types *)
- | VernacDeclareModule of bool option * lident *
- module_binder list * module_ast_inl
- | VernacDefineModule of bool option * lident *
- module_binder list * module_ast_inl module_signature * module_ast_inl list
- | VernacDeclareModuleType of lident *
- module_binder list * module_ast_inl list * module_ast_inl list
- | VernacInclude of module_ast_inl list
-
- (* Solving *)
-
- | VernacSolve of int * raw_tactic_expr * bool
- | VernacSolveExistential of int * constr_expr
-
- (* Auxiliary file and library management *)
- | VernacRequireFrom of export_flag option * specif_flag option * string
- | VernacAddLoadPath of rec_flag * string * dir_path option
- | VernacRemoveLoadPath of string
- | VernacAddMLPath of rec_flag * string
- | VernacDeclareMLModule of locality_flag * string list
- | VernacChdir of string option
-
- (* State management *)
- | VernacWriteState of string
- | VernacRestoreState of string
-
- (* Resetting *)
- | VernacResetName of lident
- | VernacResetInitial
- | VernacBack of int
- | VernacBackTo of int
-
- (* Commands *)
- | VernacDeclareTacticDefinition of
- (locality_flag * rec_flag * (reference * bool * raw_tactic_expr) list)
- | VernacCreateHintDb of locality_flag * string * bool
- | VernacRemoveHints of locality_flag * string list * reference list
- | VernacHints of locality_flag * string list * hints_expr
- | VernacSyntacticDefinition of identifier located * (identifier list * constr_expr) *
- locality_flag * onlyparsing_flag
- | VernacDeclareImplicits of locality_flag * reference or_by_notation *
- (explicitation * bool * bool) list list
- | VernacArguments of locality_flag * reference or_by_notation *
- ((name * bool * (loc * string) option * bool * bool) list) list *
- int * [ `SimplDontExposeCase | `SimplNeverUnfold | `Rename | `ExtraScopes
- | `ClearImplicits | `ClearScopes | `DefaultImplicits ] list
- | VernacArgumentsScope of locality_flag * reference or_by_notation *
- scope_name option list
- | VernacReserve of simple_binder list
- | VernacGeneralizable of locality_flag * (lident list) option
- | VernacSetOpacity of
- locality_flag * (Conv_oracle.level * reference or_by_notation list) list
- | VernacUnsetOption of full_locality_flag * Goptions.option_name
- | VernacSetOption of full_locality_flag * Goptions.option_name * option_value
- | VernacAddOption of Goptions.option_name * option_ref_value list
- | VernacRemoveOption of Goptions.option_name * option_ref_value list
- | VernacMemOption of Goptions.option_name * option_ref_value list
- | VernacPrintOption of Goptions.option_name
- | VernacCheckMayEval of raw_red_expr option * int option * constr_expr
- | VernacGlobalCheck of constr_expr
- | VernacDeclareReduction of locality_flag * string * raw_red_expr
- | VernacPrint of printable
- | VernacSearch of searchable * search_restriction
- | VernacLocate of locatable
- | VernacComments of comment list
- | VernacNop
-
- (* Proof management *)
- | VernacGoal of constr_expr
- | VernacAbort of lident option
- | VernacAbortAll
- | VernacRestart
- | VernacUndo of int
- | VernacUndoTo of int
- | VernacBacktrack of int*int*int
- | VernacFocus of int option
- | VernacUnfocus
- | VernacUnfocused
- | VernacBullet of bullet
- | VernacSubproof of int option
- | VernacEndSubproof
- | VernacShow of showable
- | VernacCheckGuard
- | VernacProof of raw_tactic_expr option * lident list option
- | VernacProofMode of string
- (* Toplevel control *)
- | VernacToplevelControl of exn
-
- (* For extension *)
- | VernacExtend of string * raw_generic_argument list
-
-and located_vernac_expr = loc * vernac_expr
-
-
-(** Categories of [vernac_expr] *)
-
-let rec strip_vernac = function
- | VernacTime c | VernacTimeout(_,c) | VernacFail c -> strip_vernac c
- | c -> c (* TODO: what about VernacList ? *)
-
-let rec is_navigation_vernac = function
- | VernacResetInitial
- | VernacResetName _
- | VernacBacktrack _
- | VernacBackTo _
- | VernacBack _ -> true
- | VernacTime c -> is_navigation_vernac c (* Time Back* is harmless *)
- | c -> is_deep_navigation_vernac c
-
-and is_deep_navigation_vernac = function
- | VernacTimeout (_,c) | VernacFail c -> is_navigation_vernac c
- | VernacList l -> List.exists (fun (_,c) -> is_navigation_vernac c) l
- | _ -> false
-
-(* NB: Reset is now allowed again as asked by A. Chlipala *)
-
-let is_reset = function
- | VernacResetInitial | VernacResetName _ -> true
- | _ -> false
-
-(* Locating errors raised just after the dot is parsed but before the
- interpretation phase *)
-
-let syntax_checking_error loc s = user_err_loc (loc,"",Pp.str s)
-
-(**********************************************************************)
-(* Managing locality *)
-
-let locality_flag = ref None
-
-let local_of_bool = function true -> Local | false -> Global
-
-let check_locality () =
- match !locality_flag with
- | Some (loc,true) ->
- syntax_checking_error loc
- "This command does not support the \"Local\" prefix.";
- | Some (loc,false) ->
- syntax_checking_error loc
- "This command does not support the \"Global\" prefix."
- | None -> ()
-
-(** Extracting the locality flag *)
-
-(* Commands which supported an inlined Local flag *)
-
-let enforce_locality_full local =
- let local =
- match !locality_flag with
- | Some (_,false) when local ->
- error "Cannot be simultaneously Local and Global."
- | Some (_,true) when local ->
- error "Use only prefix \"Local\"."
- | None ->
- if local then begin
- Flags.if_warn
- Pp.msg_warning (Pp.str"Obsolete syntax: use \"Local\" as a prefix.");
- Some true
- end else
- None
- | Some (_,b) -> Some b in
- locality_flag := None;
- local
-
-(* Commands which did not supported an inlined Local flag (synonym of
- [enforce_locality_full false]) *)
-
-let use_locality_full () =
- let r = Option.map snd !locality_flag in
- locality_flag := None;
- r
-
-(** Positioning locality for commands supporting discharging and export
- outside of modules *)
-
-(* For commands whose default is to discharge and export:
- Global is the default and is neutral;
- Local in a section deactivates discharge,
- Local not in a section deactivates export *)
-
-let make_locality = function Some true -> true | _ -> false
-
-let use_locality () = make_locality (use_locality_full ())
-
-let use_locality_exp () = local_of_bool (use_locality ())
-
-let enforce_locality local = make_locality (enforce_locality_full local)
-
-let enforce_locality_exp local = local_of_bool (enforce_locality local)
-
-(* For commands whose default is not to discharge and not to export:
- Global forces discharge and export;
- Local is the default and is neutral *)
-
-let use_non_locality () =
- match use_locality_full () with Some false -> false | _ -> true
-
-(* For commands whose default is to not discharge but to export:
- Global in sections forces discharge, Global not in section is the default;
- Local in sections is the default, Local not in section forces non-export *)
-
-let make_section_locality =
- function Some b -> b | None -> Lib.sections_are_opened ()
-
-let use_section_locality () =
- make_section_locality (use_locality_full ())
-
-let enforce_section_locality local =
- make_section_locality (enforce_locality_full local)
-
-(** Positioning locality for commands supporting export but not discharge *)
-
-(* For commands whose default is to export (if not in section):
- Global in sections is forbidden, Global not in section is neutral;
- Local in sections is the default, Local not in section forces non-export *)
-
-let make_module_locality = function
- | Some false ->
- if Lib.sections_are_opened () then
- error "This command does not support the Global option in sections.";
- false
- | Some true -> true
- | None -> false
-
-let use_module_locality () =
- make_module_locality (use_locality_full ())
-
-let enforce_module_locality local =
- make_module_locality (enforce_locality_full local)
-
-(**********************************************************************)
-
diff --git a/toplevel/vernacinterp.ml b/toplevel/vernacinterp.ml
index 0a94c050..17f971fd 100644
--- a/toplevel/vernacinterp.ml
+++ b/toplevel/vernacinterp.ml
@@ -1,35 +1,26 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Pp
open Util
-open Names
-open Libnames
-open Himsg
-open Proof_type
-open Tacinterp
-open Vernacexpr
-
-let disable_drop e =
- if e <> Drop then e
- else UserError("Vernac.disable_drop",(str"Drop is forbidden."))
+open Pp
+open Errors
(* Table of vernac entries *)
let vernac_tab =
(Hashtbl.create 51 :
- (string, Tacexpr.raw_generic_argument list -> unit -> unit) Hashtbl.t)
+ (Vernacexpr.extend_name, (Genarg.raw_generic_argument list -> unit -> unit)) Hashtbl.t)
let vinterp_add s f =
try
Hashtbl.add vernac_tab s f
with Failure _ ->
errorlabstrm "vinterp_add"
- (str"Cannot add the vernac command " ++ str s ++ str" twice.")
+ (str"Cannot add the vernac command " ++ str (fst s) ++ str" twice.")
let overwriting_vinterp_add s f =
begin
@@ -42,25 +33,28 @@ let overwriting_vinterp_add s f =
let vinterp_map s =
try
Hashtbl.find vernac_tab s
- with Not_found ->
+ with Failure _ | Not_found ->
errorlabstrm "Vernac Interpreter"
- (str"Cannot find vernac command " ++ str s ++ str".")
+ (str"Cannot find vernac command " ++ str (fst s) ++ str".")
let vinterp_init () = Hashtbl.clear vernac_tab
(* Interpretation of a vernac command *)
-let call (opn,converted_args) =
+let call ?locality (opn,converted_args) =
let loc = ref "Looking up command" in
try
let callback = vinterp_map opn in
loc:= "Checking arguments";
let hunk = callback converted_args in
loc:= "Executing command";
- hunk()
+ Locality.LocalityFixme.set locality;
+ hunk();
+ Locality.LocalityFixme.assert_consumed()
with
| Drop -> raise Drop
| reraise ->
+ let reraise = Errors.push reraise in
if !Flags.debug then
- msgnl (str"Vernac Interpreter " ++ str !loc);
- raise reraise
+ msg_debug (str"Vernac Interpreter " ++ str !loc);
+ iraise reraise
diff --git a/toplevel/vernacinterp.mli b/toplevel/vernacinterp.mli
index f3ae0b7a..38fce5d1 100644
--- a/toplevel/vernacinterp.mli
+++ b/toplevel/vernacinterp.mli
@@ -1,20 +1,16 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Tacexpr
-
(** Interpretation of extended vernac phrases. *)
-val disable_drop : exn -> exn
-
-val vinterp_add : string -> (raw_generic_argument list -> unit -> unit) -> unit
+val vinterp_add : Vernacexpr.extend_name -> (Genarg.raw_generic_argument list -> unit -> unit) -> unit
val overwriting_vinterp_add :
- string -> (raw_generic_argument list -> unit -> unit) -> unit
+ Vernacexpr.extend_name -> (Genarg.raw_generic_argument list -> unit -> unit) -> unit
val vinterp_init : unit -> unit
-val call : string * raw_generic_argument list -> unit
+val call : ?locality:bool -> Vernacexpr.extend_name * Genarg.raw_generic_argument list -> unit
diff --git a/toplevel/whelp.ml4 b/toplevel/whelp.ml4
index 888a4f1a..daedc30f 100644
--- a/toplevel/whelp.ml4
+++ b/toplevel/whelp.ml4
@@ -1,31 +1,28 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i camlp4deps: "parsing/grammar.cma" i*)
+(*i camlp4deps: "grammar/grammar.cma" i*)
open Flags
open Pp
-open Util
-open System
+open Errors
open Names
open Term
-open Environ
open Glob_term
open Libnames
+open Globnames
open Nametab
open Detyping
open Constrintern
open Dischargedhypsmap
-open Command
open Pfedit
-open Refiner
open Tacmach
-open Syntax_def
+open Misctypes
(* Coq interface to the Whelp query engine developed at
the University of Bologna *)
@@ -60,8 +57,8 @@ let make_whelp_request req c =
let b = Buffer.create 16
let url_char c =
- if 'A' <= c & c <= 'Z' or 'a' <= c & c <= 'z' or
- '0' <= c & c <= '9' or c ='.'
+ if 'A' <= c && c <= 'Z' || 'a' <= c && c <= 'z' ||
+ '0' <= c && c <= '9' || c ='.'
then Buffer.add_char b c
else Buffer.add_string b (Printf.sprintf "%%%2X" (Char.code c))
@@ -72,13 +69,13 @@ let rec url_list_with_sep sep f = function
| [a] -> f a
| a::l -> f a; url_string sep; url_list_with_sep sep f l
-let url_id id = url_string (string_of_id id)
+let url_id id = url_string (Id.to_string id)
let uri_of_dirpath dir =
url_string "cic:/"; url_list_with_sep "/" url_id (List.rev dir)
let error_whelp_unknown_reference ref =
- let qid = Nametab.shortest_qualid_of_global Idset.empty ref in
+ let qid = Nametab.shortest_qualid_of_global Id.Set.empty ref in
errorlabstrm ""
(strbrk "Definitions of the current session, like " ++ pr_qualid qid ++
strbrk ", are not supported in Whelp.")
@@ -86,7 +83,7 @@ let error_whelp_unknown_reference ref =
let uri_of_repr_kn ref (mp,dir,l) =
match mp with
| MPfile sl ->
- uri_of_dirpath (id_of_label l :: repr_dirpath dir @ repr_dirpath sl)
+ uri_of_dirpath (Label.to_id l :: DirPath.repr dir @ DirPath.repr sl)
| _ ->
error_whelp_unknown_reference ref
@@ -94,8 +91,8 @@ let url_paren f l = url_char '('; f l; url_char ')'
let url_bracket f l = url_char '['; f l; url_char ']'
let whelp_of_glob_sort = function
- | GProp Null -> "Prop"
- | GProp Pos -> "Set"
+ | GProp -> "Prop"
+ | GSet -> "Set"
| GType _ -> "Type"
let uri_int n = Buffer.add_string b (string_of_int n)
@@ -105,7 +102,7 @@ let uri_of_ind_pointer l =
let uri_of_global ref =
match ref with
- | VarRef id -> error ("Unknown Whelp reference: "^(string_of_id id)^".")
+ | VarRef id -> error ("Unknown Whelp reference: "^(Id.to_string id)^".")
| ConstRef cst ->
uri_of_repr_kn ref (repr_con cst); url_string ".con"
| IndRef (kn,i) ->
@@ -113,7 +110,7 @@ let uri_of_global ref =
| ConstructRef ((kn,i),j) ->
uri_of_repr_kn ref (repr_mind kn); uri_of_ind_pointer [1;i+1;j]
-let whelm_special = id_of_string "WHELM_ANON_VAR"
+let whelm_special = Id.of_string "WHELM_ANON_VAR"
let url_of_name = function
| Name id -> url_id id
@@ -129,9 +126,9 @@ let uri_params f = function
let get_discharged_hyp_names sp = List.map basename (get_discharged_hyps sp)
let section_parameters = function
- | GRef (_,(ConstructRef ((induri,_),_) | IndRef (induri,_))) ->
+ | GRef (_,(ConstructRef ((induri,_),_) | IndRef (induri,_)),_) ->
get_discharged_hyp_names (path_of_global (IndRef(induri,0)))
- | GRef (_,(ConstRef cst as ref)) ->
+ | GRef (_,(ConstRef cst as ref),_) ->
get_discharged_hyp_names (path_of_global ref)
| _ -> []
@@ -144,10 +141,9 @@ let merge vl al =
let rec uri_of_constr c =
match c with
| GVar (_,id) -> url_id id
- | GRef (_,ref) -> uri_of_global ref
+ | GRef (_,ref,_) -> uri_of_global ref
| GHole _ | GEvar _ -> url_string "?"
| GSort (_,s) -> url_string (whelp_of_glob_sort s)
- | _ -> url_paren (fun () -> match c with
| GApp (_,f,args) ->
let inst,rest = merge (section_parameters f) args in
uri_of_constr f; url_char ' '; uri_params uri_of_constr inst;
@@ -163,30 +159,30 @@ let rec uri_of_constr c =
| GLetIn (_,na,b,c) ->
url_string "let "; url_of_name na; url_string "\\def ";
uri_of_constr b; url_string " in "; uri_of_constr c
- | GCast (_,c, CastConv (_,t)) ->
+ | GCast (_,c, (CastConv t|CastVM t|CastNative t)) ->
uri_of_constr c; url_string ":"; uri_of_constr t
| GRec _ | GIf _ | GLetTuple _ | GCases _ ->
error "Whelp does not support pattern-matching and (co-)fixpoint."
- | GVar _ | GRef _ | GHole _ | GEvar _ | GSort _ | GCast (_,_, CastCoerce) ->
- anomaly "Written w/o parenthesis"
+ | GCast (_,_, CastCoerce) ->
+ anomaly (Pp.str "Written w/o parenthesis")
| GPatVar _ ->
- anomaly "Found constructors not supported in constr") ()
+ anomaly (Pp.str "Found constructors not supported in constr")
let make_string f x = Buffer.reset b; f x; Buffer.contents b
let send_whelp req s =
let url = make_whelp_request req s in
- let command = subst_command_placeholder browser_cmd_fmt url in
- let _ = run_command (fun x -> x) print_string command in ()
+ let command = Util.subst_command_placeholder browser_cmd_fmt url in
+ let _ = CUnix.run_command ~hook:print_string command in ()
-let whelp_constr req c =
- let c = detype false [whelm_special] [] c in
+let whelp_constr env sigma req c =
+ let c = detype false [whelm_special] env sigma c in
send_whelp req (make_string uri_of_constr c)
let whelp_constr_expr req c =
let (sigma,env)= Lemmas.get_current_context () in
- let _,c = interp_open_constr sigma env c in
- whelp_constr req c
+ let _,c = interp_open_constr env sigma c in
+ whelp_constr env sigma req c
let whelp_locate s =
send_whelp "locate" s
@@ -195,9 +191,9 @@ let whelp_elim ind =
send_whelp "elim" (make_string uri_of_global (IndRef ind))
let on_goal f =
- let { Evd.it=goals ; sigma=sigma } = Proof.V82.subgoals (get_pftreestate ()) in
- let gls = { Evd.it=List.hd goals ; sigma = sigma } in
- f (Termops.it_mkNamedProd_or_LetIn (pf_concl gls) (pf_hyps gls))
+ let gls = Proof.V82.subgoals (get_pftreestate ()) in
+ let gls = { gls with Evd.it = List.hd gls.Evd.it } in
+ f (pf_env gls) (project gls) (Termops.it_mkNamedProd_or_LetIn (pf_concl gls) (pf_hyps gls))
type whelp_request =
| Locate of string
@@ -207,21 +203,22 @@ type whelp_request =
let whelp = function
| Locate s -> whelp_locate s
| Elim ind -> whelp_elim ind
- | Constr (s,c) -> whelp_constr s c
+ | Constr (s,c) -> whelp_constr (Global.env()) (Evd.empty) s c
VERNAC ARGUMENT EXTEND whelp_constr_request
| [ "Match" ] -> [ "match" ]
| [ "Instance" ] -> [ "instance" ]
END
-VERNAC COMMAND EXTEND Whelp
+VERNAC COMMAND EXTEND Whelp CLASSIFIED AS QUERY
| [ "Whelp" "Locate" string(s) ] -> [ whelp_locate s ]
| [ "Whelp" "Locate" preident(s) ] -> [ whelp_locate s ]
| [ "Whelp" "Elim" global(r) ] -> [ whelp_elim (Smartlocate.global_inductive_with_alias r) ]
| [ "Whelp" whelp_constr_request(req) constr(c) ] -> [ whelp_constr_expr req c]
END
-VERNAC COMMAND EXTEND WhelpHint
+VERNAC COMMAND EXTEND WhelpHint CLASSIFIED AS QUERY
| [ "Whelp" "Hint" constr(c) ] -> [ whelp_constr_expr "hint" c ]
-| [ "Whelp" "Hint" ] -> [ on_goal (whelp_constr "hint") ]
+| [ "Whelp" "Hint" ] => [ Vernacexpr.VtProofStep false, Vernacexpr.VtLater ] ->
+ [ on_goal (fun env sigma -> whelp_constr env sigma "hint") ]
END
diff --git a/toplevel/whelp.mli b/toplevel/whelp.mli
index dbdecee8..62272c50 100644
--- a/toplevel/whelp.mli
+++ b/toplevel/whelp.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -11,8 +11,6 @@
open Names
open Term
-open Topconstr
-open Environ
type whelp_request =
| Locate of string