aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--checker/closure.ml45
-rw-r--r--checker/closure.mli12
-rw-r--r--checker/reduction.ml13
-rw-r--r--dev/ci/ci-basic-overlay.sh2
-rwxr-xr-xdev/ci/ci-iris-lambda-rust.sh4
-rw-r--r--dev/ci/user-overlays/06686-ccnv-no-proj.sh4
-rw-r--r--dev/doc/changes.md11
-rw-r--r--ide/coq_lex.mll6
-rw-r--r--intf/vernacexpr.ml8
-rw-r--r--kernel/cClosure.ml20
-rw-r--r--kernel/cClosure.mli1
-rw-r--r--kernel/nativecode.ml17
-rw-r--r--kernel/nativecode.mli2
-rw-r--r--kernel/nativeconv.ml11
-rw-r--r--kernel/nativeinstr.mli2
-rw-r--r--kernel/nativelambda.ml5
-rw-r--r--kernel/nativevalues.ml18
-rw-r--r--kernel/nativevalues.mli6
-rw-r--r--kernel/reduction.ml35
-rw-r--r--parsing/g_proofs.ml43
-rw-r--r--parsing/g_vernac.ml418
-rw-r--r--plugins/funind/functional_principles_proofs.ml15
-rw-r--r--plugins/ltac/rewrite.ml13
-rw-r--r--plugins/omega/coq_omega.ml2
-rw-r--r--plugins/ssr/ssrcommon.ml4
-rw-r--r--plugins/ssr/ssrvernac.ml42
-rw-r--r--pretyping/cases.ml6
-rw-r--r--pretyping/evardefine.ml4
-rw-r--r--pretyping/nativenorm.ml28
-rw-r--r--pretyping/reductionops.ml6
-rw-r--r--pretyping/reductionops.mli6
-rw-r--r--pretyping/tacred.ml12
-rw-r--r--pretyping/unification.ml4
-rw-r--r--printing/ppvernac.ml17
-rw-r--r--proofs/clenv.ml2
-rw-r--r--proofs/logic.ml4
-rw-r--r--proofs/tacmach.ml2
-rw-r--r--stm/stm.ml37
-rw-r--r--stm/stm.mli21
-rw-r--r--stm/vernac_classifier.ml17
-rw-r--r--tactics/equality.ml2
-rw-r--r--tactics/tactics.ml8
-rw-r--r--toplevel/coqinit.ml22
-rw-r--r--toplevel/coqinit.mli7
-rw-r--r--toplevel/coqtop.ml52
-rw-r--r--vernac/himsg.ml22
-rw-r--r--vernac/lemmas.ml32
-rw-r--r--vernac/lemmas.mli2
-rw-r--r--vernac/vernacentries.ml32
49 files changed, 341 insertions, 283 deletions
diff --git a/checker/closure.ml b/checker/closure.ml
index 98f8c4a82..14b31e09d 100644
--- a/checker/closure.ml
+++ b/checker/closure.ml
@@ -49,13 +49,6 @@ let with_stats c =
end else
Lazy.force c
-type transparent_state = Id.Pred.t * Cpred.t
-let all_opaque = (Id.Pred.empty, Cpred.empty)
-let all_transparent = (Id.Pred.full, Cpred.full)
-
-let is_transparent_variable (ids, _) id = Id.Pred.mem id ids
-let is_transparent_constant (_, csts) cst = Cpred.mem cst csts
-
module type RedFlagsSig = sig
type reds
type red_kind
@@ -63,8 +56,6 @@ module type RedFlagsSig = sig
val fDELTA : red_kind
val fIOTA : red_kind
val fZETA : red_kind
- val fCONST : Constant.t -> red_kind
- val fVAR : Id.t -> red_kind
val no_red : reds
val red_add : reds -> red_kind -> reds
val mkflags : red_kind list -> reds
@@ -80,51 +71,33 @@ module RedFlags = (struct
type reds = {
r_beta : bool;
r_delta : bool;
- r_const : transparent_state;
r_zeta : bool;
r_evar : bool;
r_iota : bool }
type red_kind = BETA | DELTA | IOTA | ZETA
- | CONST of Constant.t | VAR of Id.t
+
let fBETA = BETA
let fDELTA = DELTA
let fIOTA = IOTA
let fZETA = ZETA
- let fCONST kn = CONST kn
- let fVAR id = VAR id
let no_red = {
r_beta = false;
r_delta = false;
- r_const = all_opaque;
r_zeta = false;
r_evar = false;
r_iota = false }
let red_add red = function
| BETA -> { red with r_beta = true }
- | DELTA -> { red with r_delta = true; r_const = all_transparent }
- | CONST kn ->
- let (l1,l2) = red.r_const in
- { red with r_const = l1, Cpred.add kn l2 }
+ | DELTA -> { red with r_delta = true}
| IOTA -> { red with r_iota = true }
| ZETA -> { red with r_zeta = true }
- | VAR id ->
- let (l1,l2) = red.r_const in
- { red with r_const = Id.Pred.add id l1, l2 }
let mkflags = List.fold_left red_add no_red
let red_set red = function
| BETA -> incr_cnt red.r_beta beta
- | CONST kn ->
- let (_,l) = red.r_const in
- let c = Cpred.mem kn l in
- incr_cnt c delta
- | VAR id -> (* En attendant d'avoir des kn pour les Var *)
- let (l,_) = red.r_const in
- let c = Id.Pred.mem id l in
- incr_cnt c delta
| ZETA -> incr_cnt red.r_zeta zeta
| IOTA -> incr_cnt red.r_iota iota
| DELTA -> (* Used for Rel/Var defined in context *)
@@ -700,6 +673,9 @@ let contract_fix_vect fix =
in
(subs_cons(Array.init nfix make_body, env), thisbody)
+let unfold_projection env p =
+ let pb = lookup_projection p env in
+ Zproj (pb.proj_npars, pb.proj_arg, p)
(*********************************************************************)
(* A machine that inspects the head of a term until it finds an
@@ -720,10 +696,9 @@ let rec knh info m stk =
| FCast(t,_,_) -> knh info t stk
| FProj (p,c) ->
- if red_set info.i_flags (fCONST (Projection.constant p)) then
- (let pb = lookup_projection p (info.i_env) in
- knh info c (Zproj (pb.proj_npars, pb.proj_arg, p)
- :: zupdate m stk))
+ if red_set info.i_flags fDELTA then
+ let s = unfold_projection info.i_env p in
+ knh info c (s :: zupdate m stk)
else (m,stk)
(* cases where knh stops *)
@@ -755,11 +730,11 @@ let rec knr info m stk =
(match get_args n tys f e stk with
Inl e', s -> knit info e' f s
| Inr lam, s -> (lam,s))
- | FFlex(ConstKey kn) when red_set info.i_flags (fCONST (fst kn)) ->
+ | FFlex(ConstKey kn) when red_set info.i_flags fDELTA ->
(match ref_value_cache info (ConstKey kn) with
Some v -> kni info v stk
| None -> (set_norm m; (m,stk)))
- | FFlex(VarKey id) when red_set info.i_flags (fVAR id) ->
+ | FFlex(VarKey id) when red_set info.i_flags fDELTA ->
(match ref_value_cache info (VarKey id) with
Some v -> kni info v stk
| None -> (set_norm m; (m,stk)))
diff --git a/checker/closure.mli b/checker/closure.mli
index ce8c64e30..7bdc21b60 100644
--- a/checker/closure.mli
+++ b/checker/closure.mli
@@ -24,14 +24,6 @@ val with_stats: 'a Lazy.t -> 'a
Rem: reduction of a Rel/Var bound to a term is Delta, but reduction of
a LetIn expression is Letin reduction *)
-type transparent_state = Id.Pred.t * Cpred.t
-
-val all_opaque : transparent_state
-val all_transparent : transparent_state
-
-val is_transparent_variable : transparent_state -> variable -> bool
-val is_transparent_constant : transparent_state -> Constant.t -> bool
-
(* Sets of reduction kinds. *)
module type RedFlagsSig = sig
type reds
@@ -42,8 +34,6 @@ module type RedFlagsSig = sig
val fDELTA : red_kind
val fIOTA : red_kind
val fZETA : red_kind
- val fCONST : Constant.t -> red_kind
- val fVAR : Id.t -> red_kind
(* No reduction at all *)
val no_red : reds
@@ -131,6 +121,8 @@ val eta_expand_stack : stack -> stack
val eta_expand_ind_stack : env -> inductive -> fconstr -> stack ->
(fconstr * stack) -> stack * stack
+val unfold_projection : env -> Projection.t -> stack_member
+
(* To lazy reduce a constr, create a [clos_infos] with
[create_clos_infos], inject the term to reduce with [inject]; then use
a reduction function *)
diff --git a/checker/reduction.ml b/checker/reduction.ml
index d4b258f58..29bb8901e 100644
--- a/checker/reduction.ml
+++ b/checker/reduction.ml
@@ -300,11 +300,6 @@ let oracle_order infos l2r k1 k2 =
if Int.equal n1 n2 then l2r
else n1 < n2
-let unfold_projection infos p c =
- let pb = lookup_projection p (infos_env infos) in
- let s = Zproj (pb.proj_npars, pb.proj_arg, p) in
- (c, s)
-
(* Conversion between [lft1]term1 and [lft2]term2 *)
let rec ccnv univ cv_pb infos lft1 lft2 term1 term2 =
eqappr univ cv_pb infos (lft1, (term1,[])) (lft2, (term2,[]))
@@ -374,12 +369,12 @@ and eqappr univ cv_pb infos (lft1,st1) (lft2,st2) =
eqappr univ cv_pb infos app1 app2)
| (FProj (p1,c1), _) ->
- let (def1, s1) = unfold_projection infos p1 c1 in
- eqappr univ cv_pb infos (lft1, whd_stack infos def1 (s1 :: v1)) appr2
+ let s1 = unfold_projection (infos_env infos) p1 in
+ eqappr univ cv_pb infos (lft1, whd_stack infos c1 (s1 :: v1)) appr2
| (_, FProj (p2,c2)) ->
- let (def2, s2) = unfold_projection infos p2 c2 in
- eqappr univ cv_pb infos appr1 (lft2, whd_stack infos def2 (s2 :: v2))
+ let s2 = unfold_projection (infos_env infos) p2 in
+ eqappr univ cv_pb infos appr1 (lft2, whd_stack infos c2 (s2 :: v2))
(* other constructors *)
| (FLambda _, FLambda _) ->
diff --git a/dev/ci/ci-basic-overlay.sh b/dev/ci/ci-basic-overlay.sh
index 628e89291..784da6f97 100644
--- a/dev/ci/ci-basic-overlay.sh
+++ b/dev/ci/ci-basic-overlay.sh
@@ -70,7 +70,7 @@
# Flocq
########################################################################
: "${Flocq_CI_BRANCH:=master}"
-: "${Flocq_CI_GITURL:=https://scm.gforge.inria.fr/anonscm/git/flocq/flocq.git}"
+: "${Flocq_CI_GITURL:=https://gitlab.inria.fr/flocq/flocq.git}"
########################################################################
# Coquelicot
diff --git a/dev/ci/ci-iris-lambda-rust.sh b/dev/ci/ci-iris-lambda-rust.sh
index cf24d202d..267e13359 100755
--- a/dev/ci/ci-iris-lambda-rust.sh
+++ b/dev/ci/ci-iris-lambda-rust.sh
@@ -34,8 +34,8 @@ git_checkout ${stdpp_CI_BRANCH} ${stdpp_URL_PARTS[0]} ${stdpp_CI_DIR} ${stdpp_UR
# Build std++
( cd ${stdpp_CI_DIR} && make && make install )
-# Build iris
-( cd ${Iris_CI_DIR} && make && make install )
+# Build and validate (except on Travis, i.e., skip if TRAVIS is non-empty) Iris
+( cd ${Iris_CI_DIR} && make && (test -n "${TRAVIS}" || make validate) && make install )
# Build lambdaRust
( cd ${lambdaRust_CI_DIR} && make && make install )
diff --git a/dev/ci/user-overlays/06686-ccnv-no-proj.sh b/dev/ci/user-overlays/06686-ccnv-no-proj.sh
new file mode 100644
index 000000000..3a3ab44e0
--- /dev/null
+++ b/dev/ci/user-overlays/06686-ccnv-no-proj.sh
@@ -0,0 +1,4 @@
+if [ "$CI_PULL_REQUEST" = "6686" ] || [ "$CI_BRANCH" = "ccnv-no-proj" ]; then
+ Equations_CI_BRANCH=ccnv-fixes
+ Equations_CI_GITURL=https://github.com/SkySkimmer/Coq-Equations
+fi
diff --git a/dev/doc/changes.md b/dev/doc/changes.md
index e616bd566..aef62b009 100644
--- a/dev/doc/changes.md
+++ b/dev/doc/changes.md
@@ -58,6 +58,17 @@ Declaration of printers for arguments used only in vernac command
happen. An alternative is to register the corresponding argument as
a value, using "Geninterp.register_val0 wit None".
+### STM API
+
+The STM API has seen a general overhaul. The main change is the
+introduction of a "Coq document" type, which all operations now take
+as a parameter. This effectively functionalize the STM API and will
+allow in the future to handle several documents simultaneously.
+
+The main remarkable point is that key implicit global parameters such
+as load-paths and required modules are now arguments to the document
+creation function. This helps enforcing some key invariants.
+
### XML IDE Protocol
- Before 8.8, `Query` only executed the first command present in the
diff --git a/ide/coq_lex.mll b/ide/coq_lex.mll
index a7e9003db..fcc242e07 100644
--- a/ide/coq_lex.mll
+++ b/ide/coq_lex.mll
@@ -19,8 +19,12 @@ let space = [' ' '\n' '\r' '\t' '\012'] (* '\012' is form-feed *)
let number = [ '0'-'9' ]+
+let string = "\"" _+ "\""
+
let undotted_sep = (number space* ':' space*)? '{' | '}' | '-'+ | '+'+ | '*'+
+let vernac_control = "Fail" | "Time" | "Redirect" space+ string | "Timeout" space+ number
+
let dot_sep = '.' (space | eof)
let utf8_extra_byte = [ '\x80' - '\xBF' ]
@@ -67,7 +71,7 @@ and sentence initial stamp = parse
stamp (utf8_lexeme_start lexbuf) Tags.Script.sentence;
sentence true stamp lexbuf
}
- | undotted_sep {
+ | (vernac_control space+)* undotted_sep {
(* Separators like { or } and bullets * - + are only active
at the start of a sentence *)
if initial then stamp (utf8_lexeme_start lexbuf + String.length (Lexing.lexeme lexbuf) - 1) Tags.Script.sentence;
diff --git a/intf/vernacexpr.ml b/intf/vernacexpr.ml
index 8bd2da2f1..8e0fe54c5 100644
--- a/intf/vernacexpr.ml
+++ b/intf/vernacexpr.ml
@@ -167,6 +167,7 @@ type option_ref_value =
type universe_decl_expr = (Id.t Loc.located list, glob_constraint list) gen_universe_decl
type ident_decl = lident * universe_decl_expr option
+type name_decl = lname * universe_decl_expr option
type sort_expr = Sorts.family
@@ -204,12 +205,12 @@ type inductive_expr =
type one_inductive_expr =
ident_decl * local_binder_expr list * constr_expr option * constructor_expr list
-type typeclass_constraint = (Name.t Loc.located * universe_decl_expr option) * binding_kind * constr_expr
+type typeclass_constraint = name_decl * binding_kind * constr_expr
and typeclass_context = typeclass_constraint list
type proof_expr =
- ident_decl option * (local_binder_expr list * constr_expr)
+ ident_decl * (local_binder_expr list * constr_expr)
type syntax_modifier =
| SetItemLevel of string list * Extend.production_level
@@ -339,7 +340,7 @@ type nonrec vernac_expr =
| VernacNotationAddFormat of string * string * string
(* Gallina *)
- | VernacDefinition of (discharge * definition_object_kind) * ident_decl * definition_expr
+ | VernacDefinition of (discharge * definition_object_kind) * name_decl * definition_expr
| VernacStartTheoremProof of theorem_kind * proof_expr list
| VernacEndProof of proof_end
| VernacExactProof of constr_expr
@@ -449,7 +450,6 @@ type nonrec vernac_expr =
| VernacComments of comment list
(* Proof management *)
- | VernacGoal of constr_expr
| VernacAbort of lident option
| VernacAbortAll
| VernacRestart
diff --git a/kernel/cClosure.ml b/kernel/cClosure.ml
index 4fd274ae1..219ea5b24 100644
--- a/kernel/cClosure.ml
+++ b/kernel/cClosure.ml
@@ -850,6 +850,14 @@ let contract_fix_vect fix =
in
(subs_cons(Array.init nfix make_body, env), thisbody)
+let unfold_projection info p =
+ if red_projection info.i_flags p
+ then
+ let open Declarations in
+ let pb = lookup_projection p (info_env info) in
+ Some (Zproj (pb.proj_npars, pb.proj_arg, Projection.constant p))
+ else None
+
(*********************************************************************)
(* A machine that inspects the head of a term until it finds an
atom or a subterm that may produce a redex (abstraction,
@@ -868,15 +876,9 @@ let rec knh info m stk =
| (None, stk') -> (m,stk'))
| FCast(t,_,_) -> knh info t stk
| FProj (p,c) ->
- let unf = Projection.unfolded p in
- if unf || red_set info.i_flags (fCONST (Projection.constant p)) then
- (match try Some (lookup_projection p (info_env info)) with Not_found -> None with
- | None -> (m, stk)
- | Some pb ->
- knh info c (Zproj (pb.Declarations.proj_npars, pb.Declarations.proj_arg,
- Projection.constant p)
- :: zupdate m stk))
- else (m,stk)
+ (match unfold_projection info p with
+ | None -> (m, stk)
+ | Some s -> knh info c (s :: zupdate m stk))
(* cases where knh stops *)
| (FFlex _|FLetIn _|FConstruct _|FEvar _|
diff --git a/kernel/cClosure.mli b/kernel/cClosure.mli
index 41db0af75..c43fc4623 100644
--- a/kernel/cClosure.mli
+++ b/kernel/cClosure.mli
@@ -166,6 +166,7 @@ val stack_tail : int -> stack -> stack
val stack_nth : stack -> int -> fconstr
val zip_term : (fconstr -> constr) -> constr -> stack -> constr
val eta_expand_stack : stack -> stack
+val unfold_projection : 'a infos -> Projection.t -> stack_member option
(** To lazy reduce a constr, create a [clos_infos] with
[create_clos_infos], inject the term to reduce with [inject]; then use
diff --git a/kernel/nativecode.ml b/kernel/nativecode.ml
index 613b2f2ec..8fa254053 100644
--- a/kernel/nativecode.ml
+++ b/kernel/nativecode.ml
@@ -148,7 +148,7 @@ type symbol =
| SymbMatch of annot_sw
| SymbInd of inductive
| SymbMeta of metavariable
- | SymbEvar of existential
+ | SymbEvar of Evar.t
| SymbLevel of Univ.Level.t
let dummy_symb = SymbValue (dummy_value ())
@@ -162,8 +162,7 @@ let eq_symbol sy1 sy2 =
| SymbMatch sw1, SymbMatch sw2 -> eq_annot_sw sw1 sw2
| SymbInd ind1, SymbInd ind2 -> eq_ind ind1 ind2
| SymbMeta m1, SymbMeta m2 -> Int.equal m1 m2
- | SymbEvar (evk1,args1), SymbEvar (evk2,args2) ->
- Evar.equal evk1 evk2 && Array.for_all2 Constr.equal args1 args2
+ | SymbEvar evk1, SymbEvar evk2 -> Evar.equal evk1 evk2
| SymbLevel l1, SymbLevel l2 -> Univ.Level.equal l1 l2
| _, _ -> false
@@ -176,10 +175,7 @@ let hash_symbol symb =
| SymbMatch sw -> combinesmall 5 (hash_annot_sw sw)
| SymbInd ind -> combinesmall 6 (ind_hash ind)
| SymbMeta m -> combinesmall 7 m
- | SymbEvar (evk,args) ->
- let evh = Evar.hash evk in
- let hl = Array.fold_left (fun h t -> combine h (Constr.hash t)) evh args in
- combinesmall 8 hl
+ | SymbEvar evk -> combinesmall 8 (Evar.hash evk)
| SymbLevel l -> combinesmall 9 (Univ.Level.hash l)
module HashedTypeSymbol = struct
@@ -1047,11 +1043,12 @@ let ml_of_instance instance u =
let tyn = fresh_lname Anonymous in
let i = push_symbol (SymbMeta mv) in
MLapp(MLprimitive Mk_meta, [|get_meta_code i; MLlocal tyn|])
- | Levar(ev,ty) ->
+ | Levar(evk,ty,args) ->
let tyn = fresh_lname Anonymous in
- let i = push_symbol (SymbEvar ev) in
+ let i = push_symbol (SymbEvar evk) in
+ let args = MLarray(Array.map (ml_of_lam env l) args) in
MLlet(tyn, ml_of_lam env l ty,
- MLapp(MLprimitive Mk_evar, [|get_evar_code i;MLlocal tyn|]))
+ MLapp(MLprimitive Mk_evar, [|get_evar_code i;MLlocal tyn; args|]))
| Lprod(dom,codom) ->
let dom = ml_of_lam env l dom in
let codom = ml_of_lam env l codom in
diff --git a/kernel/nativecode.mli b/kernel/nativecode.mli
index d08f49095..7d20054f7 100644
--- a/kernel/nativecode.mli
+++ b/kernel/nativecode.mli
@@ -44,7 +44,7 @@ val get_ind : symbols -> int -> inductive
val get_meta : symbols -> int -> metavariable
-val get_evar : symbols -> int -> existential
+val get_evar : symbols -> int -> Evar.t
val get_level : symbols -> int -> Univ.Level.t
diff --git a/kernel/nativeconv.ml b/kernel/nativeconv.ml
index 9f9102f7d..bfa982136 100644
--- a/kernel/nativeconv.ml
+++ b/kernel/nativeconv.ml
@@ -54,13 +54,18 @@ and conv_accu env pb lvl k1 k2 cu =
conv_atom env pb lvl (atom_of_accu k1) (atom_of_accu k2) cu
else
let cu = conv_atom env pb lvl (atom_of_accu k1) (atom_of_accu k2) cu in
- List.fold_right2 (conv_val env CONV lvl) (args_of_accu k1) (args_of_accu k2) cu
+ Array.fold_right2 (conv_val env CONV lvl) (args_of_accu k1) (args_of_accu k2) cu
and conv_atom env pb lvl a1 a2 cu =
if a1 == a2 then cu
else
match a1, a2 with
- | Ameta _, _ | _, Ameta _ | Aevar _, _ | _, Aevar _ -> assert false
+ | Ameta (m1,_), Ameta (m2,_) ->
+ if Int.equal m1 m2 then cu else raise NotConvertible
+ | Aevar (ev1,_,args1), Aevar (ev2,_,args2) ->
+ if Evar.equal ev1 ev2 then
+ Array.fold_right2 (conv_val env CONV lvl) args1 args2 cu
+ else raise NotConvertible
| Arel i1, Arel i2 ->
if Int.equal i1 i2 then cu else raise NotConvertible
| Aind (ind1,u1), Aind (ind2,u2) ->
@@ -112,7 +117,7 @@ and conv_atom env pb lvl a1 a2 cu =
else conv_accu env CONV lvl ac1 ac2 cu
| Arel _, _ | Aind _, _ | Aconstant _, _ | Asort _, _ | Avar _, _
| Acase _, _ | Afix _, _ | Acofix _, _ | Acofixe _, _ | Aprod _, _
- | Aproj _, _ -> raise NotConvertible
+ | Aproj _, _ | Ameta _, _ | Aevar _, _ -> raise NotConvertible
(* Precondition length t1 = length f1 = length f2 = length t2 *)
and conv_fix env lvl t1 f1 t2 f2 cu =
diff --git a/kernel/nativeinstr.mli b/kernel/nativeinstr.mli
index 928283a4d..48ad88444 100644
--- a/kernel/nativeinstr.mli
+++ b/kernel/nativeinstr.mli
@@ -23,7 +23,7 @@ and lambda =
| Lrel of Name.t * int
| Lvar of Id.t
| Lmeta of metavariable * lambda (* type *)
- | Levar of existential * lambda (* type *)
+ | Levar of Evar.t * lambda (* type *) * lambda array (* arguments *)
| Lprod of lambda * lambda
| Llam of Name.t array * lambda
| Llet of Name.t * lambda * lambda
diff --git a/kernel/nativelambda.ml b/kernel/nativelambda.ml
index b333b0fb9..29b3e59da 100644
--- a/kernel/nativelambda.ml
+++ b/kernel/nativelambda.ml
@@ -453,11 +453,12 @@ let rec lambda_of_constr env sigma c =
let ty = meta_type sigma mv in
Lmeta (mv, lambda_of_constr env sigma ty)
- | Evar ev ->
+ | Evar (evk,args as ev) ->
(match evar_value sigma ev with
| None ->
let ty = evar_type sigma ev in
- Levar(ev, lambda_of_constr env sigma ty)
+ let args = Array.map (lambda_of_constr env sigma) args in
+ Levar(evk, lambda_of_constr env sigma ty, args)
| Some t -> lambda_of_constr env sigma t)
| Cast (c, _, _) -> lambda_of_constr env sigma c
diff --git a/kernel/nativevalues.ml b/kernel/nativevalues.ml
index ae66362ca..3d47b1672 100644
--- a/kernel/nativevalues.ml
+++ b/kernel/nativevalues.ml
@@ -61,7 +61,7 @@ type atom =
| Acofixe of t array * t array * int * t
| Aprod of Name.t * t * (t -> t)
| Ameta of metavariable * t
- | Aevar of existential * t
+ | Aevar of Evar.t * t * t array
| Aproj of Constant.t * accumulator
let accumulate_tag = 0
@@ -132,8 +132,8 @@ let mk_prod_accu s dom codom =
let mk_meta_accu mv ty =
mk_accu (Ameta (mv,ty))
-let mk_evar_accu ev ty =
- mk_accu (Aevar (ev,ty))
+let mk_evar_accu ev ty args =
+ mk_accu (Aevar (ev,ty,args))
let mk_proj_accu kn c =
mk_accu (Aproj (kn,c))
@@ -153,8 +153,7 @@ let accu_nargs (k:accumulator) =
let args_of_accu (k:accumulator) =
let nargs = accu_nargs k in
let f i = (Obj.magic (Obj.field (Obj.magic k) (nargs-i+2)) : t) in
- let t = Array.init nargs f in
- Array.to_list t
+ Array.init nargs f
let is_accu x =
let o = Obj.repr x in
@@ -179,11 +178,10 @@ let force_cofix (cofix : t) =
let atom = atom_of_accu accu in
match atom with
| Acofix(typ,norm,pos,f) ->
- let f = ref f in
- let args = List.rev (args_of_accu accu) in
- List.iter (fun x -> f := !f x) args;
- let v = !f (Obj.magic ()) in
- set_atom_of_accu accu (Acofixe(typ,norm,pos,v));
+ let args = args_of_accu accu in
+ let f = Array.fold_right (fun arg f -> f arg) args f in
+ let v = f (Obj.magic ()) in
+ set_atom_of_accu accu (Acofixe(typ,norm,pos,v));
v
| Acofixe(_,_,_,v) -> v
| _ -> cofix
diff --git a/kernel/nativevalues.mli b/kernel/nativevalues.mli
index 18b877745..993842740 100644
--- a/kernel/nativevalues.mli
+++ b/kernel/nativevalues.mli
@@ -51,7 +51,7 @@ type atom =
| Acofixe of t array * t array * int * t
| Aprod of Name.t * t * (t -> t)
| Ameta of metavariable * t
- | Aevar of existential * t
+ | Aevar of Evar.t * t (* type *) * t array (* arguments *)
| Aproj of Constant.t * accumulator
(* Constructors *)
@@ -68,7 +68,7 @@ val mk_prod_accu : Name.t -> t -> t -> t
val mk_fix_accu : rec_pos -> int -> t array -> t array -> t
val mk_cofix_accu : int -> t array -> t array -> t
val mk_meta_accu : metavariable -> t
-val mk_evar_accu : existential -> t -> t
+val mk_evar_accu : Evar.t -> t -> t array -> t
val mk_proj_accu : Constant.t -> accumulator -> t
val upd_cofix : t -> t -> unit
val force_cofix : t -> t
@@ -84,7 +84,7 @@ val napply : t -> t array -> t
val dummy_value : unit -> t
val atom_of_accu : accumulator -> atom
-val args_of_accu : accumulator -> t list
+val args_of_accu : accumulator -> t array
val accu_nargs : accumulator -> int
val cast_accu : t -> accumulator
diff --git a/kernel/reduction.ml b/kernel/reduction.ml
index 6104f56c6..1724f210d 100644
--- a/kernel/reduction.ml
+++ b/kernel/reduction.ml
@@ -308,17 +308,6 @@ let in_whnf (t,stk) =
| (FFlex _ | FProd _ | FEvar _ | FInd _ | FAtom _ | FRel _ | FProj _) -> true
| FLOCKED -> assert false
-let unfold_projection infos p c =
- let unf = Projection.unfolded p in
- if unf || RedFlags.red_set infos.i_flags (RedFlags.fCONST (Projection.constant p)) then
- (match try Some (lookup_projection p (info_env infos)) with Not_found -> None with
- | Some pb ->
- let s = Zproj (pb.Declarations.proj_npars, pb.Declarations.proj_arg,
- Projection.constant p) in
- Some (c, s)
- | None -> None)
- else None
-
(* Conversion between [lft1]term1 and [lft2]term2 *)
let rec ccnv cv_pb l2r infos lft1 lft2 term1 term2 cuniv =
eqappr cv_pb l2r infos (lft1, (term1,[])) (lft2, (term2,[])) cuniv
@@ -396,13 +385,13 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
(* Projections: prefer unfolding to first-order unification,
which will happen naturally if the terms c1, c2 are not in constructor
form *)
- (match unfold_projection infos p1 c1 with
- | Some (def1,s1) ->
- eqappr cv_pb l2r infos (lft1, (def1, (s1 :: v1))) appr2 cuniv
+ (match unfold_projection infos p1 with
+ | Some s1 ->
+ eqappr cv_pb l2r infos (lft1, (c1, (s1 :: v1))) appr2 cuniv
| None ->
- match unfold_projection infos p2 c2 with
- | Some (def2,s2) ->
- eqappr cv_pb l2r infos appr1 (lft2, (def2, (s2 :: v2))) cuniv
+ match unfold_projection infos p2 with
+ | Some s2 ->
+ eqappr cv_pb l2r infos appr1 (lft2, (c2, (s2 :: v2))) cuniv
| None ->
if Constant.equal (Projection.constant p1) (Projection.constant p2)
&& compare_stack_shape v1 v2 then
@@ -412,9 +401,9 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
raise NotConvertible)
| (FProj (p1,c1), t2) ->
- (match unfold_projection infos p1 c1 with
- | Some (def1,s1) ->
- eqappr cv_pb l2r infos (lft1, (def1, (s1 :: v1))) appr2 cuniv
+ (match unfold_projection infos p1 with
+ | Some s1 ->
+ eqappr cv_pb l2r infos (lft1, (c1, (s1 :: v1))) appr2 cuniv
| None ->
(match t2 with
| FFlex fl2 ->
@@ -425,9 +414,9 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
| _ -> raise NotConvertible))
| (t1, FProj (p2,c2)) ->
- (match unfold_projection infos p2 c2 with
- | Some (def2,s2) ->
- eqappr cv_pb l2r infos appr1 (lft2, (def2, (s2 :: v2))) cuniv
+ (match unfold_projection infos p2 with
+ | Some s2 ->
+ eqappr cv_pb l2r infos appr1 (lft2, (c2, (s2 :: v2))) cuniv
| None ->
(match t1 with
| FFlex fl1 ->
diff --git a/parsing/g_proofs.ml4 b/parsing/g_proofs.ml4
index d88f6fa0d..1c3ba7837 100644
--- a/parsing/g_proofs.ml4
+++ b/parsing/g_proofs.ml4
@@ -28,7 +28,8 @@ GEXTEND Gram
| ":"; l = LIST1 [id = IDENT -> id ] -> l ] ]
;
command:
- [ [ IDENT "Goal"; c = lconstr -> VernacGoal c
+ [ [ IDENT "Goal"; c = lconstr ->
+ VernacDefinition (Decl_kinds.(NoDischarge, Definition), ((Loc.tag ~loc:!@loc Names.Anonymous), None), ProveBody ([], c))
| IDENT "Proof" -> VernacProof (None,None)
| IDENT "Proof" ; IDENT "Mode" ; mn = string -> VernacProofMode mn
| IDENT "Proof"; c = lconstr -> VernacExactProof c
diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4
index d498bda34..3244b0ff2 100644
--- a/parsing/g_vernac.ml4
+++ b/parsing/g_vernac.ml4
@@ -133,6 +133,12 @@ let test_plural_form_types loc kwd = function
warn_plural_command ~loc:!@loc kwd
| _ -> ()
+let lname_of_lident : lident -> lname =
+ Loc.map (fun s -> Name s)
+
+let name_of_ident_decl : ident_decl -> name_decl =
+ on_fst lname_of_lident
+
(* Gallina declarations *)
GEXTEND Gram
GLOBAL: gallina gallina_ext thm_token def_body of_type_with_opt_coercion
@@ -143,17 +149,17 @@ GEXTEND Gram
[ [ thm = thm_token; id = ident_decl; bl = binders; ":"; c = lconstr;
l = LIST0
[ "with"; id = ident_decl; bl = binders; ":"; c = lconstr ->
- (Some id,(bl,c)) ] ->
- VernacStartTheoremProof (thm, (Some id,(bl,c))::l)
+ (id,(bl,c)) ] ->
+ VernacStartTheoremProof (thm, (id,(bl,c))::l)
| stre = assumption_token; nl = inline; bl = assum_list ->
VernacAssumption (stre, nl, bl)
| (kwd,stre) = assumptions_token; nl = inline; bl = assum_list ->
test_plural_form loc kwd bl;
VernacAssumption (stre, nl, bl)
| d = def_token; id = ident_decl; b = def_body ->
- VernacDefinition (d, id, b)
+ VernacDefinition (d, name_of_ident_decl id, b)
| IDENT "Let"; id = identref; b = def_body ->
- VernacDefinition ((DoDischarge, Let), (id, None), b)
+ VernacDefinition ((DoDischarge, Let), (lname_of_lident id, None), b)
(* Gallina inductive declarations *)
| cum = cumulativity_token; priv = private_token; f = finite_token;
indl = LIST1 inductive_definition SEP "with" ->
@@ -623,12 +629,12 @@ GEXTEND Gram
VernacCanonical (ByNotation ntn)
| IDENT "Canonical"; IDENT "Structure"; qid = global; d = def_body ->
let s = coerce_reference_to_id qid in
- VernacDefinition ((NoDischarge,CanonicalStructure),((Loc.tag s),None),d)
+ VernacDefinition ((NoDischarge,CanonicalStructure),((Loc.tag (Name s)),None),d)
(* Coercions *)
| IDENT "Coercion"; qid = global; d = def_body ->
let s = coerce_reference_to_id qid in
- VernacDefinition ((NoDischarge,Coercion),((Loc.tag s),None),d)
+ VernacDefinition ((NoDischarge,Coercion),((Loc.tag (Name s)),None),d)
| IDENT "Identity"; IDENT "Coercion"; f = identref; ":";
s = class_rawexpr; ">->"; t = class_rawexpr ->
VernacIdentityCoercion (f, s, t)
diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml
index 62ca70626..d04887a48 100644
--- a/plugins/funind/functional_principles_proofs.ml
+++ b/plugins/funind/functional_principles_proofs.ml
@@ -324,7 +324,7 @@ let change_eq env sigma hyp_id (context:rel_context) x t end_of_type =
context
in
let new_type_of_hyp =
- Reductionops.nf_betaiota sigma new_type_of_hyp in
+ Reductionops.nf_betaiota env sigma new_type_of_hyp in
let new_ctxt,new_end_of_type =
decompose_prod_n_assum sigma ctxt_size new_type_of_hyp
in
@@ -698,6 +698,7 @@ let build_proof
: tactic =
let rec build_proof_aux do_finalize dyn_infos : tactic =
fun g ->
+ let env = pf_env g in
let sigma = project g in
(* observe (str "proving on " ++ Printer.pr_lconstr_env (pf_env g) term);*)
match EConstr.kind sigma dyn_infos.info with
@@ -794,7 +795,7 @@ let build_proof
do_finalize dyn_infos g
| Lambda _ ->
let new_term =
- Reductionops.nf_beta sigma dyn_infos.info in
+ Reductionops.nf_beta env sigma dyn_infos.info in
build_proof do_finalize {dyn_infos with info = new_term}
g
| LetIn _ ->
@@ -1153,7 +1154,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
let bodies_with_all_params =
Array.map
(fun body ->
- Reductionops.nf_betaiota (project g)
+ Reductionops.nf_betaiota (pf_env g) (project g)
(applist(substl (List.rev (Array.to_list all_funs_with_full_params)) body,
List.rev_map var_of_decl princ_params))
)
@@ -1191,12 +1192,12 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
let body_with_param,num =
let body = get_body fnames.(i) in
let body_with_full_params =
- Reductionops.nf_betaiota (project g) (
+ Reductionops.nf_betaiota (pf_env g) (project g) (
applist(body,List.rev_map var_of_decl full_params))
in
match EConstr.kind (project g) body_with_full_params with
| Fix((_,num),(_,_,bs)) ->
- Reductionops.nf_betaiota (project g)
+ Reductionops.nf_betaiota (pf_env g) (project g)
(
(applist
(substl
@@ -1279,7 +1280,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
nb_rec_hyps = -100;
rec_hyps = [];
info =
- Reductionops.nf_betaiota (project g)
+ Reductionops.nf_betaiota (pf_env g) (project g)
(applist(fix_body,List.rev_map mkVar args_id));
eq_hyps = []
}
@@ -1339,7 +1340,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
nb_rec_hyps = -100;
rec_hyps = [];
info =
- Reductionops.nf_betaiota Evd.empty
+ Reductionops.nf_betaiota (pf_env g) Evd.empty
(applist(fbody_with_full_params,
(List.rev_map var_of_decl princ_params)@
(List.rev_map mkVar args_id)
diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml
index 3cbb11001..acd7a30c4 100644
--- a/plugins/ltac/rewrite.ml
+++ b/plugins/ltac/rewrite.ml
@@ -210,9 +210,9 @@ end) = struct
let t = Reductionops.whd_all env (goalevars evars) ty in
match EConstr.kind (goalevars evars) t, l with
| Prod (na, ty, b), obj :: cstrs ->
- let b = Reductionops.nf_betaiota (goalevars evars) b in
+ let b = Reductionops.nf_betaiota env (goalevars evars) b in
if noccurn (goalevars evars) 1 b (* non-dependent product *) then
- let ty = Reductionops.nf_betaiota (goalevars evars) ty in
+ let ty = Reductionops.nf_betaiota env (goalevars evars) ty in
let (evars, b', arg, cstrs) = aux env evars (subst1 mkProp b) cstrs in
let evars, relty = mk_relty evars env ty obj in
let evars, newarg = app_poly env evars respectful [| ty ; b' ; relty ; arg |] in
@@ -221,7 +221,7 @@ end) = struct
let (evars, b, arg, cstrs) =
aux (push_rel (LocalAssum (na, ty)) env) evars b cstrs
in
- let ty = Reductionops.nf_betaiota (goalevars evars) ty in
+ let ty = Reductionops.nf_betaiota env (goalevars evars) ty in
let pred = mkLambda (na, ty, b) in
let liftarg = mkLambda (na, ty, arg) in
let evars, arg' = app_poly env evars forall_relation [| ty ; pred ; liftarg |] in
@@ -231,7 +231,7 @@ end) = struct
| _, [] ->
(match finalcstr with
| None | Some (_, None) ->
- let t = Reductionops.nf_betaiota (fst evars) ty in
+ let t = Reductionops.nf_betaiota env (fst evars) ty in
let evars, rel = mk_relty evars env t None in
evars, t, rel, [t, Some rel]
| Some (t, Some rel) -> evars, t, rel, [t, Some rel])
@@ -1557,9 +1557,8 @@ let newfail n s =
let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause =
let open Proofview.Notations in
(** For compatibility *)
- let beta_red _ sigma c = Reductionops.nf_betaiota sigma c in
- let beta = Tactics.reduct_in_concl (beta_red, DEFAULTcast) in
- let beta_hyp id = Tactics.reduct_in_hyp beta_red (id, InHyp) in
+ let beta = Tactics.reduct_in_concl (Reductionops.nf_betaiota, DEFAULTcast) in
+ let beta_hyp id = Tactics.reduct_in_hyp Reductionops.nf_betaiota (id, InHyp) in
let treat sigma res =
match res with
| None -> newfail 0 (str "Nothing to rewrite")
diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml
index 869284246..4271c80cd 100644
--- a/plugins/omega/coq_omega.ml
+++ b/plugins/omega/coq_omega.ml
@@ -652,7 +652,7 @@ let decompile af =
(** Backward compat to emulate the old Refine: normalize the goal conclusion *)
let new_hole env sigma c =
- let c = Reductionops.nf_betaiota sigma c in
+ let c = Reductionops.nf_betaiota env sigma c in
Evarutil.new_evar env sigma c
let clever_rewrite_base_poly typ p result theorem =
diff --git a/plugins/ssr/ssrcommon.ml b/plugins/ssr/ssrcommon.ml
index 8493dbdbb..7cdf05117 100644
--- a/plugins/ssr/ssrcommon.ml
+++ b/plugins/ssr/ssrcommon.ml
@@ -563,7 +563,7 @@ let pf_abs_evars_pirrel gl (sigma, c0) =
| _ -> Constr.fold put evlist c in
let evlist = put [] c0 in
if evlist = [] then 0, c0 else
- let pr_constr t = Printer.pr_econstr_env (pf_env gl) sigma (Reductionops.nf_beta (project gl) (EConstr.of_constr t)) in
+ let pr_constr t = Printer.pr_econstr_env (pf_env gl) sigma (Reductionops.nf_beta (pf_env gl) (project gl) (EConstr.of_constr t)) in
pp(lazy(str"evlist=" ++ pr_list (fun () -> str";")
(fun (k,_) -> Evar.print k) evlist));
let evplist =
@@ -894,7 +894,7 @@ let saturate ?(beta=false) ?(bi_types=false) env sigma c ?(ty=Retyping.get_type_
let sigma = create_evar_defs sigma in
let (sigma, x) =
Evarutil.new_evar env sigma
- (if bi_types then Reductionops.nf_betaiota sigma src else src) in
+ (if bi_types then Reductionops.nf_betaiota env sigma src else src) in
loop (EConstr.Vars.subst1 x tgt) ((m - n,x) :: args) sigma (n-1)
| CastType (t, _) -> loop t args sigma n
| LetInType (_, v, _, t) -> loop (EConstr.Vars.subst1 v t) args sigma n
diff --git a/plugins/ssr/ssrvernac.ml4 b/plugins/ssr/ssrvernac.ml4
index c0479dd24..d74ad06b3 100644
--- a/plugins/ssr/ssrvernac.ml4
+++ b/plugins/ssr/ssrvernac.ml4
@@ -553,7 +553,7 @@ GEXTEND Gram
let s = coerce_reference_to_id qid in
Vernacexpr.VernacDefinition
((Decl_kinds.NoDischarge,Decl_kinds.CanonicalStructure),
- ((Loc.tag s),None), d)
+ ((Loc.tag (Name s)),None), d)
]];
END
diff --git a/pretyping/cases.ml b/pretyping/cases.ml
index 1207c967b..311c1c09e 100644
--- a/pretyping/cases.ml
+++ b/pretyping/cases.ml
@@ -1276,7 +1276,7 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn
(* This is a bit too strong I think, in the sense that what we would *)
(* really like is to have beta-iota reduction only at the positions where *)
(* parameters are substituted *)
- let typs = List.map (map_type (nf_betaiota !(pb.evdref))) typs in
+ let typs = List.map (map_type (nf_betaiota pb.env !(pb.evdref))) typs in
(* We build the matrix obtained by expanding the matching on *)
(* "C x1..xn as x" followed by a residual matching on eqn into *)
@@ -1426,7 +1426,7 @@ and match_current pb (initial,tomatch) =
find_predicate pb.caseloc pb.env pb.evdref
pred current indt (names,dep) tomatch in
let ci = make_case_info pb.env (fst mind) pb.casestyle in
- let pred = nf_betaiota !(pb.evdref) pred in
+ let pred = nf_betaiota pb.env !(pb.evdref) pred in
let case =
make_case_or_project pb.env !(pb.evdref) indf ci pred current brvals
in
@@ -1663,7 +1663,7 @@ let rec list_assoc_in_triple x = function
*)
let abstract_tycon ?loc env evdref subst tycon extenv t =
- let t = nf_betaiota !evdref t in (* it helps in some cases to remove K-redex*)
+ let t = nf_betaiota env !evdref t in (* it helps in some cases to remove K-redex*)
let src = match EConstr.kind !evdref t with
| Evar (evk,_) -> (Loc.tag ?loc @@ Evar_kinds.SubEvar evk)
| _ -> (Loc.tag ?loc @@ Evar_kinds.CasesType true) in
diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml
index b646a37f8..fd83795f5 100644
--- a/pretyping/evardefine.ml
+++ b/pretyping/evardefine.ml
@@ -28,8 +28,8 @@ let env_nf_evar sigma env =
let env_nf_betaiotaevar sigma env =
process_rel_context
- (fun d e ->
- push_rel (RelDecl.map_constr (fun c -> Reductionops.nf_betaiota sigma c) d) e) env
+ (fun d env ->
+ push_rel (RelDecl.map_constr (fun c -> Reductionops.nf_betaiota env sigma c) d) env) env
(****************************************)
(* Operations on value/type constraints *)
diff --git a/pretyping/nativenorm.ml b/pretyping/nativenorm.ml
index 79e0afa72..b41e15f5a 100644
--- a/pretyping/nativenorm.ml
+++ b/pretyping/nativenorm.ml
@@ -224,7 +224,7 @@ and nf_accu env sigma accu =
if Int.equal (accu_nargs accu) 0 then nf_atom env sigma atom
else
let a,typ = nf_atom_type env sigma atom in
- let _, args = nf_args env sigma accu typ in
+ let _, args = nf_args env sigma (args_of_accu accu) typ in
mkApp(a,Array.of_list args)
and nf_accu_type env sigma accu =
@@ -232,10 +232,10 @@ and nf_accu_type env sigma accu =
if Int.equal (accu_nargs accu) 0 then nf_atom_type env sigma atom
else
let a,typ = nf_atom_type env sigma atom in
- let t, args = nf_args env sigma accu typ in
+ let t, args = nf_args env sigma (args_of_accu accu) typ in
mkApp(a,Array.of_list args), t
-and nf_args env sigma accu t =
+and nf_args env sigma args t =
let aux arg (t,l) =
let _,dom,codom =
try decompose_prod env t with
@@ -246,7 +246,7 @@ and nf_args env sigma accu t =
let c = nf_val env sigma arg dom in
(subst1 c codom, c::l)
in
- let t,l = List.fold_right aux (args_of_accu accu) (t,[]) in
+ let t,l = Array.fold_right aux args (t,[]) in
t, List.rev l
and nf_bargs env sigma b t =
@@ -277,7 +277,6 @@ and nf_atom env sigma atom =
let codom = nf_type env sigma (codom vn) in
mkProd(n,dom,codom)
| Ameta (mv,_) -> mkMeta mv
- | Aevar (ev,_) -> mkEvar ev
| Aproj(p,c) ->
let c = nf_accu env sigma c in
mkProj(Projection.make p true,c)
@@ -347,9 +346,9 @@ and nf_atom_type env sigma atom =
let env = push_rel (LocalAssum (n,dom)) env in
let codom,s2 = nf_type_sort env sigma (codom vn) in
mkProd(n,dom,codom), Typeops.type_of_product env n s1 s2
- | Aevar(ev,ty) ->
- let ty = nf_type env sigma ty in
- mkEvar ev, ty
+ | Aevar(evk,ty,args) ->
+ let ty = nf_type env sigma ty in
+ nf_evar env sigma evk ty args
| Ameta(mv,ty) ->
let ty = nf_type env sigma ty in
mkMeta mv, ty
@@ -386,6 +385,19 @@ and nf_predicate env sigma ind mip params v pT =
true, mkLambda(name,dom,body)
| _, _ -> false, nf_type env sigma v
+and nf_evar env sigma evk ty args =
+ let evi = try Evd.find sigma evk with Not_found -> assert false in
+ let hyps = Environ.named_context_of_val (Evd.evar_filtered_hyps evi) in
+ if List.is_empty hyps then begin
+ assert (Int.equal (Array.length args) 0);
+ mkEvar (evk, [||]), ty
+ end
+ else
+ let fold accu d = Term.mkNamedProd_or_LetIn d accu in
+ let t = List.fold_left fold ty hyps in
+ let ty, args = nf_args env sigma args t in
+ mkEvar (evk, Array.of_list args), ty
+
let evars_of_evar_map sigma =
{ Nativelambda.evars_val = Evd.existential_opt_value sigma;
Nativelambda.evars_typ = Evd.existential_type sigma;
diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml
index 78de0437d..1893018a9 100644
--- a/pretyping/reductionops.ml
+++ b/pretyping/reductionops.ml
@@ -1241,9 +1241,9 @@ let clos_whd_flags flgs env sigma t =
(CClosure.inject (EConstr.Unsafe.to_constr t)))
with e when is_anomaly e -> user_err Pp.(str "Tried to normalize ill-typed term")
-let nf_beta = clos_norm_flags CClosure.beta (Global.env ())
-let nf_betaiota = clos_norm_flags CClosure.betaiota (Global.env ())
-let nf_betaiotazeta = clos_norm_flags CClosure.betaiotazeta (Global.env ())
+let nf_beta = clos_norm_flags CClosure.beta
+let nf_betaiota = clos_norm_flags CClosure.betaiota
+let nf_betaiotazeta = clos_norm_flags CClosure.betaiotazeta
let nf_all env sigma =
clos_norm_flags CClosure.all env sigma
diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli
index a277864c9..0565baf45 100644
--- a/pretyping/reductionops.mli
+++ b/pretyping/reductionops.mli
@@ -168,9 +168,9 @@ val clos_norm_flags : CClosure.RedFlags.reds -> reduction_function
val clos_whd_flags : CClosure.RedFlags.reds -> reduction_function
(** Same as [(strong whd_beta[delta][iota])], but much faster on big terms *)
-val nf_beta : local_reduction_function
-val nf_betaiota : local_reduction_function
-val nf_betaiotazeta : local_reduction_function
+val nf_beta : reduction_function
+val nf_betaiota : reduction_function
+val nf_betaiotazeta : reduction_function
val nf_all : reduction_function
val nf_evar : evar_map -> constr -> constr
diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml
index f682143f8..9b9408698 100644
--- a/pretyping/tacred.ml
+++ b/pretyping/tacred.ml
@@ -474,7 +474,7 @@ let contract_fix_use_function env sigma f
let nbodies = Array.length recindices in
let make_Fi j = (mkFix((recindices,j),typedbodies), f j) in
let lbodies = List.init nbodies make_Fi in
- substl_checking_arity env (List.rev lbodies) sigma (nf_beta sigma bodies.(bodynum))
+ substl_checking_arity env (List.rev lbodies) sigma (nf_beta env sigma bodies.(bodynum))
let reduce_fix_use_function env sigma f whfun fix stack =
match fix_recarg fix (Stack.append_app_list stack Stack.empty) with
@@ -498,7 +498,7 @@ let contract_cofix_use_function env sigma f
let make_Fi j = (mkCoFix(j,typedbodies), f j) in
let subbodies = List.init nbodies make_Fi in
substl_checking_arity env (List.rev subbodies)
- sigma (nf_beta sigma bodies.(bodynum))
+ sigma (nf_beta env sigma bodies.(bodynum))
let reduce_mind_case_use_function func env sigma mia =
match EConstr.kind sigma mia.mconstr with
@@ -695,7 +695,7 @@ let rec red_elim_const env sigma ref u largs =
let whfun = whd_construct_stack env sigma in
(match reduce_fix_use_function env sigma f whfun (destFix sigma d) lrest with
| NotReducible -> raise Redelimination
- | Reduced (c,rest) -> (nf_beta sigma c, rest), nocase)
+ | Reduced (c,rest) -> (nf_beta env sigma c, rest), nocase)
| EliminationMutualFix (min,refgoal,refinfos) when nargs >= min ->
let rec descend (ref,u) args =
let c = reference_value env sigma ref u in
@@ -710,7 +710,7 @@ let rec red_elim_const env sigma ref u largs =
let whfun = whd_construct_stack env sigma in
(match reduce_fix_use_function env sigma f whfun (destFix sigma d) lrest with
| NotReducible -> raise Redelimination
- | Reduced (c,rest) -> (nf_beta sigma c, rest), nocase)
+ | Reduced (c,rest) -> (nf_beta env sigma c, rest), nocase)
| NotAnElimination when unfold_nonelim ->
let c = reference_value env sigma ref u in
(whd_betaiotazeta sigma (applist (c, largs)), []), nocase
@@ -1101,7 +1101,7 @@ let unfoldoccs env sigma (occs,name) c =
| [] -> ()
| _ -> error_invalid_occurrence rest
in
- nf_betaiotazeta sigma uc
+ nf_betaiotazeta env sigma uc
in
match occs with
| NoOccurrences -> c
@@ -1282,7 +1282,7 @@ let reduce_to_ref_gen allow_product env sigma ref t =
else raise Not_found
with Not_found ->
try
- let t' = nf_betaiota sigma (one_step_reduce env sigma t) in
+ let t' = nf_betaiota env sigma (one_step_reduce env sigma t) in
elimrec env t' l
with NotStepReducible -> error_cannot_recognize ref
in
diff --git a/pretyping/unification.ml b/pretyping/unification.ml
index 8df8f8474..e1720ec95 100644
--- a/pretyping/unification.ml
+++ b/pretyping/unification.ml
@@ -194,7 +194,7 @@ let pose_all_metas_as_evars env evd t =
let ty = if Evd.Metaset.is_empty mvs then ty else aux ty in
let ty =
if Flags.version_strictly_greater Flags.V8_6 || Flags.version_less_or_equal Flags.VOld
- then nf_betaiota evd ty (* How it was in Coq <= 8.4 (but done in logic.ml at this time) *)
+ then nf_betaiota env evd ty (* How it was in Coq <= 8.4 (but done in logic.ml at this time) *)
else ty (* some beta-iota-normalization "regression" in 8.5 and 8.6 *) in
let src = Evd.evar_source_of_meta mv !evdref in
let ev = Evarutil.e_new_evar env evdref ~src ty in
@@ -1277,7 +1277,7 @@ let w_coerce env evd mv c =
let unify_to_type env sigma flags c status u =
let sigma, c = refresh_universes (Some false) env sigma c in
let t = get_type_of env sigma (nf_meta sigma c) in
- let t = nf_betaiota sigma (nf_meta sigma t) in
+ let t = nf_betaiota env sigma (nf_meta sigma t) in
unify_0 env sigma CUMUL flags t u
let unify_type env sigma flags mv status c =
diff --git a/printing/ppvernac.ml b/printing/ppvernac.ml
index 7e68a97e4..950246c53 100644
--- a/printing/ppvernac.ml
+++ b/printing/ppvernac.ml
@@ -71,6 +71,9 @@ open Decl_kinds
| Some loc -> let (b,_) = Loc.unloc loc in
pr_located pr_fqid @@ Loc.tag ~loc:(Loc.make_loc (b,b + String.length (string_of_fqid fqid))) fqid
+ let pr_lname_decl (n, u) =
+ pr_lname n ++ pr_universe_decl u
+
let pr_smart_global = Pputils.pr_or_by_notation pr_reference
let pr_ltac_ref = Libnames.pr_reference
@@ -388,8 +391,6 @@ open Decl_kinds
++ prlist (pr_decl_notation pr_constr) ntn
let pr_statement head (idpl,(bl,c)) =
- assert (not (Option.is_empty idpl));
- let idpl = Option.get idpl in
hov 2
(head ++ spc() ++ pr_ident_decl idpl ++ spc() ++
(match bl with [] -> mt() | _ -> pr_binders bl ++ spc()) ++
@@ -562,8 +563,6 @@ open Decl_kinds
return (keyword "Unfocus")
| VernacUnfocused ->
return (keyword "Unfocused")
- | VernacGoal c ->
- return (keyword "Goal" ++ pr_lconstrarg c)
| VernacAbort id ->
return (keyword "Abort" ++ pr_opt pr_lident id)
| VernacUndo i ->
@@ -674,7 +673,10 @@ open Decl_kinds
(* Gallina *)
| VernacDefinition ((discharge,kind),id,b) -> (* A verifier... *)
let pr_def_token dk =
- keyword (Kindops.string_of_definition_object_kind dk)
+ keyword (
+ if Name.is_anonymous (snd (fst id))
+ then "Goal"
+ else Kindops.string_of_definition_object_kind dk)
in
let pr_reduce = function
| None -> mt()
@@ -691,12 +693,13 @@ open Decl_kinds
in
(pr_binders_arg bl,ty,Some (pr_reduce red ++ pr_lconstr body))
| ProveBody (bl,t) ->
- (pr_binders_arg bl, str" :" ++ pr_spc_lconstr t, None) in
+ let typ u = if snd (fst id) = Anonymous then (assert (bl = []); u) else (str" :" ++ u) in
+ (pr_binders_arg bl, typ (pr_spc_lconstr t), None) in
let (binds,typ,c) = pr_def_body b in
return (
hov 2 (
pr_def_token kind ++ spc()
- ++ pr_ident_decl id ++ binds ++ typ
+ ++ pr_lname_decl id ++ binds ++ typ
++ (match c with
| None -> mt()
| Some cc -> str" :=" ++ spc() ++ cc))
diff --git a/proofs/clenv.ml b/proofs/clenv.ml
index 16798a1d5..9e06d913b 100644
--- a/proofs/clenv.ml
+++ b/proofs/clenv.ml
@@ -498,7 +498,7 @@ let clenv_unify_binding_type clenv c t u =
let clenv_assign_binding clenv k c =
let k_typ = clenv_hnf_constr clenv (clenv_meta_type clenv k) in
- let c_typ = nf_betaiota clenv.evd (clenv_get_type_of clenv c) in
+ let c_typ = nf_betaiota clenv.env clenv.evd (clenv_get_type_of clenv c) in
let status,clenv',c = clenv_unify_binding_type clenv c c_typ k_typ in
let c = EConstr.Unsafe.to_constr c in
{ clenv' with evd = meta_assign k (c,(Conv,status)) clenv'.evd }
diff --git a/proofs/logic.ml b/proofs/logic.ml
index 1d86a0909..5ff5fa38a 100644
--- a/proofs/logic.ml
+++ b/proofs/logic.ml
@@ -334,7 +334,7 @@ let rec mk_refgoals sigma goal goalacc conclty trm =
else
match kind trm with
| Meta _ ->
- let conclty = nf_betaiota sigma (EConstr.of_constr conclty) in
+ let conclty = nf_betaiota env sigma (EConstr.of_constr conclty) in
if !check && occur_meta sigma conclty then
raise (RefinerError (env, sigma, MetaInType conclty));
let (gl,ev,sigma) = mk_goal hyps conclty in
@@ -416,7 +416,7 @@ and mk_hdgoals sigma goal goalacc trm =
match kind trm with
| Cast (c,_, ty) when isMeta c ->
check_typability env sigma ty;
- let (gl,ev,sigma) = mk_goal hyps (nf_betaiota sigma (EConstr.of_constr ty)) in
+ let (gl,ev,sigma) = mk_goal hyps (nf_betaiota env sigma (EConstr.of_constr ty)) in
let ev = EConstr.Unsafe.to_constr ev in
gl::goalacc,ty,sigma,ev
diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml
index d41541251..bdcb4868b 100644
--- a/proofs/tacmach.ml
+++ b/proofs/tacmach.ml
@@ -87,7 +87,7 @@ let pf_e_reduce = pf_apply
let pf_whd_all = pf_reduce whd_all
let pf_hnf_constr = pf_reduce hnf_constr
let pf_nf = pf_reduce simpl
-let pf_nf_betaiota = pf_reduce (fun _ -> nf_betaiota)
+let pf_nf_betaiota = pf_reduce nf_betaiota
let pf_compute = pf_reduce compute
let pf_unfoldn ubinds = pf_reduce (unfoldn ubinds)
let pf_unsafe_type_of = pf_reduce unsafe_type_of
diff --git a/stm/stm.ml b/stm/stm.ml
index 07d8b39bd..b5848c662 100644
--- a/stm/stm.ml
+++ b/stm/stm.ml
@@ -566,8 +566,8 @@ end = struct (* {{{ *)
let reachable id = reachable !vcs id
let mk_branch_name { expr = x } = Branch.make
(match Vernacprop.under_control x with
- | VernacDefinition (_,((_,i),_),_) -> Id.to_string i
- | VernacStartTheoremProof (_,[Some ((_,i),_),_]) -> Id.to_string i
+ | VernacDefinition (_,((_, Name i),_),_) -> Id.to_string i
+ | VernacStartTheoremProof (_,[((_,i),_),_]) -> Id.to_string i
| _ -> "branch")
let edit_branch = Branch.make "edit"
let branch ?root ?pos name kind = vcs := branch !vcs ?root ?pos name kind
@@ -2532,16 +2532,28 @@ end (* }}} *)
(********************************* STM API ************************************)
(******************************************************************************)
+(* Main initalization routine *)
type stm_init_options = {
+ (* The STM will set some internal flags differently depending on the
+ specified [doc_type]. This distinction should dissappear at some
+ some point. *)
doc_type : stm_doc_type;
+
+ (* Initial load path in scope for the document. Usually extracted
+ from -R options / _CoqProject *)
+ iload_path : Mltop.coq_path list;
+
+ (* Require [require_libs] before the initial state is
+ ready. Parameters follow [Library], that is to say,
+ [lib,prefix,import_export] means require library [lib] from
+ optional [prefix] and [import_export] if [Some false/Some true]
+ is used. *)
require_libs : (string * string option * bool option) list;
+
+ (* STM options that apply to the current document. *)
stm_options : AsyncOpts.stm_opt;
-(*
- fb_handler : Feedback.feedback -> unit;
- iload_path : (string list * string * bool) list;
- implicit_std : bool;
-*)
}
+(* fb_handler : Feedback.feedback -> unit; *)
(*
let doc_type_module_name (std : stm_doc_type) =
@@ -2554,7 +2566,7 @@ let init_core () =
if !cur_opt.async_proofs_mode = APon then Control.enable_thread_delay := true;
State.register_root_state ()
-let new_doc { doc_type ; stm_options; require_libs } =
+let new_doc { doc_type ; iload_path; require_libs; stm_options } =
let load_objs libs =
let rq_file (dir, from, exp) =
@@ -2572,6 +2584,11 @@ let new_doc { doc_type ; stm_options; require_libs } =
let doc = VCS.init doc_type Stateid.initial in
+ (* Set load path; important, this has to happen before we declare
+ the library below as [Declaremods/Library] will infer the module
+ name by looking at the load path! *)
+ List.iter Mltop.add_coq_path iload_path;
+
begin match doc_type with
| Interactive ln ->
Safe_typing.allow_delayed_constants := true;
@@ -2588,9 +2605,11 @@ let new_doc { doc_type ; stm_options; require_libs } =
VCS.set_ldir ldir;
set_compilation_hints ln
end;
+
+ (* Import initial libraries. *)
load_objs require_libs;
- (* We record the state here! *)
+ (* We record the state at this point! *)
State.define ~cache:`Yes ~redefine:true (fun () -> ()) Stateid.initial;
Backtrack.record ();
Slaves.init ();
diff --git a/stm/stm.mli b/stm/stm.mli
index 8b5581979..8a4de34b4 100644
--- a/stm/stm.mli
+++ b/stm/stm.mli
@@ -46,15 +46,26 @@ type stm_doc_type =
(* Main initalization routine *)
type stm_init_options = {
+ (* The STM will set some internal flags differently depending on the
+ specified [doc_type]. This distinction should dissappear at some
+ some point. *)
doc_type : stm_doc_type;
+
+ (* Initial load path in scope for the document. Usually extracted
+ from -R options / _CoqProject *)
+ iload_path : Mltop.coq_path list;
+
+ (* Require [require_libs] before the initial state is
+ ready. Parameters follow [Library], that is to say,
+ [lib,prefix,import_export] means require library [lib] from
+ optional [prefix] and [import_export] if [Some false/Some true]
+ is used. *)
require_libs : (string * string option * bool option) list;
+
+ (* STM options that apply to the current document. *)
stm_options : AsyncOpts.stm_opt;
-(*
- fb_handler : Feedback.feedback -> unit;
- iload_path : (string list * string * bool) list;
- implicit_std : bool;
-*)
}
+(* fb_handler : Feedback.feedback -> unit; *)
(** The type of a STM document *)
type doc
diff --git a/stm/vernac_classifier.ml b/stm/vernac_classifier.ml
index 99b56d484..cbbb54e45 100644
--- a/stm/vernac_classifier.ml
+++ b/stm/vernac_classifier.ml
@@ -48,6 +48,11 @@ let declare_vernac_classifier
=
classifiers := !classifiers @ [s,f]
+let idents_of_name : Names.Name.t -> Names.Id.t list =
+ function
+ | Names.Anonymous -> []
+ | Names.Name n -> [n]
+
let classify_vernac e =
let static_classifier ~poly e = match e with
(* Univ poly compatibility: we run it now, so that we can just
@@ -83,18 +88,14 @@ let classify_vernac e =
| VernacSetOption (["Default";"Proof";"Using"],_) -> VtSideff [], VtNow
(* StartProof *)
| VernacDefinition ((Decl_kinds.DoDischarge,_),((_,i),_),ProveBody _) ->
- VtStartProof(default_proof_mode (),Doesn'tGuaranteeOpacity,[i]), VtLater
+ VtStartProof(default_proof_mode (),Doesn'tGuaranteeOpacity, idents_of_name i), VtLater
| VernacDefinition (_,((_,i),_),ProveBody _) ->
let guarantee = if poly then Doesn'tGuaranteeOpacity else GuaranteesOpacity in
- VtStartProof(default_proof_mode (),guarantee,[i]), VtLater
+ VtStartProof(default_proof_mode (),guarantee, idents_of_name i), VtLater
| VernacStartTheoremProof (_,l) ->
- let ids =
- CList.map_filter (function (Some ((_,i),pl), _) -> Some i | _ -> None) l in
+ let ids = List.map (fun (((_, i), _), _) -> i) l in
let guarantee = if poly then Doesn'tGuaranteeOpacity else GuaranteesOpacity in
VtStartProof (default_proof_mode (),guarantee,ids), VtLater
- | VernacGoal _ ->
- let guarantee = if poly then Doesn'tGuaranteeOpacity else GuaranteesOpacity in
- VtStartProof (default_proof_mode (),guarantee,[]), VtLater
| VernacFixpoint (discharge,l) ->
let guarantee =
if discharge = Decl_kinds.DoDischarge || poly then Doesn'tGuaranteeOpacity
@@ -121,7 +122,7 @@ let classify_vernac e =
| VernacAssumption (_,_,l) ->
let ids = List.flatten (List.map (fun (_,(l,_)) -> List.map (fun (id, _) -> snd id) l) l) in
VtSideff ids, VtLater
- | VernacDefinition (_,((_,id),_),DefineBody _) -> VtSideff [id], VtLater
+ | VernacDefinition (_,((_,id),_),DefineBody _) -> VtSideff (idents_of_name id), VtLater
| VernacInductive (_, _,_,l) ->
let ids = List.map (fun (((_,((_,id),_)),_,_,_,cl),_) -> id :: match cl with
| Constructors l -> List.map (fun (_,((_,id),_)) -> id) l
diff --git a/tactics/equality.ml b/tactics/equality.ml
index 22073d39b..450d68436 100644
--- a/tactics/equality.ml
+++ b/tactics/equality.ml
@@ -1583,7 +1583,7 @@ let subst_tuple_term env sigma dep_pair1 dep_pair2 b =
let body = mkApp (lambda_create env sigma (typ,pred_body),[|dep_pair1|]) in
let expected_goal = beta_applist sigma (abst_B,List.map fst e2_list) in
(* Simulate now the normalisation treatment made by Logic.mk_refgoals *)
- let expected_goal = nf_betaiota sigma expected_goal in
+ let expected_goal = nf_betaiota env sigma expected_goal in
(* Retype to get universes right *)
let sigma, expected_goal_ty = Typing.type_of env sigma expected_goal in
let sigma, _ = Typing.type_of env sigma body in
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index 4ee0a8a7b..9fded04db 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -471,7 +471,7 @@ let internal_cut_gen ?(check=true) dir replace id t =
(if check && mem_named_context_val id sign then
user_err (str "Variable " ++ Id.print id ++ str " is already declared.");
push_named_context_val (LocalAssum (id,t)) sign,t,concl,sigma) in
- let nf_t = nf_betaiota sigma t in
+ let nf_t = nf_betaiota env sigma t in
Proofview.tclTHEN
(Proofview.Unsafe.tclEVARS sigma)
(Refine.refine ~typecheck:false begin fun sigma ->
@@ -1728,7 +1728,7 @@ let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind :
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
- let thm_ty0 = nf_betaiota sigma (Retyping.get_type_of env sigma c) in
+ let thm_ty0 = nf_betaiota env sigma (Retyping.get_type_of env sigma c) in
let try_apply thm_ty nprod =
try
let n = nb_prod_modulo_zeta sigma thm_ty - nprod in
@@ -1864,7 +1864,7 @@ let explain_unable_to_apply_lemma ?loc env sigma thm innerclause =
str "."))
let apply_in_once_main flags innerclause env sigma (loc,d,lbind) =
- let thm = nf_betaiota sigma (Retyping.get_type_of env sigma d) in
+ let thm = nf_betaiota env sigma (Retyping.get_type_of env sigma d) in
let rec aux clause =
try progress_with_clause flags innerclause clause
with e when CErrors.noncritical e ->
@@ -2162,7 +2162,7 @@ let apply_type newcl args =
let env = Proofview.Goal.env gl in
let store = Proofview.Goal.extra gl in
Refine.refine ~typecheck:false begin fun sigma ->
- let newcl = nf_betaiota sigma newcl (* As in former Logic.refine *) in
+ let newcl = nf_betaiota env sigma newcl (* As in former Logic.refine *) in
let (sigma, ev) =
Evarutil.new_evar env sigma ~principal:true ~store newcl in
(sigma, applist (ev, args))
diff --git a/toplevel/coqinit.ml b/toplevel/coqinit.ml
index b3b5375bf..10205964a 100644
--- a/toplevel/coqinit.ml
+++ b/toplevel/coqinit.ml
@@ -89,22 +89,26 @@ let ml_path_if c p =
let f x = { recursive = false; path_spec = MlPath x } in
if c then List.map f p else []
-(* Initializes the LoadPath *)
-let init_load_path ~load_init =
- let open Mltop in
+(* LoadPath for toploop toplevels *)
+let toplevel_init_load_path () =
let coqlib = Envars.coqlib () in
- let user_contrib = coqlib/"user-contrib" in
- let xdg_dirs = Envars.xdg_dirs ~warn:(fun x -> Feedback.msg_warning (str x)) in
- let coqpath = Envars.coqpath in
- let coq_path = Names.DirPath.make [Libnames.coq_root] in
-
(* NOTE: These directories are searched from last to first *)
(* first, developer specific directory to open *)
ml_path_if Coq_config.local [coqlib/"dev"] @
(* main loops *)
ml_path_if (Coq_config.local || !Flags.boot) [coqlib/"stm"; coqlib/"ide"] @
- ml_path_if (System.exists_dir (coqlib/"toploop")) [coqlib/"toploop"] @
+ ml_path_if (System.exists_dir (coqlib/"toploop")) [coqlib/"toploop"]
+
+(* LoadPath for Coq user libraries *)
+let libs_init_load_path ~load_init =
+
+ let open Mltop in
+ let coqlib = Envars.coqlib () in
+ let user_contrib = coqlib/"user-contrib" in
+ let xdg_dirs = Envars.xdg_dirs ~warn:(fun x -> Feedback.msg_warning (str x)) in
+ let coqpath = Envars.coqpath in
+ let coq_path = Names.DirPath.make [Libnames.coq_root] in
(* then standard library and plugins *)
[build_stdlib_path ~load_init ~unix_path:(coqlib/"theories") ~coq_path ~with_ml:false;
diff --git a/toplevel/coqinit.mli b/toplevel/coqinit.mli
index 089847f5d..0d2da84bb 100644
--- a/toplevel/coqinit.mli
+++ b/toplevel/coqinit.mli
@@ -19,5 +19,10 @@ val push_include : string -> Names.DirPath.t -> bool -> unit
(** [push_include phys_path log_path implicit] *)
val push_ml_include : string -> unit
-val init_load_path : load_init:bool -> Mltop.coq_path list
+
val init_ocaml_path : unit -> unit
+
+(* LoadPath for toploop toplevels *)
+val toplevel_init_load_path : unit -> Mltop.coq_path list
+(* LoadPath for Coq user libraries *)
+val libs_init_load_path : load_init:bool -> Mltop.coq_path list
diff --git a/toplevel/coqtop.ml b/toplevel/coqtop.ml
index 572092e74..400f7048d 100644
--- a/toplevel/coqtop.ml
+++ b/toplevel/coqtop.ml
@@ -304,6 +304,9 @@ let compile ~time ~verbosely ~f_in ~f_out =
|> prlist_with_sep pr_comma Names.Id.print)
++ str ".")
in
+ let iload_path = Coqinit.libs_init_load_path ~load_init:!load_init in
+ let require_libs = require_libs () in
+ let stm_options = !top_stm_options in
match !compilation_mode with
| BuildVo ->
Flags.record_aux_file := true;
@@ -316,8 +319,9 @@ let compile ~time ~verbosely ~f_in ~f_out =
let doc, sid = Stm.(new_doc
{ doc_type = VoDoc long_f_dot_vo;
- require_libs = require_libs ();
- stm_options = !top_stm_options;
+ iload_path;
+ require_libs;
+ stm_options;
}) in
let doc, sid = load_init_vernaculars ~time doc sid in
@@ -352,8 +356,9 @@ let compile ~time ~verbosely ~f_in ~f_out =
let doc, sid = Stm.(new_doc
{ doc_type = VioDoc long_f_dot_vio;
- require_libs = require_libs ();
- stm_options = !top_stm_options;
+ iload_path;
+ require_libs;
+ stm_options;
}) in
let doc, sid = load_init_vernaculars ~time doc sid in
@@ -424,12 +429,15 @@ let set_vio_checking_j opt j =
prerr_endline "setting the J variable like in 'make vio2vo J=3'";
exit 1
-let schedule_vio_checking () =
- if !vio_files <> [] && !vio_checking then
- Vio_checking.schedule_vio_checking !vio_files_j !vio_files
+let schedule_vio () =
+ (* We must add update the loadpath here as the scheduling process
+ happens outside of the STM *)
+ let iload_path = Coqinit.libs_init_load_path ~load_init:!load_init in
+ List.iter Mltop.add_coq_path iload_path;
-let schedule_vio_compilation () =
- if !vio_files <> [] && not !vio_checking then
+ if !vio_checking then
+ Vio_checking.schedule_vio_checking !vio_files_j !vio_files
+ else
Vio_checking.schedule_vio_compilation !vio_files_j !vio_files
(******************************************************************************)
@@ -496,7 +504,8 @@ let usage batch =
try Envars.set_coqlib ~fail:(fun x -> raise NoCoqLib)
with NoCoqLib -> usage_no_coqlib ()
end;
- let lp = Coqinit.init_load_path ~load_init:!load_init in
+ let lp = Coqinit.toplevel_init_load_path () in
+ (* Necessary for finding the toplevels below *)
List.iter Mltop.add_coq_path lp;
if batch then Usage.print_usage_coqc ()
else begin
@@ -810,8 +819,8 @@ let init_toplevel arglist =
if !print_config then (Envars.print_config stdout Coq_config.all_src_dirs; exit (exitcode ()));
if !print_tags then (print_style_tags (); exit (exitcode ()));
if !filter_opts then (print_string (String.concat "\n" extras); exit 0);
- let lp = Coqinit.init_load_path ~load_init:!load_init in
- List.iter Mltop.add_coq_path lp;
+ let top_lp = Coqinit.toplevel_init_load_path () in
+ List.iter Mltop.add_coq_path top_lp;
Option.iter Mltop.load_ml_object_raw !toploop;
let extras = !toploop_init extras in
if not (CList.is_empty extras) then begin
@@ -840,14 +849,22 @@ let init_toplevel arglist =
We split the codepath here depending whether coqtop is called
in interactive mode or not. *)
- if (not !batch_mode || CList.is_empty !compile_list)
+ (* The condition for starting the interactive mode is a bit
+ convoluted, we should really refactor batch/compilation_mode
+ more. *)
+ if (not !batch_mode
+ || CList.(is_empty !compile_list && is_empty !vio_files && is_empty !vio_tasks))
(* Interactive *)
then begin
+ let iload_path = Coqinit.libs_init_load_path ~load_init:!load_init in
+ let require_libs = require_libs () in
+ let stm_options = !top_stm_options in
try
let doc, sid = Stm.(new_doc
{ doc_type = Interactive !toplevel_name;
- require_libs = require_libs ();
- stm_options = !top_stm_options;
+ iload_path;
+ require_libs;
+ stm_options;
}) in
Some (load_init_vernaculars ~time:!measure_time doc sid)
with any -> flush_all(); fatal_error any
@@ -855,8 +872,9 @@ let init_toplevel arglist =
end else begin
try
compile_files ~time:!measure_time;
- schedule_vio_checking ();
- schedule_vio_compilation ();
+ (* Vio compile pass *)
+ if !vio_files <> [] then schedule_vio ();
+ (* Vio task pass *)
check_vio_tasks ();
(* Allow the user to output an arbitrary state *)
outputstate ();
diff --git a/vernac/himsg.ml b/vernac/himsg.ml
index e8c5aeedd..f00c1e604 100644
--- a/vernac/himsg.ml
+++ b/vernac/himsg.ml
@@ -83,12 +83,12 @@ let rec contract3' env sigma a b c = function
(** Ad-hoc reductions *)
-let j_nf_betaiotaevar sigma j =
+let j_nf_betaiotaevar env sigma j =
{ uj_val = j.uj_val;
- uj_type = Reductionops.nf_betaiota sigma j.uj_type }
+ uj_type = Reductionops.nf_betaiota env sigma j.uj_type }
-let jv_nf_betaiotaevar sigma jl =
- Array.map (fun j -> j_nf_betaiotaevar sigma j) jl
+let jv_nf_betaiotaevar env sigma jl =
+ Array.map (fun j -> j_nf_betaiotaevar env sigma j) jl
(** Printers *)
@@ -258,7 +258,7 @@ let explain_number_branches env sigma cj expn =
str "expects " ++ int expn ++ str " branches."
let explain_ill_formed_branch env sigma c ci actty expty =
- let simp t = Reductionops.nf_betaiota sigma t in
+ let simp t = Reductionops.nf_betaiota env sigma t in
let env = make_all_name_different env sigma in
let pc = pr_leconstr_env env sigma c in
let pa, pe = pr_explicit env sigma (simp actty) (simp expty) in
@@ -295,8 +295,8 @@ let explain_unification_error env sigma p1 p2 = function
| NotSameArgSize | NotSameHead | NoCanonicalStructure ->
(* Error speaks from itself *) []
| ConversionFailed (env,t1,t2) ->
- let t1 = Reductionops.nf_betaiota sigma t1 in
- let t2 = Reductionops.nf_betaiota sigma t2 in
+ let t1 = Reductionops.nf_betaiota env sigma t1 in
+ let t2 = Reductionops.nf_betaiota env sigma t2 in
if EConstr.eq_constr sigma t1 p1 && EConstr.eq_constr sigma t2 p2 then [] else
let env = make_all_name_different env sigma in
if not (EConstr.eq_constr sigma t1 p1) || not (EConstr.eq_constr sigma t2 p2) then
@@ -336,8 +336,8 @@ let explain_unification_error env sigma p1 p2 = function
let explain_actual_type env sigma j t reason =
let env = make_all_name_different env sigma in
- let j = j_nf_betaiotaevar sigma j in
- let t = Reductionops.nf_betaiota sigma t in
+ let j = j_nf_betaiotaevar env sigma j in
+ let t = Reductionops.nf_betaiota env sigma t in
(** Actually print *)
let pe = pr_ne_context_of (str "In environment") env sigma in
let pc = pr_leconstr_env env sigma (Environ.j_val j) in
@@ -351,8 +351,8 @@ let explain_actual_type env sigma j t reason =
ppreason ++ str ".")
let explain_cant_apply_bad_type env sigma (n,exptyp,actualtyp) rator randl =
- let randl = jv_nf_betaiotaevar sigma randl in
- let actualtyp = Reductionops.nf_betaiota sigma actualtyp in
+ let randl = jv_nf_betaiotaevar env sigma randl in
+ let actualtyp = Reductionops.nf_betaiota env sigma actualtyp in
let env = make_all_name_different env sigma in
let actualtyp, exptyp = pr_explicit env sigma actualtyp exptyp in
let nargs = Array.length randl in
diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml
index 6ef310837..57436784b 100644
--- a/vernac/lemmas.ml
+++ b/vernac/lemmas.ml
@@ -210,17 +210,16 @@ let save ?export_seff id const uctx do_guard (locality,poly,kind) hook =
let default_thm_id = Id.of_string "Unnamed_thm"
-let compute_proof_name locality = function
- | Some ((loc,id),pl) ->
- (* We check existence here: it's a bit late at Qed time *)
- if Nametab.exists_cci (Lib.make_path id) || is_section_variable id ||
- locality == Global && Nametab.exists_cci (Lib.make_path_except_section id)
- then
- user_err ?loc (Id.print id ++ str " already exists.");
- id
- | None ->
- let avoid = Id.Set.of_list (Proof_global.get_all_proof_names ()) in
- next_global_ident_away default_thm_id avoid
+let fresh_name_for_anonymous_theorem () =
+ let avoid = Id.Set.of_list (Proof_global.get_all_proof_names ()) in
+ next_global_ident_away default_thm_id avoid
+
+let check_name_freshness locality (loc,id) : unit =
+ (* We check existence here: it's a bit late at Qed time *)
+ if Nametab.exists_cci (Lib.make_path id) || is_section_variable id ||
+ locality == Global && Nametab.exists_cci (Lib.make_path_except_section id)
+ then
+ user_err ?loc (Id.print id ++ str " already exists.")
let save_remaining_recthms (locality,p,kind) norm univs body opaq i (id,(t_i,(_,imps))) =
let t_i = norm t_i in
@@ -435,20 +434,17 @@ let start_proof_with_initialization kind sigma decl recguard thms snl hook =
let start_proof_com ?inference_hook kind thms hook =
let env0 = Global.env () in
let decl = fst (List.hd thms) in
- let evd, decl =
- match decl with
- | None -> Evd.from_env env0, Univdecls.default_univ_decl
- | Some decl ->
- Univdecls.interp_univ_decl_opt env0 (snd decl) in
- let evd, thms = List.fold_left_map (fun evd (sopt,(bl,t)) ->
+ let evd, decl = Univdecls.interp_univ_decl_opt env0 (snd decl) in
+ let evd, thms = List.fold_left_map (fun evd ((id, _), (bl, t)) ->
let evd, (impls, ((env, ctx), imps)) = interp_context_evars env0 evd bl in
let evd, (t', imps') = interp_type_evars_impls ~impls env evd t in
let flags = all_and_fail_flags in
let flags = { flags with use_hook = inference_hook } in
let evd = solve_remaining_evars flags env evd Evd.empty in
let ids = List.map RelDecl.get_name ctx in
+ check_name_freshness (pi1 kind) id;
(* XXX: The nf_evar is critical !! *)
- evd, (compute_proof_name (pi1 kind) sopt,
+ evd, (snd id,
(Evarutil.nf_evar evd (EConstr.it_mkProd_or_LetIn t' ctx),
(ids, imps @ lift_implicits (Context.Rel.nhyps ctx) imps'))))
evd thms in
diff --git a/vernac/lemmas.mli b/vernac/lemmas.mli
index ca92e856b..126dcd8b0 100644
--- a/vernac/lemmas.mli
+++ b/vernac/lemmas.mli
@@ -52,6 +52,8 @@ val standard_proof_terminator :
Proof_global.lemma_possible_guards -> unit declaration_hook ->
Proof_global.proof_terminator
+val fresh_name_for_anonymous_theorem : unit -> Id.t
+
(** {6 ... } *)
(** A hook the next three functions pass to cook_proof *)
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
index 1a02a22a5..be09696d1 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -471,33 +471,40 @@ let vernac_definition_hook p = function
| SubClass -> Class.add_subclass_hook p
| _ -> no_hook
-let vernac_definition ~atts discharge kind ((loc,id as lid),pl) def =
+let vernac_definition ~atts discharge kind ((loc, id), pl) def =
let local = enforce_locality_exp atts.locality discharge in
let hook = vernac_definition_hook atts.polymorphic kind in
- let () = match local with
- | Discharge -> Dumpglob.dump_definition lid true "var"
- | Local | Global -> Dumpglob.dump_definition lid false "def"
+ let () =
+ match id with
+ | Anonymous -> ()
+ | Name n -> let lid = (loc, n) in
+ match local with
+ | Discharge -> Dumpglob.dump_definition lid true "var"
+ | Local | Global -> Dumpglob.dump_definition lid false "def"
in
let program_mode = Flags.is_program_mode () in
+ let name =
+ match id with
+ | Anonymous -> fresh_name_for_anonymous_theorem ()
+ | Name n -> n
+ in
(match def with
| ProveBody (bl,t) -> (* local binders, typ *)
- start_proof_and_print (local, atts.polymorphic, DefinitionBody kind)
- [Some (lid,pl), (bl,t)] hook
+ start_proof_and_print (local, atts.polymorphic, DefinitionBody kind)
+ [((loc, name), pl), (bl, t)] hook
| DefineBody (bl,red_option,c,typ_opt) ->
let red_option = match red_option with
| None -> None
| Some r ->
let sigma, env = Pfedit.get_current_context () in
Some (snd (Hook.get f_interp_redexp env sigma r)) in
- ComDefinition.do_definition ~program_mode id (local, atts.polymorphic, kind) pl bl red_option c typ_opt hook)
+ ComDefinition.do_definition ~program_mode name
+ (local, atts.polymorphic, kind) pl bl red_option c typ_opt hook)
let vernac_start_proof ~atts kind l =
let local = enforce_locality_exp atts.locality NoDischarge in
if Dumpglob.dump () then
- List.iter (fun (id, _) ->
- match id with
- | Some (lid,_) -> Dumpglob.dump_definition lid false "prf"
- | None -> ()) l;
+ List.iter (fun ((id, _), _) -> Dumpglob.dump_definition id false "prf") l;
start_proof_and_print (local, atts.polymorphic, Proof kind) l no_hook
let vernac_end_proof ?proof = function
@@ -1587,7 +1594,7 @@ let vernac_check_may_eval ~atts redexp glopt rc =
| None ->
let evars_of_term c = Evarutil.undefined_evars_of_term sigma' c in
let l = Evar.Set.union (evars_of_term j.Environ.uj_val) (evars_of_term j.Environ.uj_type) in
- let j = { j with Environ.uj_type = Reductionops.nf_betaiota sigma' j.Environ.uj_type } in
+ let j = { j with Environ.uj_type = Reductionops.nf_betaiota env sigma' j.Environ.uj_type } in
Feedback.msg_notice (print_judgment env sigma' j ++
pr_ne_evar_set (fnl () ++ str "where" ++ fnl ()) (mt ()) sigma' l ++
Printer.pr_universe_ctx_set sigma uctx)
@@ -2068,7 +2075,6 @@ let interp ?proof ~atts ~st c =
| VernacComments l -> Flags.if_verbose Feedback.msg_info (str "Comments ok\n")
(* Proof management *)
- | VernacGoal t -> vernac_start_proof ~atts Theorem [None,([],t)]
| VernacFocus n -> vernac_focus n
| VernacUnfocus -> vernac_unfocus ()
| VernacUnfocused -> vernac_unfocused ()