aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--.gitattributes15
-rw-r--r--CHANGES3
-rw-r--r--dev/base_include7
-rw-r--r--dev/ci/user-overlays/06511-ejgallego-econstr+more_fix.sh7
-rw-r--r--dev/ci/user-overlays/06676-gares-proofview-goals-come-with-a-state.sh6
-rw-r--r--dev/top_printers.ml8
-rw-r--r--dev/vm_printers.ml3
-rw-r--r--doc/refman/RefMan-oth.tex5
-rw-r--r--doc/refman/RefMan-tac.tex2
-rw-r--r--engine/proofview.ml133
-rw-r--r--engine/proofview.mli15
-rw-r--r--engine/proofview_monad.ml30
-rw-r--r--engine/proofview_monad.mli21
-rw-r--r--engine/termops.ml4
-rw-r--r--engine/termops.mli7
-rw-r--r--ide/ide_slave.ml2
-rw-r--r--interp/constrintern.ml46
-rw-r--r--interp/constrintern.mli39
-rw-r--r--interp/declare.ml8
-rw-r--r--interp/impargs.ml128
-rw-r--r--interp/impargs.mli6
-rw-r--r--interp/modintern.ml7
-rw-r--r--interp/notation.ml59
-rw-r--r--interp/notation.mli6
-rw-r--r--kernel/cbytecodes.ml23
-rw-r--r--kernel/cbytecodes.mli3
-rw-r--r--kernel/cbytegen.ml22
-rw-r--r--kernel/cemitcodes.ml2
-rw-r--r--kernel/uGraph.ml2
-rw-r--r--kernel/vconv.ml8
-rw-r--r--kernel/vm.ml2
-rw-r--r--kernel/vmvalues.ml30
-rw-r--r--kernel/vmvalues.mli3
-rw-r--r--parsing/cLexer.ml412
-rw-r--r--parsing/cLexer.mli5
-rw-r--r--parsing/pcoq.ml16
-rw-r--r--parsing/pcoq.mli4
-rw-r--r--plugins/funind/glob_term_to_relation.ml47
-rw-r--r--plugins/funind/indfun.ml7
-rw-r--r--plugins/funind/recdef.ml9
-rw-r--r--plugins/ltac/coretactics.ml412
-rw-r--r--plugins/ltac/extratactics.ml410
-rw-r--r--plugins/ltac/rewrite.ml4
-rw-r--r--plugins/ltac/tacintern.ml6
-rw-r--r--plugins/ltac/tacinterp.ml2
-rw-r--r--plugins/setoid_ring/newring.ml2
-rw-r--r--plugins/ssr/ssrcommon.ml2
-rw-r--r--plugins/ssr/ssrvernac.ml48
-rw-r--r--plugins/ssr/ssrview.ml3
-rw-r--r--plugins/ssrmatching/ssrmatching.ml46
-rw-r--r--pretyping/inductiveops.ml7
-rw-r--r--pretyping/inductiveops.mli2
-rw-r--r--pretyping/pretyping.ml3
-rw-r--r--pretyping/pretyping.mli6
-rw-r--r--pretyping/vnorm.ml4
-rw-r--r--printing/ppconstr.ml11
-rw-r--r--printing/pputils.ml16
-rw-r--r--printing/pputils.mli6
-rw-r--r--proofs/proof.ml2
-rw-r--r--proofs/refine.ml2
-rw-r--r--stm/stm.ml1
-rw-r--r--tactics/class_tactics.ml6
-rw-r--r--tactics/equality.ml4
-rw-r--r--tactics/hints.ml8
-rw-r--r--tactics/tacticals.ml6
-rw-r--r--tactics/tacticals.mli1
-rw-r--r--tactics/tactics.ml25
-rw-r--r--tactics/tactics.mli3
-rw-r--r--test-suite/bugs/closed/6878.v8
-rw-r--r--test-suite/output/Load.out6
-rw-r--r--test-suite/output/Load.v7
-rw-r--r--test-suite/output/load/Load_noproof.v1
-rw-r--r--test-suite/output/load/Load_openproof.v1
-rw-r--r--test-suite/output/load/Load_proof.v2
-rw-r--r--test-suite/success/rewrite.v17
-rw-r--r--toplevel/coqinit.ml4
-rw-r--r--toplevel/coqloop.ml11
-rw-r--r--toplevel/coqloop.mli3
-rw-r--r--toplevel/coqtop.ml24
-rw-r--r--toplevel/coqtop.mli3
-rw-r--r--toplevel/vernac.ml155
-rw-r--r--toplevel/vernac.mli2
-rw-r--r--vernac/classes.ml4
-rw-r--r--vernac/comAssumption.ml8
-rw-r--r--vernac/comFixpoint.ml10
-rw-r--r--vernac/comInductive.ml22
-rw-r--r--vernac/comProgramFixpoint.ml4
-rw-r--r--vernac/obligations.ml6
-rw-r--r--vernac/record.ml4
-rw-r--r--vernac/vernacentries.ml33
90 files changed, 716 insertions, 549 deletions
diff --git a/.gitattributes b/.gitattributes
index 51fa208a7..db179c8d2 100644
--- a/.gitattributes
+++ b/.gitattributes
@@ -2,10 +2,18 @@
.gitignore export-ignore
.mailmap export-ignore
-*.out -whitespace
+# Because our commit hook automatically does [apply whitespace=fix] we
+# disable whitespace checking for all files except those where we want
+# it. Otherwise rogue global configuration and forgotten local
+# configuration can break commits.
+* -whitespace
+# tabs are allowed in Makefiles.
+Makefile* whitespace=trailing-space
+tools/CoqMakefile.in whitespace=trailing-space
+
+# in general we don't want tabs.
*.asciidoc whitespace=trailing-space,tab-in-indent
-*.bat whitespace=cr-at-eol,trailing-space,tab-in-indent
*.bib whitespace=trailing-space,tab-in-indent
*.c whitespace=trailing-space,tab-in-indent
*.css whitespace=trailing-space,tab-in-indent
@@ -36,3 +44,6 @@
*.v whitespace=trailing-space,tab-in-indent
*.xml whitespace=trailing-space,tab-in-indent
*.yml whitespace=trailing-space,tab-in-indent
+
+# CR is desired for these Windows files.
+*.bat whitespace=cr-at-eol,trailing-space,tab-in-indent
diff --git a/CHANGES b/CHANGES
index d9322a888..466b4cde5 100644
--- a/CHANGES
+++ b/CHANGES
@@ -54,6 +54,9 @@ Tactics
with let bindings in the parameters.
- The tactic "dtauto" now handles some inductives such as
"@sigT A (fun _ => B)" as non-dependent conjunctions.
+- A bug fixed in "rewrite H in *" and "rewrite H in * |-" may cause a
+ few rare incompatibilities (it was unintendedly recursively
+ rewriting in the side conditions generated by H).
Focusing
diff --git a/dev/base_include b/dev/base_include
index 350ccaa10..3320a2a94 100644
--- a/dev/base_include
+++ b/dev/base_include
@@ -196,7 +196,10 @@ let parse_tac = Pcoq.parse_string Ltac_plugin.Pltac.tactic;;
(* build a term of type glob_constr without type-checking or resolution of
implicit syntax *)
-let e s = Constrintern.intern_constr (Global.env()) (parse_constr s);;
+let e s =
+ let env = Global.env () in
+ let sigma = Evd.from_env env in
+ Constrintern.intern_constr env sigma (parse_constr s);;
(* build a term of type constr with type-checking and resolution of
implicit syntax *)
@@ -229,7 +232,7 @@ let _ = Flags.in_toplevel := true
let _ = Constrextern.set_extern_reference
(fun ?loc _ r -> Libnames.Qualid (loc,Nametab.shortest_qualid_of_global Id.Set.empty r));;
-let go () = Coqloop.loop ~time:false ~state:Option.(get !Coqtop.drop_last_doc)
+let go () = Coqloop.loop ~time:false ~state:Option.(get !Coqloop.drop_last_doc)
let _ =
print_string
diff --git a/dev/ci/user-overlays/06511-ejgallego-econstr+more_fix.sh b/dev/ci/user-overlays/06511-ejgallego-econstr+more_fix.sh
new file mode 100644
index 000000000..4b681909d
--- /dev/null
+++ b/dev/ci/user-overlays/06511-ejgallego-econstr+more_fix.sh
@@ -0,0 +1,7 @@
+ if [ "$CI_PULL_REQUEST" = "6511" ] || [ "$CI_BRANCH" = "econstr+more_fix" ]; then
+ ltac2_CI_BRANCH=econstr+more_fix
+ ltac2_CI_GITURL=https://github.com/ejgallego/ltac2
+
+ Equations_CI_BRANCH=econstr+more_fix
+ Equations_CI_GITURL=https://github.com/ejgallego/Coq-Equations
+fi
diff --git a/dev/ci/user-overlays/06676-gares-proofview-goals-come-with-a-state.sh b/dev/ci/user-overlays/06676-gares-proofview-goals-come-with-a-state.sh
new file mode 100644
index 000000000..2451657d4
--- /dev/null
+++ b/dev/ci/user-overlays/06676-gares-proofview-goals-come-with-a-state.sh
@@ -0,0 +1,6 @@
+if [ "$CI_PULL_REQUEST" = "6676" ] || [ "$CI_BRANCH" = "proofview/goal-w-state" ]; then
+ ltac2_CI_BRANCH=fix-for/6676
+ ltac2_CI_GITURL=https://github.com/gares/ltac2.git
+ Equations_CI_BRANCH=fix-for/6676
+ Equations_CI_GITURL=https://github.com/gares/Coq-Equations.git
+fi
diff --git a/dev/top_printers.ml b/dev/top_printers.ml
index f99e2593d..2c983fd8d 100644
--- a/dev/top_printers.ml
+++ b/dev/top_printers.ml
@@ -313,6 +313,8 @@ let constr_display csr =
in
pp (str (term_display csr) ++fnl ())
+let econstr_display c = constr_display EConstr.Unsafe.(to_constr c) ;;
+
open Format;;
let print_pure_constr csr =
@@ -452,6 +454,8 @@ let print_pure_constr csr =
print_string (Printexc.to_string e);print_flush ();
raise e
+let print_pure_econstr c = print_pure_constr EConstr.Unsafe.(to_constr c) ;;
+
let pploc x = let (l,r) = Loc.unloc x in
print_string"(";print_int l;print_string",";print_int r;print_string")"
@@ -505,7 +509,7 @@ let _ =
(function
[c] when genarg_tag c = unquote (topwit wit_constr) && true ->
let c = out_gen (rawwit wit_constr) c in
- (fun ~atts ~st -> in_current_context constr_display c; st)
+ (fun ~atts ~st -> in_current_context econstr_display c; st)
| _ -> failwith "Vernac extension: cannot occur")
with
e -> pp (CErrors.print e)
@@ -521,7 +525,7 @@ let _ =
(function
[c] when genarg_tag c = unquote (topwit wit_constr) && true ->
let c = out_gen (rawwit wit_constr) c in
- (fun ~atts ~st -> in_current_context print_pure_constr c; st)
+ (fun ~atts ~st -> in_current_context print_pure_econstr c; st)
| _ -> failwith "Vernac extension: cannot occur")
with
e -> pp (CErrors.print e)
diff --git a/dev/vm_printers.ml b/dev/vm_printers.ml
index f819d2e6a..bb7a963aa 100644
--- a/dev/vm_printers.ml
+++ b/dev/vm_printers.ml
@@ -61,7 +61,7 @@ and ppstack s =
and ppatom a =
match a with
| Aid idk -> print_idkey idk
- | Atype u -> print_string "Type(...)"
+ | Asort u -> print_string "Sort(...)"
| Aind(sp,i) -> print_string "Ind(";
print_string (MutInd.to_string sp);
print_string ","; print_int i;
@@ -69,7 +69,6 @@ and ppatom a =
and ppwhd whd =
match whd with
- | Vsort s -> ppsort s
| Vprod _ -> print_string "product"
| Vfun _ -> print_string "function"
| Vfix _ -> print_vfix()
diff --git a/doc/refman/RefMan-oth.tex b/doc/refman/RefMan-oth.tex
index 1cd23c929..bef31d3fa 100644
--- a/doc/refman/RefMan-oth.tex
+++ b/doc/refman/RefMan-oth.tex
@@ -513,6 +513,9 @@ This command loads the file named {\ident}{\tt .v}, searching
successively in each of the directories specified in the {\em
loadpath}. (see Section~\ref{loadpath})
+Files loaded this way cannot leave proofs open, and neither the {\tt
+ Load} command can be use inside a proof.
+
\begin{Variants}
\item {\tt Load {\str}.}\label{Load-str}\\
Loads the file denoted by the string {\str}, where {\str} is any
@@ -530,6 +533,8 @@ successively in each of the directories specified in the {\em
\begin{ErrMsgs}
\item \errindex{Can't find file {\ident} on loadpath}
+\item \errindex{Load is not supported inside proofs}
+\item \errindex{Files processed by Load cannot leave open proofs}
\end{ErrMsgs}
\section[Compiled files]{Compiled files\label{compiled}\index{Compiled files}}
diff --git a/doc/refman/RefMan-tac.tex b/doc/refman/RefMan-tac.tex
index 9635b3ab1..6dca314b4 100644
--- a/doc/refman/RefMan-tac.tex
+++ b/doc/refman/RefMan-tac.tex
@@ -2904,7 +2904,7 @@ This happens if \term$_1$ does not occur in the goal.
rewrite H in H2 at - 2}. In particular a failure will happen if any of
these three simpler tactics fails.
\item \texttt{rewrite H in * |- } will do \texttt{rewrite H in
- H$_i$} for all hypothesis \texttt{H$_i$ <> H}. A success will happen
+ H$_i$} for all hypotheses \texttt{H$_i$} different from \texttt{H}. A success will happen
as soon as at least one of these simpler tactics succeeds.
\item \texttt{rewrite H in *} is a combination of \texttt{rewrite H}
and \texttt{rewrite H in * |-} that succeeds if at
diff --git a/engine/proofview.ml b/engine/proofview.ml
index c41b0b0dc..77a884121 100644
--- a/engine/proofview.ml
+++ b/engine/proofview.ml
@@ -33,7 +33,7 @@ type entry = (EConstr.constr * EConstr.types) list
(* In this version: returns the list of focused goals together with
the [evar_map] context. *)
let proofview p =
- p.comb , p.solution
+ List.map drop_state p.comb , p.solution
let compact el ({ solution } as pv) =
let nf c = Evarutil.nf_evar solution c in
@@ -74,7 +74,7 @@ let dependent_init =
let (gl, _) = EConstr.destEvar sigma econstr in
let ret, { solution = sol; comb = comb } = aux (t sigma econstr) in
let entry = (econstr, typ) :: ret in
- entry, { solution = sol; comb = gl :: comb; shelf = [] }
+ entry, { solution = sol; comb = with_empty_state gl :: comb; shelf = [] }
in
fun t ->
let entry, v = aux t in
@@ -110,7 +110,7 @@ let partial_proof entry pv = CList.map (return_constr pv) (CList.map fst entry)
(* First component is a reverse list of the goals which come before
and second component is the list of the goals which go after (in
the expected order). *)
-type focus_context = Evar.t list * Evar.t list
+type focus_context = goal_with_state list * goal_with_state list
(** Returns a stylised view of a focus_context for use by, for
@@ -120,7 +120,8 @@ type focus_context = Evar.t list * Evar.t list
new nearly identical function everytime. Hence the generic name. *)
(* In this version: the goals in the context, as a "zipper" (the first
list is in reversed order). *)
-let focus_context f = f
+let focus_context (left,right) =
+ (List.map drop_state left, List.map drop_state right)
(** This (internal) function extracts a sublist between two indices,
and returns this sublist together with its context: if it returns
@@ -149,21 +150,35 @@ let unfocus_sublist (left,right) s =
proofview. It returns the focused proofview, and a context for
the focus stack. *)
let focus i j sp =
- let (new_comb, context) = focus_sublist i j sp.comb in
- ( { sp with comb = new_comb } , context )
+ let (new_comb, (left, right)) = focus_sublist i j sp.comb in
+ ( { sp with comb = new_comb } , (left, right) )
+
+let cleared_alias evd g =
+ let evk = drop_state g in
+ let state = get_state g in
+ Option.map (fun g -> goal_with_state g state) (Evarutil.advance evd evk)
(** [undefined defs l] is the list of goals in [l] which are still
unsolved (after advancing cleared goals). Note that order matters. *)
-let undefined defs l =
+let undefined_evars defs l =
List.fold_right (fun evk l ->
match Evarutil.advance defs evk with
| Some evk -> List.add_set Evar.equal evk l
| None -> l) l []
+let goal_with_state_equal x y = Evar.equal (drop_state x) (drop_state y)
+let undefined defs l =
+ List.fold_right (fun evk l ->
+ match cleared_alias defs evk with
+ | Some evk -> List.add_set goal_with_state_equal evk l
+ | None -> l) l []
(** Unfocuses a proofview with respect to a context. *)
-let unfocus c sp =
- { sp with comb = undefined sp.solution (unfocus_sublist c sp.comb) }
+let unfocus (left, right) sp =
+ { sp with comb = undefined sp.solution (unfocus_sublist (left, right) sp.comb) }
+let with_empty_state = Proofview_monad.with_empty_state
+let drop_state = Proofview_monad.drop_state
+let goal_with_state = Proofview_monad.goal_with_state
(** {6 The tactic monad} *)
@@ -406,7 +421,8 @@ let tclFOCUSID id t =
try
let ev = Evd.evar_key id initial.solution in
try
- let n = CList.index Evar.equal ev initial.comb in
+ let comb = CList.map drop_state initial.comb in
+ let n = CList.index Evar.equal ev comb in
(* goal is already under focus *)
let (focused,context) = focus n n initial in
Pv.set focused >>
@@ -415,7 +431,7 @@ let tclFOCUSID id t =
return result
with Not_found ->
(* otherwise, save current focus and work purely on the shelve *)
- Comb.set [ev] >>
+ Comb.set [with_empty_state ev] >>
t >>= fun result ->
Comb.set initial.comb >>
return result
@@ -445,7 +461,7 @@ let iter_goal i =
Comb.get >>= fun initial ->
Proof.List.fold_left begin fun (subgoals as cur) goal ->
Solution.get >>= fun step ->
- match Evarutil.advance step goal with
+ match cleared_alias step goal with
| None -> return cur
| Some goal ->
Comb.set [goal] >>
@@ -462,7 +478,7 @@ let map_goal i =
Comb.get >>= fun initial ->
Proof.List.fold_left begin fun (acc, subgoals as cur) goal ->
Solution.get >>= fun step ->
- match Evarutil.advance step goal with
+ match cleared_alias step goal with
| None -> return cur
| Some goal ->
Comb.set [goal] >>
@@ -488,7 +504,7 @@ let fold_left2_goal i s l =
in
Proof.List.fold_left2 err begin fun ((r,subgoals) as cur) goal a ->
Solution.get >>= fun step ->
- match Evarutil.advance step goal with
+ match cleared_alias step goal with
| None -> return cur
| Some goal ->
Comb.set [goal] >>
@@ -532,7 +548,7 @@ let tclDISPATCHGEN0 join tacs =
let open Proof in
Pv.get >>= function
| { comb=[goal] ; solution } ->
- begin match Evarutil.advance solution goal with
+ begin match cleared_alias solution goal with
| None -> tclUNIT (join [])
| Some _ -> Proof.map (fun res -> join [res]) tac
end
@@ -624,12 +640,12 @@ let shelve =
Comb.get >>= fun initial ->
Comb.set [] >>
InfoL.leaf (Info.Tactic (fun () -> Pp.str"shelve")) >>
- Shelf.modify (fun gls -> gls @ initial)
+ Shelf.modify (fun gls -> gls @ CList.map drop_state initial)
let shelve_goals l =
let open Proof in
Comb.get >>= fun initial ->
- let comb = CList.filter (fun g -> not (CList.mem g l)) initial in
+ let comb = CList.filter (fun g -> not (CList.mem (drop_state g) l)) initial in
Comb.set comb >>
InfoL.leaf (Info.Tactic (fun () -> Pp.str"shelve_goals")) >>
Shelf.modify (fun gls -> gls @ l)
@@ -656,9 +672,27 @@ let free_evars sigma l =
in
List.map map l
+let free_evars_with_state sigma l =
+ let cache = Evarutil.create_undefined_evars_cache () in
+ let map ev =
+ (** Computes the set of evars appearing in the hypotheses, the conclusion or
+ the body of the evar_info [evi]. Note: since we want to use it on goals,
+ the body is actually supposed to be empty. *)
+ let ev = drop_state ev in
+ let evi = Evd.find sigma ev in
+ let fevs = lazy (Evarutil.filtered_undefined_evars_of_evar_info ~cache sigma evi) in
+ (ev, fevs)
+ in
+ List.map map l
+
+
(** [unifiable sigma g l] checks whether [g] appears in another
subgoal of [l]. The list [l] may contain [g], but it does not
affect the result. *)
+let unifiable_delayed_with_state sigma g l =
+ let g = drop_state g in
+ unifiable_delayed g l
+
let unifiable sigma g l =
let l = free_evars sigma l in
unifiable_delayed g l
@@ -668,8 +702,8 @@ let unifiable sigma g l =
whose definition other goals of [l] depend, and [n] are the
non-unifiable goals. *)
let partition_unifiable sigma l =
- let fevs = free_evars sigma l in
- CList.partition (fun g -> unifiable_delayed g fevs) l
+ let fevs = free_evars_with_state sigma l in
+ CList.partition (fun g -> unifiable_delayed_with_state sigma g fevs) l
(** Shelves the unifiable goals under focus, i.e. the goals which
appear in other goals under focus (the unfocused goals are not
@@ -680,7 +714,7 @@ let shelve_unifiable =
let (u,n) = partition_unifiable initial.solution initial.comb in
Comb.set n >>
InfoL.leaf (Info.Tactic (fun () -> Pp.str"shelve_unifiable")) >>
- Shelf.modify (fun gls -> gls @ u)
+ Shelf.modify (fun gls -> gls @ CList.map drop_state u)
(** [guard_no_unifiable] returns the list of unifiable goals if some
goals are unifiable (see {!shelve_unifiable}) in the current focus. *)
@@ -691,13 +725,14 @@ let guard_no_unifiable =
match u with
| [] -> tclUNIT None
| gls ->
- let l = CList.map (fun g -> Evd.dependent_evar_ident g initial.solution) gls in
+ let l = CList.map (fun g -> Evd.dependent_evar_ident (drop_state g) initial.solution) gls in
let l = CList.map (fun id -> Names.Name id) l in
tclUNIT (Some l)
(** [unshelve l p] adds all the goals in [l] at the end of the focused
goals of p *)
let unshelve l p =
+ let l = List.map with_empty_state l in
(* advance the goals in case of clear *)
let l = undefined p.solution l in
{ p with comb = p.comb@l }
@@ -736,7 +771,7 @@ let with_shelf tac =
let pgoal = Evd.principal_future_goal solution in
let sigma = Evd.restore_future_goals sigma fgoals pgoal in
(* Ensure we mark and return only unsolved goals *)
- let gls' = undefined sigma (CList.rev_append gls' gls) in
+ let gls' = undefined_evars sigma (CList.rev_append gls' gls) in
let sigma = CList.fold_left (mark_in_evm ~goal:false) sigma gls' in
let npv = { npv with shelf; solution = sigma } in
Pv.set npv >> tclUNIT (gls', ans)
@@ -818,7 +853,7 @@ let give_up =
Comb.set [] >>
mark_as_unsafe >>
InfoL.leaf (Info.Tactic (fun () -> Pp.str"give_up")) >>
- Giveup.put initial
+ Giveup.put (CList.map drop_state initial)
@@ -859,8 +894,8 @@ module Progress = struct
(** Equality function on goals *)
let goal_equal evars1 gl1 evars2 gl2 =
- let evi1 = Evd.find evars1 gl1 in
- let evi2 = Evd.find evars2 gl2 in
+ let evi1 = Evd.find evars1 (drop_state gl1) in
+ let evi2 = Evd.find evars2 (drop_state gl2) in
eq_evar_info evars1 evars2 evi1 evi2
end
@@ -1027,10 +1062,12 @@ module Goal = struct
env : Environ.env;
sigma : Evd.evar_map;
concl : EConstr.constr ;
+ state : StateStore.t;
self : Evar.t ; (* for compatibility with old-style definitions *)
}
let assume (gl : t) = (gl : t)
+ let state { state=state } = state
let env {env} = env
let sigma {sigma} = sigma
@@ -1038,16 +1075,19 @@ module Goal = struct
let concl {concl} = concl
let extra {sigma; self} = goal_extra sigma self
- let gmake_with info env sigma goal =
+ let gmake_with info env sigma goal state =
{ env = Environ.reset_with_named_context (Evd.evar_filtered_hyps info) env ;
sigma = sigma ;
concl = EConstr.of_constr (Evd.evar_concl info);
+ state = state ;
self = goal }
let nf_gmake env sigma goal =
+ let state = get_state goal in
+ let goal = drop_state goal in
let info = Evarutil.nf_evar_info sigma (Evd.find sigma goal) in
let sigma = Evd.add sigma goal info in
- gmake_with info env sigma goal , sigma
+ gmake_with info env sigma goal state , sigma
let nf_enter f =
InfoL.tag (Info.Dispatch) begin
@@ -1063,15 +1103,17 @@ module Goal = struct
end
end
- let normalize { self } =
+ let normalize { self; state } =
Env.get >>= fun env ->
tclEVARMAP >>= fun sigma ->
- let (gl,sigma) = nf_gmake env sigma self in
+ let (gl,sigma) = nf_gmake env sigma (goal_with_state self state) in
tclTHEN (Unsafe.tclEVARS sigma) (tclUNIT gl)
let gmake env sigma goal =
+ let state = get_state goal in
+ let goal = drop_state goal in
let info = Evd.find sigma goal in
- gmake_with info env sigma goal
+ gmake_with info env sigma goal state
let enter f =
let f gl = InfoL.tag (Info.DBranch) (f gl) in
@@ -1104,7 +1146,7 @@ module Goal = struct
Pv.get >>= fun step ->
let sigma = step.solution in
let map goal =
- match Evarutil.advance sigma goal with
+ match cleared_alias sigma goal with
| None -> None (** ppedrot: Is this check really necessary? *)
| Some goal ->
let gl =
@@ -1164,18 +1206,22 @@ module V82 = struct
let open Proof in
Pv.get >>= fun ps ->
try
- let tac gl evd =
+ let tac g_w_s evd =
+ let g, w = drop_state g_w_s, get_state g_w_s in
let glsigma =
- tac { Evd.it = gl ; sigma = evd; } in
+ tac { Evd.it = g ; sigma = evd; } in
let sigma = glsigma.Evd.sigma in
- let g = glsigma.Evd.it in
+ let g = CList.map (fun g -> goal_with_state g w) glsigma.Evd.it in
( g, sigma )
in
(* Old style tactics expect the goals normalized with respect to evars. *)
- let (initgoals,initevd) =
- Evd.Monad.List.map (fun g s -> goal_nf_evar s g) ps.comb ps.solution
+ let (initgoals_w_state, initevd) =
+ Evd.Monad.List.map (fun g_w_s s ->
+ let g, w = drop_state g_w_s, get_state g_w_s in
+ let g, s = goal_nf_evar s g in
+ goal_with_state g w, s) ps.comb ps.solution
in
- let (goalss,evd) = Evd.Monad.List.map tac initgoals initevd in
+ let (goalss,evd) = Evd.Monad.List.map tac initgoals_w_state initevd in
let sgs = CList.flatten goalss in
let sgs = undefined evd sgs in
InfoL.leaf (Info.Tactic (fun () -> Pp.str"<unknown>")) >>
@@ -1190,8 +1236,9 @@ module V82 = struct
let nf_evar_goals =
Pv.modify begin fun ps ->
let map g s = goal_nf_evar s g in
- let (goals,evd) = Evd.Monad.List.map map ps.comb ps.solution in
- { ps with solution = evd; comb = goals; }
+ let comb = CList.map drop_state ps.comb in
+ let (_goals,evd) = Evd.Monad.List.map map comb ps.solution in
+ { ps with solution = evd; }
end
let has_unresolved_evar pv =
@@ -1201,14 +1248,14 @@ module V82 = struct
let grab pv =
let undef = Evd.undefined_map pv.solution in
let goals = CList.rev_map fst (Evar.Map.bindings undef) in
- { pv with comb = goals }
+ { pv with comb = List.map with_empty_state goals }
(* Returns the open goals of the proofview together with the evar_map to
interpret them. *)
let goals { comb = comb ; solution = solution; } =
- { Evd.it = comb ; sigma = solution }
+ { Evd.it = List.map drop_state comb ; sigma = solution }
let top_goals initial { solution=solution; } =
let goals = CList.map (fun (t,_) -> fst (Constr.destEvar (EConstr.Unsafe.to_constr t))) initial in
@@ -1222,9 +1269,9 @@ module V82 = struct
let of_tactic t gls =
try
- let init = { shelf = []; solution = gls.Evd.sigma ; comb = [gls.Evd.it] } in
+ let init = { shelf = []; solution = gls.Evd.sigma ; comb = [with_empty_state gls.Evd.it] } in
let (_,final,_,_) = apply (goal_env gls.Evd.sigma gls.Evd.it) t init in
- { Evd.sigma = final.solution ; it = final.comb }
+ { Evd.sigma = final.solution ; it = CList.map drop_state final.comb }
with Logic_monad.TacticFailure e as src ->
let (_, info) = CErrors.push src in
iraise (e, info)
diff --git a/engine/proofview.mli b/engine/proofview.mli
index 721ce507d..77f30746d 100644
--- a/engine/proofview.mli
+++ b/engine/proofview.mli
@@ -72,7 +72,15 @@ val return : proofview -> Evd.evar_map
val partial_proof : entry -> proofview -> constr list
val initial_goals : entry -> (constr * types) list
+(** goal <-> goal_with_state *)
+val with_empty_state :
+ Proofview_monad.goal -> Proofview_monad.goal_with_state
+val drop_state :
+ Proofview_monad.goal_with_state -> Proofview_monad.goal
+val goal_with_state :
+ Proofview_monad.goal -> Proofview_monad.StateStore.t ->
+ Proofview_monad.goal_with_state
(** {6 Focusing commands} *)
@@ -416,14 +424,14 @@ module Unsafe : sig
(** [tclNEWGOALS gls] adds the goals [gls] to the ones currently
being proved, appending them to the list of focused goals. If a
goal is already solved, it is not added. *)
- val tclNEWGOALS : Evar.t list -> unit tactic
+ val tclNEWGOALS : Proofview_monad.goal_with_state list -> unit tactic
(** [tclSETGOALS gls] sets goals [gls] as the goals being under focus. If a
goal is already solved, it is not set. *)
- val tclSETGOALS : Evar.t list -> unit tactic
+ val tclSETGOALS : Proofview_monad.goal_with_state list -> unit tactic
(** [tclGETGOALS] returns the list of goals under focus. *)
- val tclGETGOALS : Evar.t list tactic
+ val tclGETGOALS : Proofview_monad.goal_with_state list tactic
(** Sets the evar universe context. *)
val tclEVARUNIVCONTEXT : UState.t -> unit tactic
@@ -480,6 +488,7 @@ module Goal : sig
val env : t -> Environ.env
val sigma : t -> Evd.evar_map
val extra : t -> Evd.Store.t
+ val state : t -> Proofview_monad.StateStore.t
(** [nf_enter t] applies the goal-dependent tactic [t] in each goal
independently, in the manner of {!tclINDEPENDENT} except that
diff --git a/engine/proofview_monad.ml b/engine/proofview_monad.ml
index d0f471225..494765fc4 100644
--- a/engine/proofview_monad.ml
+++ b/engine/proofview_monad.ml
@@ -149,13 +149,25 @@ module Info = struct
CList.map_append (collapse_tree n) f
end
+module StateStore = Store.Make(struct end)
+
+(* let (set_state, get_state) = StateDyn.Easy.make_dyn "goal_state" *)
+
+type goal = Evar.t
+type goal_with_state = Evar.t * StateStore.t
+
+let drop_state = fst
+let get_state = snd
+let goal_with_state g s = (g, s)
+let with_empty_state g = (g, StateStore.empty)
+let map_goal_with_state f (g, s) = (f g, s)
(** Type of proof views: current [evar_map] together with the list of
focused goals. *)
type proofview = {
solution : Evd.evar_map;
- comb : Evar.t list;
- shelf : Evar.t list;
+ comb : goal_with_state list;
+ shelf : goal list;
}
(** {6 Instantiation of the logic monad} *)
@@ -169,7 +181,7 @@ module P = struct
type e = bool
(** Status (safe/unsafe) * shelved goals * given up *)
- type w = bool * Evar.t list
+ type w = bool * goal list
let wunit = true , []
let wprod (b1, g1) (b2, g2) = b1 && b2 , g1@g2
@@ -209,9 +221,9 @@ module Solution : State with type t := Evd.evar_map = struct
let modify f = Pv.modify (fun pv -> { pv with solution = f pv.solution })
end
-module Comb : State with type t = Evar.t list = struct
+module Comb : State with type t = goal_with_state list = struct
(* spiwack: I don't know why I cannot substitute ([:=]) [t] with a type expression. *)
- type t = Evar.t list
+ type t = goal_with_state list
let get = Logical.map (fun {comb} -> comb) Pv.get
let set c = Pv.modify (fun pv -> { pv with comb = c })
let modify f = Pv.modify (fun pv -> { pv with comb = f pv.comb })
@@ -227,17 +239,17 @@ module Status : Writer with type t := bool = struct
let put s = Logical.put (s, [])
end
-module Shelf : State with type t = Evar.t list = struct
+module Shelf : State with type t = goal list = struct
(* spiwack: I don't know why I cannot substitute ([:=]) [t] with a type expression. *)
- type t = Evar.t list
+ type t = goal list
let get = Logical.map (fun {shelf} -> shelf) Pv.get
let set c = Pv.modify (fun pv -> { pv with shelf = c })
let modify f = Pv.modify (fun pv -> { pv with shelf = f pv.shelf })
end
-module Giveup : Writer with type t = Evar.t list = struct
+module Giveup : Writer with type t = goal list = struct
(* spiwack: I don't know why I cannot substitute ([:=]) [t] with a type expression. *)
- type t = Evar.t list
+ type t = goal list
let put gs = Logical.put (true, gs)
end
diff --git a/engine/proofview_monad.mli b/engine/proofview_monad.mli
index e7123218b..d26816fa6 100644
--- a/engine/proofview_monad.mli
+++ b/engine/proofview_monad.mli
@@ -67,12 +67,21 @@ module Info : sig
end
+module StateStore : Store.S
+type goal = Evar.t
+type goal_with_state
+val drop_state : goal_with_state -> goal
+val get_state : goal_with_state -> StateStore.t
+val goal_with_state : goal -> StateStore.t -> goal_with_state
+val with_empty_state : goal -> goal_with_state
+val map_goal_with_state : (goal -> goal) -> goal_with_state -> goal_with_state
+
(** Type of proof views: current [evar_map] together with the list of
focused goals. *)
type proofview = {
solution : Evd.evar_map;
- comb : Evar.t list;
- shelf : Evar.t list;
+ comb : goal_with_state list;
+ shelf : goal list;
}
(** {6 Instantiation of the logic monad} *)
@@ -81,7 +90,7 @@ module P : sig
type s = proofview * Environ.env
(** Status (safe/unsafe) * given up *)
- type w = bool * Evar.t list
+ type w = bool * goal list
val wunit : w
val wprod : w -> w -> w
@@ -118,7 +127,7 @@ module Pv : State with type t := proofview
module Solution : State with type t := Evd.evar_map
(** Lens to the list of focused goals. *)
-module Comb : State with type t = Evar.t list
+module Comb : State with type t = goal_with_state list
(** Lens to the global environment. *)
module Env : State with type t := Environ.env
@@ -128,11 +137,11 @@ module Status : Writer with type t := bool
(** Lens to the list of goals which have been shelved during the
execution of the tactic. *)
-module Shelf : State with type t = Evar.t list
+module Shelf : State with type t = goal list
(** Lens to the list of goals which were given up during the execution
of the tactic. *)
-module Giveup : Writer with type t = Evar.t list
+module Giveup : Writer with type t = goal list
(** Lens and utilies pertaining to the info trace *)
module InfoL : sig
diff --git a/engine/termops.ml b/engine/termops.ml
index 40b3d0d8b..668ae8777 100644
--- a/engine/termops.ml
+++ b/engine/termops.ml
@@ -797,9 +797,9 @@ let fold_constr_with_binders sigma g f n acc c =
each binder traversal; it is not recursive and the order with which
subterms are processed is not specified *)
-let iter_constr_with_full_binders g f l c =
+let iter_constr_with_full_binders sigma g f l c =
let open RelDecl in
- match kind c with
+ match EConstr.kind sigma c with
| (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
| Construct _) -> ()
| Cast (c,_, t) -> f l c; f l t
diff --git a/engine/termops.mli b/engine/termops.mli
index a3559a693..5aa6235f5 100644
--- a/engine/termops.mli
+++ b/engine/termops.mli
@@ -76,9 +76,10 @@ val fold_constr_with_full_binders : Evd.evar_map ->
(Context.Rel.Declaration.t -> 'a -> 'a) -> ('a -> 'b -> constr -> 'b) ->
'a -> 'b -> constr -> 'b
-val iter_constr_with_full_binders :
- (Context.Rel.Declaration.t -> 'a -> 'a) -> ('a -> Constr.constr -> unit) -> 'a ->
- Constr.constr -> unit
+val iter_constr_with_full_binders : Evd.evar_map ->
+ (rel_declaration -> 'a -> 'a) ->
+ ('a -> constr -> unit) -> 'a ->
+ constr -> unit
(**********************************************************************)
diff --git a/ide/ide_slave.ml b/ide/ide_slave.ml
index fe86df084..8f6ae9760 100644
--- a/ide/ide_slave.ml
+++ b/ide/ide_slave.ml
@@ -281,7 +281,7 @@ let pattern_of_string ?env s =
| Some e -> e
in
let constr = Pcoq.parse_string Pcoq.Constr.lconstr_pattern s in
- let (_, pat) = Constrintern.intern_constr_pattern env constr in
+ let (_, pat) = Constrintern.intern_constr_pattern env Evd.empty constr in
pat
let dirpath_of_string_list s =
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index d03aa3552..9053cdc24 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -184,14 +184,14 @@ let compute_explicitable_implicit imps = function
(* Unable to know in advance what the implicit arguments will be *)
[]
-let compute_internalization_data env ty typ impl =
- let impl = compute_implicits_with_manual env typ (is_implicit_args()) impl in
+let compute_internalization_data env sigma ty typ impl =
+ let impl = compute_implicits_with_manual env sigma typ (is_implicit_args()) impl in
let expls_impl = compute_explicitable_implicit impl ty in
- (ty, expls_impl, impl, compute_arguments_scope typ)
+ (ty, expls_impl, impl, compute_arguments_scope sigma typ)
-let compute_internalization_env env ?(impls=empty_internalization_env) ty =
+let compute_internalization_env env sigma ?(impls=empty_internalization_env) ty =
List.fold_left3
- (fun map id typ impl -> Id.Map.add id (compute_internalization_data env ty typ impl) map)
+ (fun map id typ impl -> Id.Map.add id (compute_internalization_data env sigma ty typ impl) map)
impls
(**********************************************************************)
@@ -2207,9 +2207,9 @@ let extract_ids env =
(Termops.ids_of_rel_context (Environ.rel_context env))
Id.Set.empty
-let scope_of_type_kind = function
+let scope_of_type_kind sigma = function
| IsType -> Notation.current_type_scope_name ()
- | OfType typ -> compute_type_scope (EConstr.Unsafe.to_constr typ)
+ | OfType typ -> compute_type_scope sigma typ
| WithoutTypeConstraint -> None
let empty_ltac_sign = {
@@ -2218,19 +2218,17 @@ let empty_ltac_sign = {
ltac_extra = Genintern.Store.empty;
}
-let intern_gen kind env
+let intern_gen kind env sigma
?(impls=empty_internalization_env) ?(pattern_mode=false) ?(ltacvars=empty_ltac_sign)
c =
- let tmp_scope = scope_of_type_kind kind in
+ let tmp_scope = scope_of_type_kind sigma kind in
internalize env {ids = extract_ids env; unb = false;
tmp_scope = tmp_scope; scopes = [];
impls = impls}
pattern_mode (ltacvars, Id.Map.empty) c
-let intern_constr env c = intern_gen WithoutTypeConstraint env c
-
-let intern_type env c = intern_gen IsType env c
-
+let intern_constr env sigma c = intern_gen WithoutTypeConstraint env sigma c
+let intern_type env sigma c = intern_gen IsType env sigma c
let intern_pattern globalenv patt =
try
intern_cases_pattern globalenv Id.Map.empty (None,[]) empty_alias patt
@@ -2245,7 +2243,7 @@ let intern_pattern globalenv patt =
(* All evars resolved *)
let interp_gen kind env sigma ?(impls=empty_internalization_env) c =
- let c = intern_gen kind ~impls env c in
+ let c = intern_gen kind ~impls env sigma c in
understand ~expected_type:kind env sigma c
let interp_constr env sigma ?(impls=empty_internalization_env) c =
@@ -2260,13 +2258,13 @@ let interp_casted_constr env sigma ?(impls=empty_internalization_env) c typ =
(* Not all evars expected to be resolved *)
let interp_open_constr env sigma c =
- understand_tcc env sigma (intern_constr env c)
+ understand_tcc env sigma (intern_constr env sigma c)
(* Not all evars expected to be resolved and computation of implicit args *)
let interp_constr_evars_gen_impls env sigma
?(impls=empty_internalization_env) expected_type c =
- let c = intern_gen expected_type ~impls env c in
+ let c = intern_gen expected_type ~impls env sigma c in
let imps = Implicit_quantifiers.implicits_of_glob_constr ~with_products:(expected_type == IsType) c in
let sigma, c = understand_tcc env sigma ~expected_type c in
sigma, (c, imps)
@@ -2283,7 +2281,7 @@ let interp_type_evars_impls env sigma ?(impls=empty_internalization_env) c =
(* Not all evars expected to be resolved, with side-effect on evars *)
let interp_constr_evars_gen env sigma ?(impls=empty_internalization_env) expected_type c =
- let c = intern_gen expected_type ~impls env c in
+ let c = intern_gen expected_type ~impls env sigma c in
understand_tcc env sigma ~expected_type c
let interp_constr_evars env evdref ?(impls=empty_internalization_env) c =
@@ -2297,9 +2295,9 @@ let interp_type_evars env sigma ?(impls=empty_internalization_env) c =
(* Miscellaneous *)
-let intern_constr_pattern env ?(as_type=false) ?(ltacvars=empty_ltac_sign) c =
+let intern_constr_pattern env sigma ?(as_type=false) ?(ltacvars=empty_ltac_sign) c =
let c = intern_gen (if as_type then IsType else WithoutTypeConstraint)
- ~pattern_mode:true ~ltacvars env c in
+ ~pattern_mode:true ~ltacvars env sigma c in
pattern_of_glob_constr c
let interp_notation_constr env ?(impls=empty_internalization_env) nenv a =
@@ -2323,16 +2321,14 @@ let interp_notation_constr env ?(impls=empty_internalization_env) nenv a =
(* Interpret binders and contexts *)
let interp_binder env sigma na t =
- let t = intern_gen IsType env t in
+ let t = intern_gen IsType env sigma t in
let t' = locate_if_hole ?loc:(loc_of_glob_constr t) na t in
understand ~expected_type:IsType env sigma t'
-let interp_binder_evars env evdref na t =
- let t = intern_gen IsType env t in
+let interp_binder_evars env sigma na t =
+ let t = intern_gen IsType env sigma t in
let t' = locate_if_hole ?loc:(loc_of_glob_constr t) na t in
- let evd, c = understand_tcc env !evdref ~expected_type:IsType t' in
- evdref := evd;
- c
+ understand_tcc env sigma ~expected_type:IsType t'
let my_intern_constr env lvar acc c =
internalize env acc false lvar c
diff --git a/interp/constrintern.mli b/interp/constrintern.mli
index 7411fb84b..563230b59 100644
--- a/interp/constrintern.mli
+++ b/interp/constrintern.mli
@@ -7,17 +7,17 @@
(************************************************************************)
open Names
-open Constr
open Evd
open Environ
+open Misctypes
open Libnames
open Globnames
open Glob_term
open Pattern
+open EConstr
open Constrexpr
open Notation_term
open Pretyping
-open Misctypes
(** Translation from front abstract syntax of term to untyped terms (glob_constr) *)
@@ -58,10 +58,10 @@ type internalization_env = var_internalization_data Id.Map.t
val empty_internalization_env : internalization_env
-val compute_internalization_data : env -> var_internalization_type ->
+val compute_internalization_data : env -> evar_map -> var_internalization_type ->
types -> Impargs.manual_explicitation list -> var_internalization_data
-val compute_internalization_env : env -> ?impls:internalization_env -> var_internalization_type ->
+val compute_internalization_env : env -> evar_map -> ?impls:internalization_env -> var_internalization_type ->
Id.t list -> types list -> Impargs.manual_explicitation list list ->
internalization_env
@@ -78,11 +78,10 @@ val empty_ltac_sign : ltac_sign
(** {6 Internalization performs interpretation of global names and notations } *)
-val intern_constr : env -> constr_expr -> glob_constr
-
-val intern_type : env -> constr_expr -> glob_constr
+val intern_constr : env -> evar_map -> constr_expr -> glob_constr
+val intern_type : env -> evar_map -> constr_expr -> glob_constr
-val intern_gen : typing_constraint -> env ->
+val intern_gen : typing_constraint -> env -> evar_map ->
?impls:internalization_env -> ?pattern_mode:bool -> ?ltacvars:ltac_sign ->
constr_expr -> glob_constr
@@ -100,7 +99,7 @@ val interp_constr : env -> evar_map -> ?impls:internalization_env ->
constr_expr -> constr Evd.in_evar_universe_context
val interp_casted_constr : env -> evar_map -> ?impls:internalization_env ->
- constr_expr -> EConstr.types -> constr Evd.in_evar_universe_context
+ constr_expr -> types -> constr Evd.in_evar_universe_context
val interp_type : env -> evar_map -> ?impls:internalization_env ->
constr_expr -> types Evd.in_evar_universe_context
@@ -108,37 +107,37 @@ val interp_type : env -> evar_map -> ?impls:internalization_env ->
(** Main interpretation function expecting all postponed problems to
be resolved, but possibly leaving evars. *)
-val interp_open_constr : env -> evar_map -> constr_expr -> evar_map * EConstr.constr
+val interp_open_constr : env -> evar_map -> constr_expr -> evar_map * constr
(** Accepting unresolved evars *)
val interp_constr_evars : env -> evar_map ->
- ?impls:internalization_env -> constr_expr -> evar_map * EConstr.constr
+ ?impls:internalization_env -> constr_expr -> evar_map * constr
val interp_casted_constr_evars : env -> evar_map ->
- ?impls:internalization_env -> constr_expr -> EConstr.types -> evar_map * EConstr.constr
+ ?impls:internalization_env -> constr_expr -> types -> evar_map * constr
val interp_type_evars : env -> evar_map ->
- ?impls:internalization_env -> constr_expr -> evar_map * EConstr.types
+ ?impls:internalization_env -> constr_expr -> evar_map * types
(** Accepting unresolved evars and giving back the manual implicit arguments *)
val interp_constr_evars_impls : env -> evar_map ->
?impls:internalization_env -> constr_expr ->
- evar_map * (EConstr.constr * Impargs.manual_implicits)
+ evar_map * (constr * Impargs.manual_implicits)
val interp_casted_constr_evars_impls : env -> evar_map ->
- ?impls:internalization_env -> constr_expr -> EConstr.types ->
- evar_map * (EConstr.constr * Impargs.manual_implicits)
+ ?impls:internalization_env -> constr_expr -> types ->
+ evar_map * (constr * Impargs.manual_implicits)
val interp_type_evars_impls : env -> evar_map ->
?impls:internalization_env -> constr_expr ->
- evar_map * (EConstr.types * Impargs.manual_implicits)
+ evar_map * (types * Impargs.manual_implicits)
(** Interprets constr patterns *)
val intern_constr_pattern :
- env -> ?as_type:bool -> ?ltacvars:ltac_sign ->
+ env -> evar_map -> ?as_type:bool -> ?ltacvars:ltac_sign ->
constr_pattern_expr -> patvar list * constr_pattern
(** Raise Not_found if syndef not bound to a name and error if unexisting ref *)
@@ -152,14 +151,14 @@ val interp_reference : ltac_sign -> reference -> glob_constr
val interp_binder : env -> evar_map -> Name.t -> constr_expr ->
types Evd.in_evar_universe_context
-val interp_binder_evars : env -> evar_map ref -> Name.t -> constr_expr -> EConstr.types
+val interp_binder_evars : env -> evar_map -> Name.t -> constr_expr -> evar_map * types
(** Interpret contexts: returns extended env and context *)
val interp_context_evars :
?global_level:bool -> ?impl_env:internalization_env -> ?shift:int ->
env -> evar_map -> local_binder_expr list ->
- evar_map * (internalization_env * ((env * EConstr.rel_context) * Impargs.manual_implicits))
+ evar_map * (internalization_env * ((env * rel_context) * Impargs.manual_implicits))
(* val interp_context_gen : (env -> glob_constr -> unsafe_type_judgment Evd.in_evar_universe_context) -> *)
(* (env -> Evarutil.type_constraint -> glob_constr -> unsafe_judgment Evd.in_evar_universe_context) -> *)
diff --git a/interp/declare.ml b/interp/declare.ml
index dfa84f278..f6d7b45c3 100644
--- a/interp/declare.ml
+++ b/interp/declare.ml
@@ -135,7 +135,7 @@ let set_declare_scheme f = declare_scheme := f
let update_tables c =
declare_constant_implicits c;
Heads.declare_head (EvalConstRef c);
- Notation.declare_ref_arguments_scope (ConstRef c)
+ Notation.declare_ref_arguments_scope Evd.empty (ConstRef c)
let register_side_effect (c, role) =
let o = inConstant {
@@ -275,7 +275,7 @@ let inVariable : variable_obj -> obj =
let declare_variable id obj =
let oname = add_leaf id (inVariable (Inr (id,obj))) in
declare_var_implicits id;
- Notation.declare_ref_arguments_scope (VarRef id);
+ Notation.declare_ref_arguments_scope Evd.empty (VarRef id);
Heads.declare_head (EvalVarRef id);
oname
@@ -283,9 +283,9 @@ let declare_variable id obj =
let declare_inductive_argument_scopes kn mie =
List.iteri (fun i {mind_entry_consnames=lc} ->
- Notation.declare_ref_arguments_scope (IndRef (kn,i));
+ Notation.declare_ref_arguments_scope Evd.empty (IndRef (kn,i));
for j=1 to List.length lc do
- Notation.declare_ref_arguments_scope (ConstructRef ((kn,i),j));
+ Notation.declare_ref_arguments_scope Evd.empty (ConstructRef ((kn,i),j));
done) mie.mind_entry_inds
let inductive_names sp kn mie =
diff --git a/interp/impargs.ml b/interp/impargs.ml
index ed1cd5276..6767af6d8 100644
--- a/interp/impargs.ml
+++ b/interp/impargs.ml
@@ -7,21 +7,20 @@
(************************************************************************)
open CErrors
+open Pp
open Util
open Names
-open Globnames
open Constr
-open Reduction
+open Globnames
open Declarations
-open Environ
-open Libobject
+open Decl_kinds
open Lib
-open Pp
-open Constrexpr
+open Libobject
+open EConstr
open Termops
+open Reductionops
+open Constrexpr
open Namegen
-open Decl_kinds
-open Context.Named.Declaration
module NamedDecl = Context.Named.Declaration
@@ -165,8 +164,8 @@ let update pos rig (na,st) =
in na, Some e
(* modified is_rigid_reference with a truncated env *)
-let is_flexible_reference env bound depth f =
- match kind f with
+let is_flexible_reference env sigma bound depth f =
+ match kind sigma f with
| Rel n when n >= bound+depth -> (* inductive type *) false
| Rel n when n >= depth -> (* previous argument *) true
| Rel n -> (* since local definitions have been expanded *) false
@@ -174,102 +173,101 @@ let is_flexible_reference env bound depth f =
let cb = Environ.lookup_constant kn env in
(match cb.const_body with Def _ -> true | _ -> false)
| Var id ->
- env |> Environ.lookup_named id |> is_local_def
+ env |> Environ.lookup_named id |> NamedDecl.is_local_def
| Ind _ | Construct _ -> false
| _ -> true
let push_lift d (e,n) = (push_rel d e,n+1)
-let is_reversible_pattern bound depth f l =
- isRel f && let n = destRel f in (n < bound+depth) && (n >= depth) &&
- Array.for_all (fun c -> isRel c && destRel c < depth) l &&
+let is_reversible_pattern sigma bound depth f l =
+ isRel sigma f && let n = destRel sigma f in (n < bound+depth) && (n >= depth) &&
+ Array.for_all (fun c -> isRel sigma c && destRel sigma c < depth) l &&
Array.distinct l
(* Precondition: rels in env are for inductive types only *)
-let add_free_rels_until strict strongly_strict revpat bound env m pos acc =
+let add_free_rels_until strict strongly_strict revpat bound env sigma m pos acc =
let rec frec rig (env,depth as ed) c =
- let hd = if strict then whd_all env c else c in
+ let hd = if strict then whd_all env sigma c else c in
let c = if strongly_strict then hd else c in
- match kind hd with
+ match kind sigma hd with
| Rel n when (n < bound+depth) && (n >= depth) ->
let i = bound + depth - n - 1 in
acc.(i) <- update pos rig acc.(i)
- | App (f,l) when revpat && is_reversible_pattern bound depth f l ->
- let i = bound + depth - destRel f - 1 in
+ | App (f,l) when revpat && is_reversible_pattern sigma bound depth f l ->
+ let i = bound + depth - EConstr.destRel sigma f - 1 in
acc.(i) <- update pos rig acc.(i)
- | App (f,_) when rig && is_flexible_reference env bound depth f ->
+ | App (f,_) when rig && is_flexible_reference env sigma bound depth f ->
if strict then () else
- iter_constr_with_full_binders push_lift (frec false) ed c
+ iter_constr_with_full_binders sigma push_lift (frec false) ed c
| Proj (p,c) when rig ->
if strict then () else
- iter_constr_with_full_binders push_lift (frec false) ed c
+ iter_constr_with_full_binders sigma push_lift (frec false) ed c
| Case _ when rig ->
if strict then () else
- iter_constr_with_full_binders push_lift (frec false) ed c
+ iter_constr_with_full_binders sigma push_lift (frec false) ed c
| Evar _ -> ()
| _ ->
- iter_constr_with_full_binders push_lift (frec rig) ed c
+ iter_constr_with_full_binders sigma push_lift (frec rig) ed c
in
- let () = if not (Vars.noccur_between 1 bound m) then frec true (env,1) m in
+ let () = if not (Vars.noccur_between sigma 1 bound m) then frec true (env,1) m in
acc
-let rec is_rigid_head t = match kind t with
+let rec is_rigid_head sigma t = match kind sigma t with
| Rel _ | Evar _ -> false
| Ind _ | Const _ | Var _ | Sort _ -> true
- | Case (_,_,f,_) -> is_rigid_head f
+ | Case (_,_,f,_) -> is_rigid_head sigma f
| Proj (p,c) -> true
| App (f,args) ->
- (match kind f with
- | Fix ((fi,i),_) -> is_rigid_head (args.(fi.(i)))
- | _ -> is_rigid_head f)
+ (match kind sigma f with
+ | Fix ((fi,i),_) -> is_rigid_head sigma (args.(fi.(i)))
+ | _ -> is_rigid_head sigma f)
| Lambda _ | LetIn _ | Construct _ | CoFix _ | Fix _
| Prod _ | Meta _ | Cast _ -> assert false
(* calcule la liste des arguments implicites *)
let find_displayed_name_in all avoid na (env, b) =
- let b = EConstr.of_constr b in
let envnames_b = (env, b) in
let flag = RenamingElsewhereFor envnames_b in
if all then compute_and_force_displayed_name_in Evd.empty flag avoid na b
else compute_displayed_name_in Evd.empty flag avoid na b
-let compute_implicits_gen strict strongly_strict revpat contextual all env t =
+let compute_implicits_gen strict strongly_strict revpat contextual all env sigma (t : EConstr.t) =
let rigid = ref true in
let open Context.Rel.Declaration in
- let rec aux env avoid n names t =
- let t = whd_all env t in
- match kind t with
+ let rec aux env avoid n names (t : EConstr.t) =
+ let t = whd_all env sigma t in
+ match kind sigma t with
| Prod (na,a,b) ->
let na',avoid' = find_displayed_name_in all avoid na (names,b) in
- add_free_rels_until strict strongly_strict revpat n env a (Hyp (n+1))
+ add_free_rels_until strict strongly_strict revpat n env sigma a (Hyp (n+1))
(aux (push_rel (LocalAssum (na',a)) env) avoid' (n+1) (na'::names) b)
| _ ->
- rigid := is_rigid_head t;
+ rigid := is_rigid_head sigma t;
let names = List.rev names in
let v = Array.map (fun na -> na,None) (Array.of_list names) in
if contextual then
- add_free_rels_until strict strongly_strict revpat n env t Conclusion v
+ add_free_rels_until strict strongly_strict revpat n env sigma t Conclusion v
else v
in
- match kind (whd_all env t) with
+ match kind sigma (whd_all env sigma t) with
| Prod (na,a,b) ->
let na',avoid = find_displayed_name_in all Id.Set.empty na ([],b) in
let v = aux (push_rel (LocalAssum (na',a)) env) avoid 1 [na'] b in
!rigid, Array.to_list v
| _ -> true, []
-let compute_implicits_flags env f all t =
+let compute_implicits_flags env sigma f all t =
compute_implicits_gen
(f.strict || f.strongly_strict) f.strongly_strict
- f.reversible_pattern f.contextual all env t
+ f.reversible_pattern f.contextual all env sigma t
-let compute_auto_implicits env flags enriching t =
- if enriching then compute_implicits_flags env flags true t
- else compute_implicits_gen false false false true true env t
+let compute_auto_implicits env sigma flags enriching t =
+ if enriching then compute_implicits_flags env sigma flags true t
+ else compute_implicits_gen false false false true true env sigma t
-let compute_implicits_names env t =
- let _, impls = compute_implicits_gen false false false false true env t in
+let compute_implicits_names env sigma t =
+ let _, impls = compute_implicits_gen false false false false true env sigma t in
List.map fst impls
(* Extra information about implicit arguments *)
@@ -398,24 +396,25 @@ let set_manual_implicits env flags enriching autoimps l =
in
merge 1 l autoimps
-let compute_semi_auto_implicits env f manual t =
+let compute_semi_auto_implicits env sigma f manual t =
match manual with
| [] ->
if not f.auto then [DefaultImpArgs, []]
- else let _,l = compute_implicits_flags env f false t in
+ else let _,l = compute_implicits_flags env sigma f false t in
[DefaultImpArgs, prepare_implicits f l]
| _ ->
- let _,autoimpls = compute_auto_implicits env f f.auto t in
+ let _,autoimpls = compute_auto_implicits env sigma f f.auto t in
[DefaultImpArgs, set_manual_implicits env f f.auto autoimpls manual]
(*s Constants. *)
let compute_constant_implicits flags manual cst =
let env = Global.env () in
+ let sigma = Evd.from_env env in
let cb = Environ.lookup_constant cst env in
- let ty = cb.const_type in
- let impls = compute_semi_auto_implicits env flags manual ty in
- impls
+ let ty = of_constr cb.const_type in
+ let impls = compute_semi_auto_implicits env sigma flags manual ty in
+ impls
(*s Inductives and constructors. Their implicit arguments are stored
in an array, indexed by the inductive number, of pairs $(i,v)$ where
@@ -424,7 +423,8 @@ let compute_constant_implicits flags manual cst =
let compute_mib_implicits flags manual kn =
let env = Global.env () in
- let mib = lookup_mind kn env in
+ let sigma = Evd.from_env env in
+ let mib = Environ.lookup_mind kn env in
let ar =
Array.to_list
(Array.mapi (* No need to lift, arities contain no de Bruijn *)
@@ -433,14 +433,14 @@ let compute_mib_implicits flags manual kn =
let ty, _ = Global.type_of_global_in_context env (IndRef (kn,i)) in
Context.Rel.Declaration.LocalAssum (Name mip.mind_typename, ty))
mib.mind_packets) in
- let env_ar = push_rel_context ar env in
+ let env_ar = Environ.push_rel_context ar env in
let imps_one_inductive i mip =
let ind = (kn,i) in
let ar, _ = Global.type_of_global_in_context env (IndRef ind) in
- ((IndRef ind,compute_semi_auto_implicits env flags manual ar),
+ ((IndRef ind,compute_semi_auto_implicits env sigma flags manual (of_constr ar)),
Array.mapi (fun j c ->
- (ConstructRef (ind,j+1),compute_semi_auto_implicits env_ar flags manual c))
- mip.mind_nf_lc)
+ (ConstructRef (ind,j+1),compute_semi_auto_implicits env_ar sigma flags manual c))
+ (Array.map of_constr mip.mind_nf_lc))
in
Array.mapi imps_one_inductive mib.mind_packets
@@ -453,7 +453,8 @@ let compute_all_mib_implicits flags manual kn =
let compute_var_implicits flags manual id =
let env = Global.env () in
- compute_semi_auto_implicits env flags manual (NamedDecl.get_type (lookup_named id env))
+ let sigma = Evd.from_env env in
+ compute_semi_auto_implicits env sigma flags manual (NamedDecl.get_type (lookup_named id env))
(* Implicits of a global reference. *)
@@ -524,7 +525,7 @@ let impls_of_context ctx =
| Implicit -> Some (NamedDecl.get_id decl, Manual, (true, true))
| _ -> None
in
- List.rev_map map (List.filter (fst %> is_local_assum) ctx)
+ List.rev_map map (List.filter (fst %> NamedDecl.is_local_assum) ctx)
let adjust_side_condition p = function
| LessArgsThan n -> LessArgsThan (n+p)
@@ -649,8 +650,8 @@ type manual_explicitation = Constrexpr.explicitation * (bool * bool * bool)
type manual_implicits = manual_explicitation list
-let compute_implicits_with_manual env typ enriching l =
- let _,autoimpls = compute_auto_implicits env !implicit_args enriching typ in
+let compute_implicits_with_manual env sigma typ enriching l =
+ let _,autoimpls = compute_auto_implicits env sigma !implicit_args enriching typ in
set_manual_implicits env !implicit_args enriching autoimpls l
let check_inclusion l =
@@ -674,9 +675,10 @@ let projection_implicits env p impls =
let declare_manual_implicits local ref ?enriching l =
let flags = !implicit_args in
let env = Global.env () in
- let t, _ = Global.type_of_global_in_context (Global.env ()) ref in
+ let sigma = Evd.from_env env in
+ let t, _ = Global.type_of_global_in_context env ref in
let enriching = Option.default flags.auto enriching in
- let isrigid,autoimpls = compute_auto_implicits env flags enriching t in
+ let isrigid,autoimpls = compute_auto_implicits env sigma flags enriching (of_constr t) in
let l' = match l with
| [] -> assert false
| [l] ->
diff --git a/interp/impargs.mli b/interp/impargs.mli
index 40fa4cb26..dcfe527b0 100644
--- a/interp/impargs.mli
+++ b/interp/impargs.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Constr
+open EConstr
open Globnames
open Environ
@@ -90,10 +90,10 @@ type manual_explicitation = Constrexpr.explicitation *
type manual_implicits = manual_explicitation list
-val compute_implicits_with_manual : env -> types -> bool ->
+val compute_implicits_with_manual : env -> Evd.evar_map -> types -> bool ->
manual_implicits -> implicit_status list
-val compute_implicits_names : env -> types -> Name.t list
+val compute_implicits_names : env -> Evd.evar_map -> types -> Name.t list
(** {6 Computation of implicits (done using the global environment). } *)
diff --git a/interp/modintern.ml b/interp/modintern.ml
index e631b3ea4..128152cc2 100644
--- a/interp/modintern.ml
+++ b/interp/modintern.ml
@@ -61,13 +61,16 @@ let transl_with_decl env = function
| CWith_Module ((_,fqid),qid) ->
WithMod (fqid,lookup_module qid), Univ.ContextSet.empty
| CWith_Definition ((_,fqid),c) ->
- let c, ectx = interp_constr env (Evd.from_env env) c in
+ let sigma = Evd.from_env env in
+ let c, ectx = interp_constr env sigma c in
if Flags.is_universe_polymorphism () then
let ctx = UState.context ectx in
let inst, ctx = Univ.abstract_universes ctx in
- let c = Vars.subst_univs_level_constr (Univ.make_instance_subst inst) c in
+ let c = EConstr.Vars.subst_univs_level_constr (Univ.make_instance_subst inst) c in
+ let c = EConstr.to_constr sigma c in
WithDef (fqid,(c, Some ctx)), Univ.ContextSet.empty
else
+ let c = EConstr.to_constr sigma c in
WithDef (fqid,(c, None)), UState.context_set ectx
let loc_of_module l = l.CAst.loc
diff --git a/interp/notation.ml b/interp/notation.ml
index ea7ef21b1..bf39c726a 100644
--- a/interp/notation.ml
+++ b/interp/notation.ml
@@ -668,8 +668,8 @@ type scope_class = cl_typ
let scope_class_compare : scope_class -> scope_class -> int =
cl_typ_ord
-let compute_scope_class t =
- let (cl,_,_) = find_class_type Evd.empty (EConstr.of_constr t) in
+let compute_scope_class sigma t =
+ let (cl,_,_) = find_class_type sigma t in
cl
module ScopeClassOrd =
@@ -698,22 +698,22 @@ let find_scope_class_opt = function
(**********************************************************************)
(* Special scopes associated to arguments of a global reference *)
-let rec compute_arguments_classes t =
- match Constr.kind (EConstr.Unsafe.to_constr (Reductionops.whd_betaiotazeta Evd.empty (EConstr.of_constr t))) with
+let rec compute_arguments_classes sigma t =
+ match EConstr.kind sigma (Reductionops.whd_betaiotazeta sigma t) with
| Prod (_,t,u) ->
- let cl = try Some (compute_scope_class t) with Not_found -> None in
- cl :: compute_arguments_classes u
+ let cl = try Some (compute_scope_class sigma t) with Not_found -> None in
+ cl :: compute_arguments_classes sigma u
| _ -> []
-let compute_arguments_scope_full t =
- let cls = compute_arguments_classes t in
+let compute_arguments_scope_full sigma t =
+ let cls = compute_arguments_classes sigma t in
let scs = List.map find_scope_class_opt cls in
scs, cls
-let compute_arguments_scope t = fst (compute_arguments_scope_full t)
+let compute_arguments_scope sigma t = fst (compute_arguments_scope_full sigma t)
-let compute_type_scope t =
- find_scope_class_opt (try Some (compute_scope_class t) with Not_found -> None)
+let compute_type_scope sigma t =
+ find_scope_class_opt (try Some (compute_scope_class sigma t) with Not_found -> None)
let current_type_scope_name () =
find_scope_class_opt (Some CL_SORT)
@@ -779,20 +779,24 @@ let discharge_arguments_scope (_,(req,r,n,l,_)) =
let classify_arguments_scope (req,_,_,_,_ as obj) =
if req == ArgsScopeNoDischarge then Dispose else Substitute obj
-let rebuild_arguments_scope (req,r,n,l,_) =
+let rebuild_arguments_scope sigma (req,r,n,l,_) =
match req with
| ArgsScopeNoDischarge -> assert false
| ArgsScopeAuto ->
- let scs,cls = compute_arguments_scope_full (fst(Global.type_of_global_in_context (Global.env ()) r)(*FIXME?*)) in
- (req,r,List.length scs,scs,cls)
+ let env = Global.env () in (*FIXME?*)
+ let typ = EConstr.of_constr @@ fst (Global.type_of_global_in_context env r) in
+ let scs,cls = compute_arguments_scope_full sigma typ in
+ (req,r,List.length scs,scs,cls)
| ArgsScopeManual ->
- (* Add to the manually given scopes the one found automatically
- for the extra parameters of the section. Discard the classes
- of the manually given scopes to avoid further re-computations. *)
- let l',cls = compute_arguments_scope_full (fst (Global.type_of_global_in_context (Global.env ()) r)) in
- let l1 = List.firstn n l' in
- let cls1 = List.firstn n cls in
- (req,r,0,l1@l,cls1)
+ (* Add to the manually given scopes the one found automatically
+ for the extra parameters of the section. Discard the classes
+ of the manually given scopes to avoid further re-computations. *)
+ let env = Global.env () in (*FIXME?*)
+ let typ = EConstr.of_constr @@ fst (Global.type_of_global_in_context env r) in
+ let l',cls = compute_arguments_scope_full sigma typ in
+ let l1 = List.firstn n l' in
+ let cls1 = List.firstn n cls in
+ (req,r,0,l1@l,cls1)
type arguments_scope_obj =
arguments_scope_discharge_request * global_reference *
@@ -807,7 +811,8 @@ let inArgumentsScope : arguments_scope_obj -> obj =
subst_function = subst_arguments_scope;
classify_function = classify_arguments_scope;
discharge_function = discharge_arguments_scope;
- rebuild_function = rebuild_arguments_scope }
+ (* XXX: Should we pass the sigma here or not, see @herbelin's comment in 6511 *)
+ rebuild_function = rebuild_arguments_scope Evd.empty }
let is_local local ref = local || isVarRef ref && Lib.is_in_section ref
@@ -819,7 +824,7 @@ let declare_arguments_scope local r scl =
(* We empty the list of argument classes to disable further scope
re-computations and keep these manually given scopes. *)
declare_arguments_scope_gen req r 0 (scl,[])
-
+
let find_arguments_scope r =
try
let (scl,cls,stamp) = Refmap.find r !arguments_scope in
@@ -832,12 +837,12 @@ let find_arguments_scope r =
scl'
with Not_found -> []
-let declare_ref_arguments_scope ref =
- let t, _ = Global.type_of_global_in_context (Global.env ()) ref in
- let (scs,cls as o) = compute_arguments_scope_full t in
+let declare_ref_arguments_scope sigma ref =
+ let env = Global.env () in (* FIXME? *)
+ let typ = EConstr.of_constr @@ fst @@ Global.type_of_global_in_context env ref in
+ let (scs,cls as o) = compute_arguments_scope_full sigma typ in
declare_arguments_scope_gen ArgsScopeAuto ref (List.length scs) o
-
(********************************)
(* Encoding notations as string *)
diff --git a/interp/notation.mli b/interp/notation.mli
index a4c79d6d3..79b56bd41 100644
--- a/interp/notation.mli
+++ b/interp/notation.mli
@@ -163,10 +163,10 @@ val subst_scope_class :
Mod_subst.substitution -> scope_class -> scope_class option
val declare_scope_class : scope_name -> scope_class -> unit
-val declare_ref_arguments_scope : global_reference -> unit
+val declare_ref_arguments_scope : Evd.evar_map -> global_reference -> unit
-val compute_arguments_scope : Constr.types -> scope_name option list
-val compute_type_scope : Constr.types -> scope_name option
+val compute_arguments_scope : Evd.evar_map -> EConstr.types -> scope_name option list
+val compute_type_scope : Evd.evar_map -> EConstr.types -> scope_name option
(** Get the current scope bound to Sortclass, if it exists *)
val current_type_scope_name : unit -> scope_name option
diff --git a/kernel/cbytecodes.ml b/kernel/cbytecodes.ml
index aa6c49bc7..586ef1709 100644
--- a/kernel/cbytecodes.ml
+++ b/kernel/cbytecodes.ml
@@ -32,13 +32,12 @@ let cofix_evaluated_tag = 7
let last_variant_tag = 245
type structured_constant =
- | Const_sorts of Sorts.t
+ | Const_sort of Sorts.t
| Const_ind of inductive
| Const_proj of Constant.t
| Const_b0 of tag
| Const_bn of tag * structured_constant array
| Const_univ_level of Univ.Level.t
- | Const_type of Univ.Universe.t
type reloc_table = (tag * int) array
@@ -46,8 +45,8 @@ type annot_switch =
{ci : case_info; rtbl : reloc_table; tailcall : bool; max_stack_size : int}
let rec eq_structured_constant c1 c2 = match c1, c2 with
-| Const_sorts s1, Const_sorts s2 -> Sorts.equal s1 s2
-| Const_sorts _, _ -> false
+| Const_sort s1, Const_sort s2 -> Sorts.equal s1 s2
+| Const_sort _, _ -> false
| Const_ind i1, Const_ind i2 -> eq_ind i1 i2
| Const_ind _, _ -> false
| Const_proj p1, Const_proj p2 -> Constant.equal p1 p2
@@ -59,13 +58,11 @@ let rec eq_structured_constant c1 c2 = match c1, c2 with
| Const_bn _, _ -> false
| Const_univ_level l1 , Const_univ_level l2 -> Univ.Level.equal l1 l2
| Const_univ_level _ , _ -> false
-| Const_type u1 , Const_type u2 -> Univ.Universe.equal u1 u2
-| Const_type _ , _ -> false
let rec hash_structured_constant c =
let open Hashset.Combine in
match c with
- | Const_sorts s -> combinesmall 1 (Sorts.hash s)
+ | Const_sort s -> combinesmall 1 (Sorts.hash s)
| Const_ind i -> combinesmall 2 (ind_hash i)
| Const_proj p -> combinesmall 3 (Constant.hash p)
| Const_b0 t -> combinesmall 4 (Int.hash t)
@@ -74,7 +71,6 @@ let rec hash_structured_constant c =
let h = Array.fold_left fold 0 a in
combinesmall 5 (combine (Int.hash t) h)
| Const_univ_level l -> combinesmall 6 (Univ.Level.hash l)
- | Const_type u -> combinesmall 7 (Univ.Universe.hash u)
let eq_annot_switch asw1 asw2 =
let eq_ci ci1 ci2 =
@@ -236,20 +232,19 @@ open Util
let pp_sort s =
let open Sorts in
- match family s with
- | InSet -> str "Set"
- | InProp -> str "Prop"
- | InType -> str "Type"
+ match s with
+ | Prop Null -> str "Prop"
+ | Prop Pos -> str "Set"
+ | Type u -> str "Type@{" ++ Univ.pr_uni u ++ str "}"
let rec pp_struct_const = function
- | Const_sorts s -> pp_sort s
+ | Const_sort s -> pp_sort s
| Const_ind (mind, i) -> MutInd.print mind ++ str"#" ++ int i
| Const_proj p -> Constant.print p
| Const_b0 i -> int i
| Const_bn (i,t) ->
int i ++ surround (prvect_with_sep pr_comma pp_struct_const t)
| Const_univ_level l -> Univ.Level.pr l
- | Const_type u -> str "Type@{" ++ Univ.pr_uni u ++ str "}"
let pp_lbl lbl = str "L" ++ int lbl
diff --git a/kernel/cbytecodes.mli b/kernel/cbytecodes.mli
index c8fbb27a9..71dd65186 100644
--- a/kernel/cbytecodes.mli
+++ b/kernel/cbytecodes.mli
@@ -26,13 +26,12 @@ val cofix_evaluated_tag : tag
val last_variant_tag : tag
type structured_constant =
- | Const_sorts of Sorts.t
+ | Const_sort of Sorts.t
| Const_ind of inductive
| Const_proj of Constant.t
| Const_b0 of tag
| Const_bn of tag * structured_constant array
| Const_univ_level of Univ.Level.t
- | Const_type of Univ.Universe.t
val pp_struct_const : structured_constant -> Pp.t
diff --git a/kernel/cbytegen.ml b/kernel/cbytegen.ml
index 3104d5751..0d7619e9f 100644
--- a/kernel/cbytegen.ml
+++ b/kernel/cbytegen.ml
@@ -480,23 +480,23 @@ let rec compile_lam env reloc lam sz cont =
(Const_ind ind) (Univ.Instance.to_array u) sz cont
| Lsort (Sorts.Prop _ as s) ->
- compile_structured_constant reloc (Const_sorts s) sz cont
+ compile_structured_constant reloc (Const_sort s) sz cont
| Lsort (Sorts.Type u) ->
- (* We separate global and local universes in [u]. The former will be part
- of the structured constant, while the later (if any) will be applied as
- arguments. *)
- let open Univ in begin
+ (* We separate global and local universes in [u]. The former will be part
+ of the structured constant, while the later (if any) will be applied as
+ arguments. *)
+ let open Univ in begin
let u,s = Universe.compact u in
(* We assume that [Universe.type0m] is a neutral element for [Universe.sup] *)
+ let compile_get_univ reloc idx sz cont =
+ set_max_stack_size sz;
+ compile_fv_elem reloc (FVuniv_var idx) sz cont
+ in
if List.is_empty s then
- compile_structured_constant reloc (Const_sorts (Sorts.Type u)) sz cont
+ compile_structured_constant reloc (Const_sort (Sorts.Type u)) sz cont
else
- let compile_get_univ reloc idx sz cont =
- set_max_stack_size sz;
- compile_fv_elem reloc (FVuniv_var idx) sz cont
- in
comp_app compile_structured_constant compile_get_univ reloc
- (Const_type u) (Array.of_list s) sz cont
+ (Const_sort (Sorts.Type u)) (Array.of_list s) sz cont
end
| Llet (id,def,body) ->
diff --git a/kernel/cemitcodes.ml b/kernel/cemitcodes.ml
index 856b0b465..32d5387b2 100644
--- a/kernel/cemitcodes.ml
+++ b/kernel/cemitcodes.ml
@@ -350,7 +350,7 @@ type to_patch = emitcodes * patches * fv
(* Substitution *)
let rec subst_strcst s sc =
match sc with
- | Const_sorts _ | Const_b0 _ | Const_univ_level _ | Const_type _ -> sc
+ | Const_sort _ | Const_b0 _ | Const_univ_level _ -> sc
| Const_proj p -> Const_proj (subst_constant s p)
| Const_bn(tag,args) -> Const_bn(tag,Array.map (subst_strcst s) args)
| Const_ind ind -> let kn,i = ind in Const_ind (subst_mind s kn, i)
diff --git a/kernel/uGraph.ml b/kernel/uGraph.ml
index f1e8d1031..5dca69c16 100644
--- a/kernel/uGraph.ml
+++ b/kernel/uGraph.ml
@@ -833,7 +833,7 @@ let sort_universes g =
(** Subtyping of polymorphic contexts *)
let check_subtype univs ctxT ctx =
- if AUContext.size ctx == AUContext.size ctx then
+ if AUContext.size ctxT == AUContext.size ctx then
let (inst, cst) = UContext.dest (AUContext.repr ctx) in
let cstT = UContext.constraints (AUContext.repr ctxT) in
let push accu v = add_universe v false accu in
diff --git a/kernel/vconv.ml b/kernel/vconv.ml
index 8c7658147..ad9aa4267 100644
--- a/kernel/vconv.ml
+++ b/kernel/vconv.ml
@@ -44,7 +44,6 @@ let rec conv_val env pb k v1 v2 cu =
and conv_whd env pb k whd1 whd2 cu =
(* Pp.(msg_debug (str "conv_whd(" ++ pr_whd whd1 ++ str ", " ++ pr_whd whd2 ++ str ")")) ; *)
match whd1, whd2 with
- | Vsort s1, Vsort s2 -> sort_cmp_universes env pb s1 s2 cu
| Vuniv_level _ , _
| _ , Vuniv_level _ ->
(** Both of these are invalid since universes are handled via
@@ -81,7 +80,7 @@ and conv_whd env pb k whd1 whd2 cu =
(* on the fly eta expansion *)
conv_val env CONV (k+1) (apply_whd k whd1) (apply_whd k whd2) cu
- | Vsort _, _ | Vprod _, _ | Vfix _, _ | Vcofix _, _ | Vconstr_const _, _
+ | Vprod _, _ | Vfix _, _ | Vcofix _, _ | Vconstr_const _, _
| Vconstr_block _, _ | Vatom_stk _, _ -> raise NotConvertible
@@ -119,8 +118,9 @@ and conv_atom env pb k a1 stk1 a2 stk2 cu =
if Vars.eq_id_key ik1 ik2 && compare_stack stk1 stk2 then
conv_stack env k stk1 stk2 cu
else raise NotConvertible
- | Atype _ , _ | _, Atype _ -> assert false
- | Aind _, _ | Aid _, _ -> raise NotConvertible
+ | Asort s1, Asort s2 ->
+ sort_cmp_universes env pb s1 s2 cu
+ | Asort _ , _ | Aind _, _ | Aid _, _ -> raise NotConvertible
and conv_stack env k stk1 stk2 cu =
match stk1, stk2 with
diff --git a/kernel/vm.ml b/kernel/vm.ml
index 352ea74a4..f0bae98dc 100644
--- a/kernel/vm.ml
+++ b/kernel/vm.ml
@@ -168,7 +168,7 @@ let rec apply_stack a stk v =
let apply_whd k whd =
let v = val_of_rel k in
match whd with
- | Vsort _ | Vprod _ | Vconstr_const _ | Vconstr_block _ -> assert false
+ | Vprod _ | Vconstr_const _ | Vconstr_block _ -> assert false
| Vfun f -> reduce_fun k f
| Vfix(f, None) ->
push_ra stop;
diff --git a/kernel/vmvalues.ml b/kernel/vmvalues.ml
index 2d8a1d976..2a784fdf4 100644
--- a/kernel/vmvalues.ml
+++ b/kernel/vmvalues.ml
@@ -121,7 +121,7 @@ type vswitch = {
type atom =
| Aid of Vars.id_key
| Aind of inductive
- | Atype of Univ.Universe.t
+ | Asort of Sorts.t
(* Zippers *)
@@ -136,7 +136,6 @@ type stack = zipper list
type to_update = values
type whd =
- | Vsort of Sorts.t
| Vprod of vprod
| Vfun of vfun
| Vfix of vfix * arguments option
@@ -167,7 +166,6 @@ let uni_lvl_val (v : values) : Univ.Level.t =
let pr =
let open Pp in
match whd with
- | Vsort _ -> str "Vsort"
| Vprod _ -> str "Vprod"
| Vfun _ -> str "Vfun"
| Vfix _ -> str "Vfix"
@@ -189,12 +187,17 @@ let rec whd_accu a stk =
match Obj.tag at with
| i when Int.equal i type_atom_tag ->
begin match stk with
+ | [] -> Vatom_stk(Obj.magic at, stk)
| [Zapp args] ->
let args = Array.init (nargs args) (arg args) in
- let u = Obj.obj (Obj.field at 0) in
- let inst = Instance.of_array (Array.map uni_lvl_val args) in
- let u = Univ.subst_instance_universe inst u in
- Vsort (Type u)
+ let s = Obj.obj (Obj.field at 0) in
+ begin match s with
+ | Type u ->
+ let inst = Instance.of_array (Array.map uni_lvl_val args) in
+ let u = Univ.subst_instance_universe inst u in
+ Vatom_stk (Asort (Type u), [])
+ | _ -> assert false
+ end
| _ -> assert false
end
| i when i <= max_atom_tag ->
@@ -243,11 +246,8 @@ let whd_val : values -> whd =
else
let tag = Obj.tag o in
if tag = accu_tag then
- (
- if Int.equal (Obj.size o) 1 then Obj.obj o (* sort *)
- else
- if is_accumulate (fun_code o) then whd_accu o []
- else Vprod(Obj.obj o))
+ if is_accumulate (fun_code o) then whd_accu o []
+ else Vprod(Obj.obj o)
else
if tag = Obj.closure_tag || tag = Obj.infix_tag then
(match kind_of_closure o with
@@ -273,7 +273,7 @@ let obj_of_atom : atom -> Obj.t =
(* obj_of_str_const : structured_constant -> Obj.t *)
let rec obj_of_str_const str =
match str with
- | Const_sorts s -> Obj.repr (Vsort s)
+ | Const_sort s -> obj_of_atom (Asort s)
| Const_ind ind -> obj_of_atom (Aind ind)
| Const_proj p -> Obj.repr p
| Const_b0 tag -> Obj.repr tag
@@ -285,7 +285,6 @@ let rec obj_of_str_const str =
done;
res
| Const_univ_level l -> Obj.repr (Vuniv_level l)
- | Const_type u -> obj_of_atom (Atype u)
let val_of_obj o = ((Obj.obj o) : values)
@@ -502,10 +501,9 @@ let rec pr_atom a =
| RelKey i -> str "#" ++ int i
| _ -> str "...") ++ str ")"
| Aind (mi,i) -> str "Aind(" ++ MutInd.print mi ++ str "#" ++ int i ++ str ")"
- | Atype _ -> str "Atype(")
+ | Asort _ -> str "Asort(")
and pr_whd w =
Pp.(match w with
- | Vsort _ -> str "Vsort"
| Vprod _ -> str "Vprod"
| Vfun _ -> str "Vfun"
| Vfix _ -> str "Vfix"
diff --git a/kernel/vmvalues.mli b/kernel/vmvalues.mli
index 350f71372..570e3606a 100644
--- a/kernel/vmvalues.mli
+++ b/kernel/vmvalues.mli
@@ -57,7 +57,7 @@ val cofix_upd_code : to_update -> tcode
type atom =
| Aid of Vars.id_key
| Aind of inductive
- | Atype of Univ.Universe.t
+ | Asort of Sorts.t
(** Zippers *)
@@ -70,7 +70,6 @@ type zipper =
type stack = zipper list
type whd =
- | Vsort of Sorts.t
| Vprod of vprod
| Vfun of vfun
| Vfix of vfix * arguments option
diff --git a/parsing/cLexer.ml4 b/parsing/cLexer.ml4
index 52a6fe16c..d7e3751fd 100644
--- a/parsing/cLexer.ml4
+++ b/parsing/cLexer.ml4
@@ -384,16 +384,6 @@ let comments = ref []
let current_comment = Buffer.create 8192
let between_commands = ref true
-let rec split_comments comacc acc pos = function
- [] -> comments := List.rev acc; comacc
- | ((b,e),c as com)::coms ->
- (* Take all comments that terminates before pos, or begin exactly
- at pos (used to print comments attached after an expression) *)
- if e<=pos || pos=b then split_comments (c::comacc) acc pos coms
- else split_comments comacc (com::acc) pos coms
-
-let extract_comments pos = split_comments [] [] pos !comments
-
(* The state of the lexer visible from outside *)
type lexer_state = int option * string * bool * ((int * int) * string) list * Loc.source
@@ -410,6 +400,8 @@ let release_lexer_state = get_lexer_state
let drop_lexer_state () =
set_lexer_state (init_lexer_state Loc.ToplevelInput)
+let get_comment_state (_,_,_,c,_) = c
+
let real_push_char c = Buffer.add_char current_comment c
(* Add a char if it is between two commands, if it is a newline or
diff --git a/parsing/cLexer.mli b/parsing/cLexer.mli
index 5f4e10f14..2d35571f1 100644
--- a/parsing/cLexer.mli
+++ b/parsing/cLexer.mli
@@ -55,7 +55,4 @@ val get_lexer_state : unit -> lexer_state
val release_lexer_state : unit -> lexer_state
[@@ocaml.deprecated "Use get_lexer_state"]
val drop_lexer_state : unit -> unit
-
-(* Retrieve the comments lexed at a given location of the stream
- currently being processeed *)
-val extract_comments : int -> string list
+val get_comment_state : lexer_state -> ((int * int) * string) list
diff --git a/parsing/pcoq.ml b/parsing/pcoq.ml
index 7a51908d9..fec26f941 100644
--- a/parsing/pcoq.ml
+++ b/parsing/pcoq.ml
@@ -88,7 +88,9 @@ module type S =
val entry_create : string -> 'a entry
val entry_parse : 'a entry -> coq_parsable -> 'a
val entry_print : Format.formatter -> 'a entry -> unit
- val with_parsable : coq_parsable -> ('a -> 'b) -> 'a -> 'b
+
+ val comment_state : coq_parsable -> ((int * int) * string) list
+
val srules' : production_rule list -> symbol
val parse_tokens_after_filter : 'a entry -> Tok.t Stream.t -> 'a
@@ -105,6 +107,7 @@ end with type 'a Entry.e = 'a Grammar.GMake(CLexer).Entry.e = struct
string option * Gramext.g_assoc option * production_rule list
type extend_statment =
Gramext.position option * single_extend_statment list
+
type coq_parsable = parsable * CLexer.lexer_state ref
let parsable ?(file=Loc.ToplevelInput) c =
@@ -129,15 +132,8 @@ end with type 'a Entry.e = 'a Grammar.GMake(CLexer).Entry.e = struct
let loc = match loc' with None -> to_coqloc loc | Some loc -> loc in
Loc.raise ~loc e
- let with_parsable (p,state) f x =
- CLexer.set_lexer_state !state;
- try
- let a = f x in
- state := CLexer.get_lexer_state ();
- a
- with e ->
- CLexer.drop_lexer_state ();
- raise e
+ let comment_state (p,state) =
+ CLexer.get_comment_state !state
let entry_print ft x = Entry.print ft x
diff --git a/parsing/pcoq.mli b/parsing/pcoq.mli
index f36250176..1b330aa04 100644
--- a/parsing/pcoq.mli
+++ b/parsing/pcoq.mli
@@ -78,7 +78,9 @@ module type S =
val entry_create : string -> 'a entry
val entry_parse : 'a entry -> coq_parsable -> 'a
val entry_print : Format.formatter -> 'a entry -> unit
- val with_parsable : coq_parsable -> ('a -> 'b) -> 'a -> 'b
+
+ (* Get comment parsing information from the Lexer *)
+ val comment_state : coq_parsable -> ((int * int) * string) list
(* Apparently not used *)
val srules' : production_rule list -> symbol
diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml
index 22881c32c..7159614d9 100644
--- a/plugins/funind/glob_term_to_relation.ml
+++ b/plugins/funind/glob_term_to_relation.ml
@@ -352,9 +352,9 @@ let raw_push_named (na,raw_value,raw_typ) env =
let typ,_ = Pretyping.understand env (Evd.from_env env) ~expected_type:Pretyping.IsType raw_typ in
(match raw_value with
| None ->
- Environ.push_named (NamedDecl.LocalAssum (id,typ)) env
+ EConstr.push_named (NamedDecl.LocalAssum (id,typ)) env
| Some value ->
- Environ.push_named (NamedDecl.LocalDef (id, value, typ)) env)
+ EConstr.push_named (NamedDecl.LocalDef (id, value, typ)) env)
let add_pat_variables pat typ env : Environ.env =
@@ -519,7 +519,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
The "value" of this branch is then simply [res]
*)
let rt_as_constr,ctx = Pretyping.understand env (Evd.from_env env) rt in
- let rt_typ = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr rt_as_constr) in
+ let rt_typ = Typing.unsafe_type_of env (Evd.from_env env) rt_as_constr in
let res_raw_type = Detyping.detype Detyping.Now false Id.Set.empty env (Evd.from_env env) rt_typ in
let res = fresh_id args_res.to_avoid "_res" in
let new_avoid = res::args_res.to_avoid in
@@ -631,12 +631,11 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
let v = match typ with None -> v | Some t -> DAst.make ?loc:rt.loc @@ GCast (v,CastConv t) in
let v_res = build_entry_lc env funnames avoid v in
let v_as_constr,ctx = Pretyping.understand env (Evd.from_env env) v in
- let v_type = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr v_as_constr) in
- let v_type = EConstr.Unsafe.to_constr v_type in
+ let v_type = Typing.unsafe_type_of env (Evd.from_env env) v_as_constr in
let new_env =
match n with
Anonymous -> env
- | Name id -> Environ.push_named (NamedDecl.LocalDef (id,v_as_constr,v_type)) env
+ | Name id -> EConstr.push_named (NamedDecl.LocalDef (id,v_as_constr,v_type)) env
in
let b_res = build_entry_lc new_env funnames avoid b in
combine_results (combine_letin n) v_res b_res
@@ -648,7 +647,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
build_entry_lc_from_case env funnames make_discr el brl avoid
| GIf(b,(na,e_option),lhs,rhs) ->
let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in
- let b_typ = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr b_as_constr) in
+ let b_typ = Typing.unsafe_type_of env (Evd.from_env env) b_as_constr in
let (ind,_) =
try Inductiveops.find_inductive env (Evd.from_env env) b_typ
with Not_found ->
@@ -680,7 +679,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
nal
in
let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in
- let b_typ = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr b_as_constr) in
+ let b_typ = Typing.unsafe_type_of env (Evd.from_env env) b_as_constr in
let (ind,_) =
try Inductiveops.find_inductive env (Evd.from_env env) b_typ
with Not_found ->
@@ -726,7 +725,7 @@ and build_entry_lc_from_case env funname make_discr
let types =
List.map (fun (case_arg,_) ->
let case_arg_as_constr,ctx = Pretyping.understand env (Evd.from_env env) case_arg in
- EConstr.Unsafe.to_constr (Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr case_arg_as_constr))
+ EConstr.Unsafe.to_constr (Typing.unsafe_type_of env (Evd.from_env env) case_arg_as_constr)
) el
in
(****** The next works only if the match is not dependent ****)
@@ -948,7 +947,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
mkGApp(mkGVar(mk_rel_id this_relname),List.tl args'@[res_rt])
in
let t',ctx = Pretyping.understand env (Evd.from_env env) new_t in
- let new_env = Environ.push_rel (LocalAssum (n,t')) env in
+ let new_env = EConstr.push_rel (LocalAssum (n,t')) env in
let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
@@ -983,7 +982,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let subst_b =
if is_in_b then b else replace_var_by_term id rt b
in
- let new_env = Environ.push_rel (LocalAssum (n,t')) env in
+ let new_env = EConstr.push_rel (LocalAssum (n,t')) env in
let new_b,id_to_exclude =
rebuild_cons
new_env
@@ -995,7 +994,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
with Continue ->
let jmeq = Globnames.IndRef (fst (EConstr.destInd Evd.empty (jmeq ()))) in
let ty',ctx = Pretyping.understand env (Evd.from_env env) ty in
- let ind,args' = Inductive.find_inductive env ty' in
+ let ind,args' = Inductiveops.find_inductive env Evd.(from_env env) ty' in
let mib,_ = Global.lookup_inductive (fst ind) in
let nparam = mib.Declarations.mind_nparams in
let params,arg' =
@@ -1017,14 +1016,14 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
observe (str "computing new type for jmeq : " ++ pr_glob_constr_env env eq');
let eq'_as_constr,ctx = Pretyping.understand env (Evd.from_env env) eq' in
observe (str " computing new type for jmeq : done") ;
+ let sigma = Evd.(from_env env) in
let new_args =
- match Constr.kind eq'_as_constr with
+ match EConstr.kind sigma eq'_as_constr with
| App(_,[|_;_;ty;_|]) ->
- let ty = Array.to_list (snd (destApp ty)) in
+ let ty = Array.to_list (snd (EConstr.destApp sigma ty)) in
let ty' = snd (Util.List.chop nparam ty) in
List.fold_left2
(fun acc var_as_constr arg ->
- let arg = EConstr.of_constr arg in
if isRel var_as_constr
then
let na = RelDecl.get_name (Environ.lookup_rel (destRel var_as_constr) env) in
@@ -1065,7 +1064,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
in
let new_env =
let t',ctx = Pretyping.understand env (Evd.from_env env) eq' in
- Environ.push_rel (LocalAssum (n,t')) env
+ EConstr.push_rel (LocalAssum (n,t')) env
in
let new_b,id_to_exclude =
rebuild_cons
@@ -1103,7 +1102,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
with Continue ->
observe (str "computing new type for prod : " ++ pr_glob_constr_env env rt);
let t',ctx = Pretyping.understand env (Evd.from_env env) t in
- let new_env = Environ.push_rel (LocalAssum (n,t')) env in
+ let new_env = EConstr.push_rel (LocalAssum (n,t')) env in
let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
@@ -1119,7 +1118,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
| _ ->
observe (str "computing new type for prod : " ++ pr_glob_constr_env env rt);
let t',ctx = Pretyping.understand env (Evd.from_env env) t in
- let new_env = Environ.push_rel (LocalAssum (n,t')) env in
+ let new_env = EConstr.push_rel (LocalAssum (n,t')) env in
let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
@@ -1140,7 +1139,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let t',ctx = Pretyping.understand env (Evd.from_env env) t in
match n with
| Name id ->
- let new_env = Environ.push_rel (LocalAssum (n,t')) env in
+ let new_env = EConstr.push_rel (LocalAssum (n,t')) env in
let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
@@ -1163,7 +1162,8 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let evd = (Evd.from_env env) in
let t',ctx = Pretyping.understand env evd t in
let evd = Evd.from_ctx ctx in
- let type_t' = Typing.unsafe_type_of env evd (EConstr.of_constr t') in
+ let type_t' = Typing.unsafe_type_of env evd t' in
+ let t' = EConstr.Unsafe.to_constr t' in
let type_t' = EConstr.Unsafe.to_constr type_t' in
let new_env = Environ.push_rel (LocalDef (n,t',type_t')) env in
let new_b,id_to_exclude =
@@ -1189,7 +1189,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
depth t
in
let t',ctx = Pretyping.understand env (Evd.from_env env) new_t in
- let new_env = Environ.push_rel (LocalAssum (na,t')) env in
+ let new_env = EConstr.push_rel (LocalAssum (na,t')) env in
let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
@@ -1369,8 +1369,9 @@ let do_build_inductive
*)
let rel_arities = Array.mapi rel_arity funsargs in
Util.Array.fold_left2 (fun env rel_name rel_ar ->
- Environ.push_named (LocalAssum (rel_name,
- fst (with_full_print (Constrintern.interp_constr env evd) rel_ar))) env) env relnames rel_arities
+ let rex = fst (with_full_print (Constrintern.interp_constr env evd) rel_ar) in
+ let rex = EConstr.Unsafe.to_constr rex in
+ Environ.push_named (LocalAssum (rel_name,rex)) env) env relnames rel_arities
in
(* and of the real constructors*)
let constr i res =
diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml
index e19fc9b62..13eda3952 100644
--- a/plugins/funind/indfun.ml
+++ b/plugins/funind/indfun.ml
@@ -141,8 +141,7 @@ let rec abstract_glob_constr c = function
| Constrexpr.CLocalPattern _::bl -> assert false
let interp_casted_constr_with_implicits env sigma impls c =
- Constrintern.intern_gen Pretyping.WithoutTypeConstraint env ~impls
- c
+ Constrintern.intern_gen Pretyping.WithoutTypeConstraint env sigma ~impls c
(*
Construct a fixpoint as a Glob_term
@@ -160,9 +159,9 @@ let build_newrecursive
let arity,ctx = Constrintern.interp_type env0 sigma arityc in
let evd = Evd.from_env env0 in
let evd, (_, (_, impls')) = Constrintern.interp_context_evars env evd bl in
- let impl = Constrintern.compute_internalization_data env0 Constrintern.Recursive arity impls' in
+ let impl = Constrintern.compute_internalization_data env0 evd Constrintern.Recursive arity impls' in
let open Context.Named.Declaration in
- (Environ.push_named (LocalAssum (recname,arity)) env, Id.Map.add recname impl impls))
+ (EConstr.push_named (LocalAssum (recname,arity)) env, Id.Map.add recname impl impls))
(env0,Constrintern.empty_internalization_env) lnameargsardef in
let recdef =
(* Declare local notations *)
diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml
index 8fe05b497..e2d7de6d5 100644
--- a/plugins/funind/recdef.ml
+++ b/plugins/funind/recdef.ml
@@ -210,6 +210,7 @@ let (value_f: Constr.t list -> global_reference -> Constr.t) =
DAst.make @@ GVar v_id)])
in
let body = fst (understand env (Evd.from_env env) glob_body)(*FIXME*) in
+ let body = EConstr.Unsafe.to_constr body in
it_mkLambda_or_LetIn body context
let (declare_f : Id.t -> logical_kind -> Constr.t list -> global_reference -> global_reference) =
@@ -1400,7 +1401,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp
(fun c ->
Proofview.V82.of_tactic (Tacticals.New.tclTHENLIST
[intros;
- Simple.apply (EConstr.of_constr (fst (interp_constr (Global.env()) Evd.empty c))) (*FIXME*);
+ Simple.apply (fst (interp_constr (Global.env()) Evd.empty c)) (*FIXME*);
Tacticals.New.tclCOMPLETE Auto.default_auto
])
)
@@ -1599,7 +1600,9 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
and functional_ref = destConst (constr_of_global functional_ref)
and eq_ref = destConst (constr_of_global eq_ref) in
generate_induction_principle f_ref tcc_lemma_constr
- functional_ref eq_ref rec_arg_num (EConstr.of_constr rec_arg_type) (nb_prod evd (EConstr.of_constr res)) (EConstr.of_constr relation);
+ functional_ref eq_ref rec_arg_num
+ (EConstr.of_constr rec_arg_type)
+ (nb_prod evd (EConstr.of_constr res)) relation;
Flags.if_verbose
msgnl (h 1 (Ppconstr.pr_id function_name ++
spc () ++ str"is defined" )++ fnl () ++
@@ -1614,7 +1617,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
tcc_lemma_constr
is_mes functional_ref
(EConstr.of_constr rec_arg_type)
- (EConstr.of_constr relation) rec_arg_num
+ relation rec_arg_num
term_id
using_lemmas
(List.length res_vars)
diff --git a/plugins/ltac/coretactics.ml4 b/plugins/ltac/coretactics.ml4
index 794a28dd4..ad6c6df5f 100644
--- a/plugins/ltac/coretactics.ml4
+++ b/plugins/ltac/coretactics.ml4
@@ -239,12 +239,20 @@ END
(** Simple induction / destruct *)
+let simple_induct h =
+ Tacticals.New.tclTHEN (Tactics.intros_until h)
+ (Tacticals.New.onLastHyp Tactics.simplest_elim)
+
TACTIC EXTEND simple_induction
- [ "simple" "induction" quantified_hypothesis(h) ] -> [ Tactics.simple_induct h ]
+ [ "simple" "induction" quantified_hypothesis(h) ] -> [ simple_induct h ]
END
+let simple_destruct h =
+ Tacticals.New.tclTHEN (Tactics.intros_until h)
+ (Tacticals.New.onLastHyp Tactics.simplest_case)
+
TACTIC EXTEND simple_destruct
- [ "simple" "destruct" quantified_hypothesis(h) ] -> [ Tactics.simple_destruct h ]
+ [ "simple" "destruct" quantified_hypothesis(h) ] -> [ simple_destruct h ]
END
(** Double induction *)
diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4
index 10be8a842..6c42518b1 100644
--- a/plugins/ltac/extratactics.ml4
+++ b/plugins/ltac/extratactics.ml4
@@ -253,6 +253,7 @@ let add_rewrite_hint ~poly bases ort t lcsr =
let sigma = Evd.from_env env in
let f ce =
let c, ctx = Constrintern.interp_constr env sigma ce in
+ let c = EConstr.to_constr sigma c in
let ctx =
let ctx = UState.context_set ctx in
if poly then ctx
@@ -554,6 +555,7 @@ let add_transitivity_lemma left lem =
let env = Global.env () in
let sigma = Evd.from_env env in
let lem',ctx (*FIXME*) = Constrintern.interp_constr env sigma lem in
+ let lem' = EConstr.to_constr sigma lem' in
add_anonymous_leaf (inTransitivity (left,lem'))
(* Vernacular syntax *)
@@ -611,8 +613,10 @@ END
VERNAC COMMAND EXTEND RetroknowledgeRegister CLASSIFIED AS SIDEFF
| [ "Register" constr(c) "as" retroknowledge_field(f) "by" constr(b)] ->
- [ let tc,ctx = Constrintern.interp_constr (Global.env ()) Evd.empty c in
- let tb,ctx(*FIXME*) = Constrintern.interp_constr (Global.env ()) Evd.empty b in
+ [ let tc,_ctx = Constrintern.interp_constr (Global.env ()) Evd.empty c in
+ let tb,_ctx(*FIXME*) = Constrintern.interp_constr (Global.env ()) Evd.empty b in
+ let tc = EConstr.to_constr Evd.empty tc in
+ let tb = EConstr.to_constr Evd.empty tb in
Global.register f tc tb ]
END
@@ -705,7 +709,6 @@ let hResolve id c occ t =
resolve_hole (subst_hole_with_term loc_begin c_raw t_hole)
in
let t_constr,ctx = resolve_hole (subst_var_with_hole occ id t_raw) in
- let t_constr = EConstr.of_constr t_constr in
let sigma = Evd.merge_universe_context sigma ctx in
let t_constr_type = Retyping.get_type_of env sigma t_constr in
Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma)
@@ -973,6 +976,7 @@ TACTIC EXTEND unshelve
| [ "unshelve" tactic1(t) ] ->
[
Proofview.with_shelf (Tacinterp.tactic_of_value ist t) >>= fun (gls, ()) ->
+ let gls = List.map Proofview.with_empty_state gls in
Proofview.Unsafe.tclGETGOALS >>= fun ogls ->
Proofview.Unsafe.tclSETGOALS (gls @ ogls)
]
diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml
index e73a18b79..1f2189015 100644
--- a/plugins/ltac/rewrite.ml
+++ b/plugins/ltac/rewrite.ml
@@ -1568,7 +1568,8 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause =
let (undef, prf, newt) = res in
let fold ev _ accu = if Evd.mem sigma ev then accu else ev :: accu in
let gls = List.rev (Evd.fold_undefined fold undef []) in
- match clause, prf with
+ let gls = List.map Proofview.with_empty_state gls in
+ match clause, prf with
| Some id, Some p ->
let tac = tclTHENLIST [
Refine.refine ~typecheck:true (fun h -> (h,p));
@@ -1896,7 +1897,6 @@ let declare_projection n instance_id r =
let build_morphism_signature env sigma m =
let m,ctx = Constrintern.interp_constr env sigma m in
- let m = EConstr.of_constr m in
let sigma = Evd.from_ctx ctx in
let t = Typing.unsafe_type_of env sigma m in
let cstrs =
diff --git a/plugins/ltac/tacintern.ml b/plugins/ltac/tacintern.ml
index 22ec6c5b1..9a10fb805 100644
--- a/plugins/ltac/tacintern.ml
+++ b/plugins/ltac/tacintern.ml
@@ -199,7 +199,7 @@ let intern_constr_gen pattern_mode isarity {ltacvars=lfun; genv=env; extra} c =
ltac_extra = extra;
} in
let c' =
- warn (Constrintern.intern_gen scope ~pattern_mode ~ltacvars env) c
+ warn (Constrintern.intern_gen scope ~pattern_mode ~ltacvars env Evd.(from_env env)) c
in
(c',if !strict_check then None else Some c)
@@ -316,7 +316,7 @@ let intern_constr_pattern ist ~as_type ~ltacvars pc =
ltac_extra = ist.extra;
} in
let metas,pat = Constrintern.intern_constr_pattern
- ist.genv ~as_type ~ltacvars pc
+ ist.genv Evd.(from_env ist.genv) ~as_type ~ltacvars pc
in
let (glob,_ as c) = intern_constr_gen true false ist pc in
let bound_names = Glob_ops.bound_glob_vars glob in
@@ -335,7 +335,7 @@ let intern_typed_pattern ist ~as_type ~ltacvars p =
ltac_bound = Id.Set.empty;
ltac_extra = ist.extra;
} in
- Constrintern.intern_constr_pattern ist.genv ~as_type ~ltacvars p
+ Constrintern.intern_constr_pattern ist.genv Evd.(from_env ist.genv) ~as_type ~ltacvars p
else
[], dummy_pat in
let (glob,_ as c) = intern_constr_gen true false ist p in
diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml
index 79b5c1622..6a04badf3 100644
--- a/plugins/ltac/tacinterp.ml
+++ b/plugins/ltac/tacinterp.ml
@@ -584,7 +584,7 @@ let interp_glob_closure ist env sigma ?(kind=WithoutTypeConstraint) ?(pattern_mo
ltac_bound = Id.Map.domain ist.lfun;
ltac_extra = Genintern.Store.empty;
} in
- { closure ; term = intern_gen kind ~pattern_mode ~ltacvars env term_expr }
+ { closure ; term = intern_gen kind ~pattern_mode ~ltacvars env sigma term_expr }
let interp_uconstr ist env sigma c = interp_glob_closure ist env sigma c
diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml
index 125afb1a0..0f0dbfff3 100644
--- a/plugins/setoid_ring/newring.ml
+++ b/plugins/setoid_ring/newring.ml
@@ -148,7 +148,7 @@ let ic c =
let ic_unsafe c = (*FIXME remove *)
let env = Global.env() in
let sigma = Evd.from_env env in
- EConstr.of_constr (fst (Constrintern.interp_constr env sigma c))
+ fst (Constrintern.interp_constr env sigma c)
let decl_constant na univs c =
let open Constr in
diff --git a/plugins/ssr/ssrcommon.ml b/plugins/ssr/ssrcommon.ml
index 4ae746288..1a02fb91c 100644
--- a/plugins/ssr/ssrcommon.ml
+++ b/plugins/ssr/ssrcommon.ml
@@ -203,7 +203,7 @@ let glob_constr ist genv = function
let vars = Id.Map.fold (fun x _ accu -> Id.Set.add x accu) ist.Tacinterp.lfun Id.Set.empty in
let ltacvars = {
Constrintern.empty_ltac_sign with Constrintern.ltac_vars = vars } in
- Constrintern.intern_gen Pretyping.WithoutTypeConstraint ~ltacvars genv ce
+ Constrintern.intern_gen Pretyping.WithoutTypeConstraint ~ltacvars genv Evd.(from_env genv) ce
| rc, None -> rc
let pf_intern_term ist gl (_, c) = glob_constr ist (pf_env gl) c
diff --git a/plugins/ssr/ssrvernac.ml4 b/plugins/ssr/ssrvernac.ml4
index 8e6e0347f..96ded5abf 100644
--- a/plugins/ssr/ssrvernac.ml4
+++ b/plugins/ssr/ssrvernac.ml4
@@ -335,7 +335,8 @@ let push_rels_assum l e =
push_rels_assum l e
let coerce_search_pattern_to_sort hpat =
- let env = Global.env () and sigma = Evd.empty in
+ let env = Global.env () in
+ let sigma = Evd.(from_env env) in
let mkPApp fp n_imps args =
let args' = Array.append (Array.make n_imps (Pattern.PMeta None)) args in
Pattern.PApp (fp, args') in
@@ -393,8 +394,9 @@ let interp_search_arg arg =
interp_search_notation ~loc s key
| RGlobSearchSubPattern p ->
try
- let intern = Constrintern.intern_constr_pattern in
- Search.GlobSearchSubPattern (snd (intern (Global.env()) p))
+ let env = Global.env () in
+ let _, p = Constrintern.intern_constr_pattern env (Evd.from_env env) p in
+ Search.GlobSearchSubPattern p
with e -> let e = CErrors.push e in iraise (ExplainErr.process_vernac_interp_error e)) arg in
let hpat, a1 = match arg with
| (_, Search.GlobSearchSubPattern (Pattern.PMeta _)) :: a' -> all_true, a'
diff --git a/plugins/ssr/ssrview.ml b/plugins/ssr/ssrview.ml
index 61b65e347..f2990070b 100644
--- a/plugins/ssr/ssrview.ml
+++ b/plugins/ssr/ssrview.ml
@@ -55,7 +55,8 @@ let in_viewhint =
Libobject.classify_function = classify_viewhint }
let glob_view_hints lvh =
- List.map (Constrintern.intern_constr (Global.env ())) lvh
+ let env = Global.env () in
+ List.map (Constrintern.intern_constr env Evd.(from_env env)) lvh
let add_view_hints lvh i = Lib.add_anonymous_leaf (in_viewhint (i, lvh))
diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4
index 73e212365..d05f50e6d 100644
--- a/plugins/ssrmatching/ssrmatching.ml4
+++ b/plugins/ssrmatching/ssrmatching.ml4
@@ -73,11 +73,11 @@ let env_size env = List.length (Environ.named_context env)
let safeDestApp c =
match kind c with App (f, a) -> f, a | _ -> c, [| |]
(* Toplevel constr must be globalized twice ! *)
-let glob_constr ist genv = function
+let glob_constr ist genv sigma = function
| _, Some ce ->
let vars = Id.Map.fold (fun x _ accu -> Id.Set.add x accu) ist.lfun Id.Set.empty in
let ltacvars = { Constrintern.empty_ltac_sign with Constrintern.ltac_vars = vars } in
- Constrintern.intern_gen WithoutTypeConstraint ~ltacvars:ltacvars genv ce
+ Constrintern.intern_gen WithoutTypeConstraint ~ltacvars:ltacvars genv sigma ce
| rc, None -> rc
(* Term printing utilities functions for deciding bracketing. *)
@@ -1009,7 +1009,7 @@ let interp_wit wit ist gl x =
sigma, Value.cast (topwit wit) arg
let interp_open_constr ist gl gc =
interp_wit wit_open_constr ist gl gc
-let pf_intern_term ist gl (_, c) = glob_constr ist (pf_env gl) c
+let pf_intern_term ist gl (_, c) = glob_constr ist (pf_env gl) gl.sigma c
let interp_term ist gl (_, c) = on_snd EConstr.Unsafe.to_constr (interp_open_constr ist gl c)
let pr_ssrterm _ _ _ = pr_term
let input_ssrtermkind strm = match stream_nth 0 strm with
diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml
index 275a079d5..33d674ff5 100644
--- a/pretyping/inductiveops.ml
+++ b/pretyping/inductiveops.ml
@@ -643,8 +643,9 @@ let type_of_projection_knowing_arg env sigma p c ty =
(* A function which checks that a term well typed verifies both
syntactic conditions *)
-let control_only_guard env c =
- let check_fix_cofix e c = match kind c with
+let control_only_guard env sigma c =
+ let check_fix_cofix e c =
+ match kind (EConstr.to_constr sigma c) with
| CoFix (_,(_,_,_) as cofix) ->
Inductive.check_cofix e cofix
| Fix (_,(_,_,_) as fix) ->
@@ -653,6 +654,6 @@ let control_only_guard env c =
in
let rec iter env c =
check_fix_cofix env c;
- iter_constr_with_full_binders push_rel iter env c
+ iter_constr_with_full_binders sigma EConstr.push_rel iter env c
in
iter env c
diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli
index 55149552a..8efa087a7 100644
--- a/pretyping/inductiveops.mli
+++ b/pretyping/inductiveops.mli
@@ -198,4 +198,4 @@ val type_of_inductive_knowing_conclusion :
env -> evar_map -> Inductive.mind_specif Univ.puniverses -> EConstr.types -> evar_map * EConstr.types
(********************)
-val control_only_guard : env -> types -> unit
+val control_only_guard : env -> Evd.evar_map -> EConstr.types -> unit
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index bfc04893d..4363fc7cc 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -1205,8 +1205,7 @@ let all_no_fail_flags = default_inference_flags false
let ise_pretype_gen_ctx flags env sigma lvar kind c =
let evd, c, _ = ise_pretype_gen flags env sigma lvar kind c in
- let evd, f = Evarutil.nf_evars_and_universes evd in
- f (EConstr.Unsafe.to_constr c), Evd.evar_universe_context evd
+ c, Evd.evar_universe_context evd
(** Entry points of the high-level type synthesis algorithm *)
diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli
index 864768fe5..ee568fbc4 100644
--- a/pretyping/pretyping.mli
+++ b/pretyping/pretyping.mli
@@ -12,7 +12,6 @@
into elementary ones, insertion of coercions and resolution of
implicit arguments. *)
-open Constr
open Environ
open Evd
open EConstr
@@ -26,7 +25,7 @@ val interp_known_glob_level : ?loc:Loc.t -> Evd.evar_map ->
(** An auxiliary function for searching for fixpoint guard indexes *)
val search_guard :
- ?loc:Loc.t -> env -> int list list -> rec_declaration -> int array
+ ?loc:Loc.t -> env -> int list list -> Constr.rec_declaration -> int array
type typing_constraint = OfType of types | IsType | WithoutTypeConstraint
@@ -85,9 +84,8 @@ val understand_ltac : inference_flags ->
heuristics (but no external tactic solver hook), as well as to
ensure that conversion problems are all solved and that no
unresolved evar remains, expanding evars. *)
-
val understand : ?flags:inference_flags -> ?expected_type:typing_constraint ->
- env -> evar_map -> glob_constr -> Constr.constr Evd.in_evar_universe_context
+ env -> evar_map -> glob_constr -> constr Evd.in_evar_universe_context
(** Trying to solve remaining evars and remaining conversion problems
possibly using type classes, heuristics, external tactic solver
diff --git a/pretyping/vnorm.ml b/pretyping/vnorm.ml
index c93b41786..6b7ea2996 100644
--- a/pretyping/vnorm.ml
+++ b/pretyping/vnorm.ml
@@ -141,7 +141,6 @@ and nf_vtype env sigma v = nf_val env sigma v crazy_type
and nf_whd env sigma whd typ =
match whd with
- | Vsort s -> mkSort s
| Vprod p ->
let dom = nf_vtype env sigma (dom p) in
let name = Name (Id.of_string "x") in
@@ -182,7 +181,8 @@ and nf_whd env sigma whd typ =
let pind = (ind, u) in (mkIndU pind, type_of_ind env pind)
in
nf_univ_args ~nb_univs mk env sigma stk
- | Vatom_stk(Atype u, stk) -> assert false
+ | Vatom_stk(Asort s, stk) ->
+ assert (List.is_empty stk); mkSort s
| Vuniv_level lvl ->
assert false
diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml
index 3c7095505..da1fe6d3d 100644
--- a/printing/ppconstr.ml
+++ b/printing/ppconstr.ml
@@ -149,7 +149,7 @@ let tag_var = tag Tag.variable
str "`" ++ str hd ++ c ++ str tl
let pr_com_at n =
- if !Flags.beautify && not (Int.equal n 0) then comment (CLexer.extract_comments n)
+ if !Flags.beautify && not (Int.equal n 0) then comment (Pputils.extract_comments n)
else mt()
let pr_with_comments ?loc pp = pr_located (fun x -> x) (loc, pp)
@@ -701,13 +701,16 @@ let tag_var = tag Tag.variable
| { CAst.v = CAppExpl ((None,f,us),[]) } -> str "@" ++ pr_cref f us
| c -> pr prec c
- let transf env c =
+ let transf env sigma c =
if !Flags.beautify_file then
- let r = Constrintern.for_grammar (Constrintern.intern_constr env) c in
+ let r = Constrintern.for_grammar (Constrintern.intern_constr env sigma) c in
Constrextern.extern_glob_constr (Termops.vars_of_env env) r
else c
- let pr_expr prec c = pr prec (transf (Global.env()) c)
+ let pr_expr prec c =
+ let env = Global.env () in
+ let sigma = Evd.from_env env in
+ pr prec (transf env sigma c)
let pr_simpleconstr = pr_expr lsimpleconstr
diff --git a/printing/pputils.ml b/printing/pputils.ml
index e779fc5fc..3d72c8da3 100644
--- a/printing/pputils.ml
+++ b/printing/pputils.ml
@@ -13,14 +13,26 @@ open Misctypes
open Locus
open Genredexpr
+let beautify_comments = ref []
+
+let rec split_comments comacc acc pos = function
+ | [] -> beautify_comments := List.rev acc; comacc
+ | ((b,e),c as com)::coms ->
+ (* Take all comments that terminates before pos, or begin exactly
+ at pos (used to print comments attached after an expression) *)
+ if e<=pos || pos=b then split_comments (c::comacc) acc pos coms
+ else split_comments comacc (com::acc) pos coms
+
+let extract_comments pos = split_comments [] [] pos !beautify_comments
+
let pr_located pr (loc, x) =
match loc with
| Some loc when !Flags.beautify ->
let (b, e) = Loc.unloc loc in
(* Side-effect: order matters *)
- let before = Pp.comment (CLexer.extract_comments b) in
+ let before = Pp.comment (extract_comments b) in
let x = pr x in
- let after = Pp.comment (CLexer.extract_comments e) in
+ let after = Pp.comment (extract_comments e) in
before ++ x ++ after
| _ -> pr x
diff --git a/printing/pputils.mli b/printing/pputils.mli
index ec5c32fc4..483bd0ae3 100644
--- a/printing/pputils.mli
+++ b/printing/pputils.mli
@@ -37,3 +37,9 @@ val pr_red_expr_env : Environ.env -> Evd.evar_map ->
val pr_raw_generic : Environ.env -> rlevel generic_argument -> Pp.t
val pr_glb_generic : Environ.env -> glevel generic_argument -> Pp.t
+
+(* The comments interface is imperative due to the printer not
+ threading it, this could be solved using a better data
+ structure. *)
+val beautify_comments : ((int * int) * string) list ref
+val extract_comments : int -> string list
diff --git a/proofs/proof.ml b/proofs/proof.ml
index 04e707cd6..f5ff35eeb 100644
--- a/proofs/proof.ml
+++ b/proofs/proof.ml
@@ -432,7 +432,7 @@ module V82 = struct
CList.nth evl (n-1)
in
let env = Evd.evar_filtered_env evi in
- let rawc = Constrintern.intern_constr env com in
+ let rawc = Constrintern.intern_constr env sigma com in
let ltac_vars = Glob_ops.empty_lvar in
let sigma = Evar_refiner.w_refine (evk, evi) (ltac_vars, rawc) sigma in
Proofview.Unsafe.tclEVARS sigma
diff --git a/proofs/refine.ml b/proofs/refine.ml
index 90276951b..73ed70a1a 100644
--- a/proofs/refine.ml
+++ b/proofs/refine.ml
@@ -73,6 +73,7 @@ let generic_refine ~typecheck f gl =
let sigma = Proofview.Goal.sigma gl in
let env = Proofview.Goal.env gl in
let concl = Proofview.Goal.concl gl in
+ let state = Proofview.Goal.state gl in
(** Save the [future_goals] state to restore them after the
refinement. *)
let prev_future_goals = Evd.future_goals sigma in
@@ -120,6 +121,7 @@ let generic_refine ~typecheck f gl =
(** Select the goals *)
let comb = CList.map_filter (Proofview.Unsafe.advance sigma) (CList.rev evs) in
let sigma = CList.fold_left Proofview.Unsafe.mark_as_goal sigma comb in
+ let comb = CList.map (fun x -> Proofview.goal_with_state x state) comb in
let trace () = Pp.(hov 2 (str"simple refine"++spc()++ Hook.get pr_constrv env sigma c)) in
Proofview.Trace.name_tactic trace (Proofview.tclUNIT v) >>= fun v ->
Proofview.Unsafe.tclSETENV (Environ.reset_context env) <*>
diff --git a/stm/stm.ml b/stm/stm.ml
index e7c371798..fd22a8388 100644
--- a/stm/stm.ml
+++ b/stm/stm.ml
@@ -2949,6 +2949,7 @@ let get_ast ~doc id =
match VCS.visit id with
| { step = `Cmd { cast = { loc; expr } } }
| { step = `Fork (({ loc; expr }, _, _, _), _) }
+ | { step = `Sideff ((ReplayCommand {loc; expr}) , _) }
| { step = `Qed ({ qast = { loc; expr } }, _) } ->
Some (Loc.tag ?loc expr)
| _ -> None
diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml
index a95e6b941..b8860d3a5 100644
--- a/tactics/class_tactics.ml
+++ b/tactics/class_tactics.ml
@@ -581,7 +581,7 @@ let make_resolve_hyp env sigma st flags only_classes pri decl =
(fun (path,info,c) ->
let info =
{ info with Vernacexpr.hint_pattern =
- Option.map (Constrintern.intern_constr_pattern env)
+ Option.map (Constrintern.intern_constr_pattern env sigma)
info.Vernacexpr.hint_pattern }
in
make_resolves env sigma ~name:(PathHints path)
@@ -1131,6 +1131,7 @@ module Search = struct
let rec result (shelf, ()) i k =
foundone := true;
Proofview.Unsafe.tclGETGOALS >>= fun gls ->
+ let gls = CList.map Proofview.drop_state gls in
let j = List.length gls in
(if !typeclasses_debug > 0 then
Feedback.msg_debug
@@ -1179,7 +1180,7 @@ module Search = struct
(if List.is_empty goals then tclUNIT ()
else
let sigma' = mark_unresolvables sigma goals in
- with_shelf (Unsafe.tclEVARS sigma' <*> Unsafe.tclNEWGOALS goals) >>=
+ with_shelf (Unsafe.tclEVARS sigma' <*> Unsafe.tclNEWGOALS (CList.map Proofview.with_empty_state goals)) >>=
fun s -> result s i (Some (Option.default 0 k + j)))
end
in with_shelf res >>= fun (sh, ()) ->
@@ -1272,6 +1273,7 @@ module Search = struct
search_tac_gl ~st only_classes dep hints depth (succ i) sigma gls gl end
in
Proofview.Unsafe.tclGETGOALS >>= fun gls ->
+ let gls = CList.map Proofview.drop_state gls in
Proofview.tclEVARMAP >>= fun sigma ->
let j = List.length gls in
(tclDISPATCH (List.init j (fun i -> tac sigma gls i)))
diff --git a/tactics/equality.ml b/tactics/equality.ml
index 0ce7fa022..96a53a8b1 100644
--- a/tactics/equality.ml
+++ b/tactics/equality.ml
@@ -533,7 +533,7 @@ let general_rewrite_clause l2r with_evars ?tac c cl =
let rec do_hyps_atleastonce = function
| [] -> tclZEROMSG (Pp.str"Nothing to rewrite.")
| id :: l ->
- tclIFTHENTRYELSEMUST
+ tclIFTHENFIRSTTRYELSEMUST
(general_rewrite_ebindings_in l2r AllOccurrences false true ?tac id c with_evars)
(do_hyps_atleastonce l)
in
@@ -549,7 +549,7 @@ let general_rewrite_clause l2r with_evars ?tac c cl =
end
in
if cl.concl_occs == NoOccurrences then do_hyps else
- tclIFTHENTRYELSEMUST
+ tclIFTHENFIRSTTRYELSEMUST
(general_rewrite_ebindings l2r (occs_of cl.concl_occs) false true ?tac c with_evars)
do_hyps
diff --git a/tactics/hints.ml b/tactics/hints.ml
index 7f9b5ef34..10cd7e1f6 100644
--- a/tactics/hints.ml
+++ b/tactics/hints.ml
@@ -1270,7 +1270,7 @@ let prepare_hint check (poly,local) env init (sigma,c) =
let interp_hints poly =
fun h ->
- let env = (Global.env()) in
+ let env = Global.env () in
let sigma = Evd.from_env env in
let f poly c =
let evd,c = Constrintern.interp_open_constr env sigma c in
@@ -1279,9 +1279,7 @@ let interp_hints poly =
let gr = global_with_alias r in
Dumpglob.add_glob ?loc:(loc_of_reference r) gr;
gr in
- let fr r =
- evaluable_of_global_reference (Global.env()) (fref r)
- in
+ let fr r = evaluable_of_global_reference env (fref r) in
let fi c =
match c with
| HintsReference c ->
@@ -1289,7 +1287,7 @@ let interp_hints poly =
(PathHints [gr], poly, IsGlobRef gr)
| HintsConstr c -> (PathAny, poly, f poly c)
in
- let fp = Constrintern.intern_constr_pattern (Global.env()) in
+ let fp = Constrintern.intern_constr_pattern env sigma in
let fres (info, b, r) =
let path, poly, gr = fi r in
let info = { info with hint_pattern = Option.map fp info.hint_pattern } in
diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml
index e7da17cff..0cc0001c1 100644
--- a/tactics/tacticals.ml
+++ b/tactics/tacticals.ml
@@ -408,8 +408,14 @@ module New = struct
Proofview.tclIFCATCH t1
(fun () -> tclDISPATCH (Array.to_list a))
(fun _ -> t3)
+ let tclIFTHENFIRSTELSE t1 t2 t3 =
+ Proofview.tclIFCATCH t1
+ (fun () -> tclEXTEND [t2] (tclUNIT ()) [])
+ (fun _ -> t3)
let tclIFTHENTRYELSEMUST t1 t2 =
tclIFTHENELSE t1 (tclTRY t2) t2
+ let tclIFTHENFIRSTTRYELSEMUST t1 t2 =
+ tclIFTHENFIRSTELSE t1 (tclTRY t2) t2
(* Try the first tactic that does not fail in a list of tactics *)
let rec tclFIRST = function
diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli
index c5d5c8c12..a3bc4707e 100644
--- a/tactics/tacticals.mli
+++ b/tactics/tacticals.mli
@@ -210,6 +210,7 @@ module New : sig
val tclIFTHENELSE : unit tactic -> unit tactic -> unit tactic -> unit tactic
val tclIFTHENSVELSE : unit tactic -> unit tactic array -> unit tactic -> unit tactic
val tclIFTHENTRYELSEMUST : unit tactic -> unit tactic -> unit tactic
+ val tclIFTHENFIRSTTRYELSEMUST : unit tactic -> unit tactic -> unit tactic
val tclDO : int -> unit tactic -> unit tactic
val tclREPEAT : unit tactic -> unit tactic
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index bf9271ba9..61ca3e319 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -1991,7 +1991,6 @@ let exact_proof c =
Proofview.Goal.enter begin fun gl ->
Refine.refine ~typecheck:false begin fun sigma ->
let (c, ctx) = Constrintern.interp_casted_constr (pf_env gl) sigma c (pf_concl gl) in
- let c = EConstr.of_constr c in
let sigma = Evd.merge_universe_context sigma ctx in
(sigma, c)
end
@@ -4662,30 +4661,6 @@ let destruct ev clr c l e =
induction_gen clr false ev e
((Evd.empty,(c,NoBindings)),(None,l)) None
-(* The registered tactic, which calls the default elimination
- * if no elimination constant is provided. *)
-
-(* Induction tactics *)
-
-(* This was Induction before 6.3 (induction only in quantified premisses) *)
-let simple_induct_id s = Tacticals.New.tclTHEN (intros_until_id s) (Tacticals.New.onLastHyp simplest_elim)
-let simple_induct_nodep n = Tacticals.New.tclTHEN (intros_until_n n) (Tacticals.New.onLastHyp simplest_elim)
-
-let simple_induct = function
- | NamedHyp id -> simple_induct_id id
- | AnonHyp n -> simple_induct_nodep n
-
-(* Destruction tactics *)
-
-let simple_destruct_id s =
- (Tacticals.New.tclTHEN (intros_until_id s) (Tacticals.New.onLastHyp simplest_case))
-let simple_destruct_nodep n =
- (Tacticals.New.tclTHEN (intros_until_n n) (Tacticals.New.onLastHyp simplest_case))
-
-let simple_destruct = function
- | NamedHyp id -> simple_destruct_id id
- | AnonHyp n -> simple_destruct_nodep n
-
(*
* Eliminations giving the type instead of the proof.
* These tactics use the default elimination constant and
diff --git a/tactics/tactics.mli b/tactics/tactics.mli
index 74415f8d0..100ddf17f 100644
--- a/tactics/tactics.mli
+++ b/tactics/tactics.mli
@@ -280,8 +280,6 @@ val simplest_elim : constr -> unit Proofview.tactic
val elim :
evars_flag -> clear_flag -> constr with_bindings -> constr with_bindings option -> unit Proofview.tactic
-val simple_induct : quantified_hypothesis -> unit Proofview.tactic
-
val induction : evars_flag -> clear_flag -> constr -> or_and_intro_pattern option ->
constr with_bindings option -> unit Proofview.tactic
@@ -290,7 +288,6 @@ val induction : evars_flag -> clear_flag -> constr -> or_and_intro_pattern optio
val general_case_analysis : evars_flag -> clear_flag -> constr with_bindings -> unit Proofview.tactic
val simplest_case : constr -> unit Proofview.tactic
-val simple_destruct : quantified_hypothesis -> unit Proofview.tactic
val destruct : evars_flag -> clear_flag -> constr -> or_and_intro_pattern option ->
constr with_bindings option -> unit Proofview.tactic
diff --git a/test-suite/bugs/closed/6878.v b/test-suite/bugs/closed/6878.v
new file mode 100644
index 000000000..70f1b3127
--- /dev/null
+++ b/test-suite/bugs/closed/6878.v
@@ -0,0 +1,8 @@
+
+Set Universe Polymorphism.
+Module Type T.
+ Axiom foo : Prop.
+End T.
+
+(** Used to anomaly *)
+Fail Module M : T with Definition foo := Type.
diff --git a/test-suite/output/Load.out b/test-suite/output/Load.out
new file mode 100644
index 000000000..0904d5540
--- /dev/null
+++ b/test-suite/output/Load.out
@@ -0,0 +1,6 @@
+f = 2
+ : nat
+u = I
+ : True
+The command has indeed failed with message:
+Files processed by Load cannot leave open proofs.
diff --git a/test-suite/output/Load.v b/test-suite/output/Load.v
new file mode 100644
index 000000000..967507415
--- /dev/null
+++ b/test-suite/output/Load.v
@@ -0,0 +1,7 @@
+Load "output/load/Load_noproof.v".
+Print f.
+
+Load "output/load/Load_proof.v".
+Print u.
+
+Fail Load "output/load/Load_openproof.v".
diff --git a/test-suite/output/load/Load_noproof.v b/test-suite/output/load/Load_noproof.v
new file mode 100644
index 000000000..aaf1ffe26
--- /dev/null
+++ b/test-suite/output/load/Load_noproof.v
@@ -0,0 +1 @@
+Definition f := 2.
diff --git a/test-suite/output/load/Load_openproof.v b/test-suite/output/load/Load_openproof.v
new file mode 100644
index 000000000..204d4ecbf
--- /dev/null
+++ b/test-suite/output/load/Load_openproof.v
@@ -0,0 +1 @@
+Lemma k : True.
diff --git a/test-suite/output/load/Load_proof.v b/test-suite/output/load/Load_proof.v
new file mode 100644
index 000000000..e47f66a19
--- /dev/null
+++ b/test-suite/output/load/Load_proof.v
@@ -0,0 +1,2 @@
+Lemma u : True.
+Proof. exact I. Qed.
diff --git a/test-suite/success/rewrite.v b/test-suite/success/rewrite.v
index 62249666b..448d0082d 100644
--- a/test-suite/success/rewrite.v
+++ b/test-suite/success/rewrite.v
@@ -151,10 +151,25 @@ Abort.
(* Check that rewriting within evars still work (was broken in 8.5beta1) *)
-
Goal forall (a: unit) (H: a = tt), exists x y:nat, x = y.
intros; eexists; eexists.
rewrite H.
Undo.
subst.
Abort.
+
+(* Check that iterated rewriting does not rewrite in the side conditions *)
+(* Example from Sigurd Schneider, extracted from contrib containers *)
+
+Lemma EQ
+ : forall (e e' : nat), True -> e = e'.
+Admitted.
+
+Lemma test (v1 v2 v3: nat) (v' : v1 = v2) : v2 = v1.
+Proof.
+ rewrite <- (EQ v1 v2) in *.
+ exact v'.
+ (* There should be only two side conditions *)
+ exact I.
+ exact I.
+Qed.
diff --git a/toplevel/coqinit.ml b/toplevel/coqinit.ml
index d8aaf3db8..6f461fd53 100644
--- a/toplevel/coqinit.ml
+++ b/toplevel/coqinit.ml
@@ -26,7 +26,7 @@ let load_rcfile ~rcfile ~time ~state =
match rcfile with
| Some rcfile ->
if CUnix.file_readable_p rcfile then
- Vernac.load_vernac ~time ~verbosely:false ~interactive:false ~check:true ~state rcfile
+ Vernac.load_vernac ~time ~echo:false ~interactive:false ~check:true ~state rcfile
else raise (Sys_error ("Cannot read rcfile: "^ rcfile))
| None ->
try
@@ -37,7 +37,7 @@ let load_rcfile ~rcfile ~time ~state =
Envars.home ~warn / "."^rcdefaultname^"."^Coq_config.version;
Envars.home ~warn / "."^rcdefaultname
] in
- Vernac.load_vernac ~time ~verbosely:false ~interactive:false ~check:true ~state inferedrc
+ Vernac.load_vernac ~time ~echo:false ~interactive:false ~check:true ~state inferedrc
with Not_found -> state
(*
Flags.if_verbose
diff --git a/toplevel/coqloop.ml b/toplevel/coqloop.ml
index ae0b94476..7ad8e2c05 100644
--- a/toplevel/coqloop.ml
+++ b/toplevel/coqloop.ml
@@ -339,6 +339,8 @@ let loop_flush_all () =
Format.pp_print_flush !Topfmt.std_ft ();
Format.pp_print_flush !Topfmt.err_ft ()
+let drop_last_doc = ref None
+
let rec loop ~time ~state =
let open Vernac.State in
Sys.catch_break true;
@@ -353,7 +355,14 @@ let rec loop ~time ~state =
not possible due exceptions. *)
in vernac_loop ~state
with
- | CErrors.Drop -> state
+ | CErrors.Drop ->
+ (* Due to using exceptions as a form of control, state here goes
+ out of sync as [do_vernac] will never return. We must thus do
+ this hack until we make `Drop` a toplevel-only command. See
+ bug #6872. *)
+ let state = { state with sid = Stm.get_current_state ~doc:state.doc } in
+ drop_last_doc := Some state;
+ state
| CErrors.Quit -> exit 0
| any ->
top_stderr (str "Anomaly: main loop exited with exception: " ++
diff --git a/toplevel/coqloop.mli b/toplevel/coqloop.mli
index 1c1309051..928f3609a 100644
--- a/toplevel/coqloop.mli
+++ b/toplevel/coqloop.mli
@@ -35,3 +35,6 @@ val do_vernac : time:bool -> state:Vernac.State.t -> Vernac.State.t
(** Main entry point of Coq: read and execute vernac commands. *)
val loop : time:bool -> state:Vernac.State.t -> Vernac.State.t
+
+(** Last document seen after `Drop` *)
+val drop_last_doc : Vernac.State.t option ref
diff --git a/toplevel/coqtop.ml b/toplevel/coqtop.ml
index 26ede1834..5afc4e182 100644
--- a/toplevel/coqtop.ml
+++ b/toplevel/coqtop.ml
@@ -34,7 +34,6 @@ let warning s = Flags.(with_option warn Feedback.msg_warning (strbrk s))
will not be generally be initialized, thus stateid, etc... may be
bogus. For now we just print to the console too *)
let coqtop_init_feed = Coqloop.coqloop_feed
-let drop_last_doc = ref None
(* Default toplevel loop *)
let console_toploop_run opts ~state =
@@ -44,9 +43,8 @@ let console_toploop_run opts ~state =
Flags.if_verbose warning "Dumpglob cannot be used in interactive mode.";
Dumpglob.noglob ()
end;
- let state = Coqloop.loop ~time:opts.time ~state in
+ let _ = Coqloop.loop ~time:opts.time ~state in
(* Initialise and launch the Ocaml toplevel *)
- drop_last_doc := Some state;
Coqinit.init_ocaml_path();
Mltop.ocaml_toploop();
(* We let the feeder in place for users of Drop *)
@@ -92,10 +90,10 @@ let outputstate opts =
(******************************************************************************)
let load_vernacular opts ~state =
List.fold_left
- (fun state (f_in, verbosely) ->
+ (fun state (f_in, echo) ->
let s = Loadpath.locate_file f_in in
(* Should make the beautify logic clearer *)
- let load_vernac f = Vernac.load_vernac ~time:opts.time ~verbosely ~interactive:false ~check:true ~state f in
+ let load_vernac f = Vernac.load_vernac ~time:opts.time ~echo ~interactive:false ~check:true ~state f in
if !Flags.beautify
then Flags.with_option Flags.beautify_file load_vernac f_in
else load_vernac s
@@ -194,7 +192,7 @@ let ensure_exists f =
fatal_error (hov 0 (str "Can't find file" ++ spc () ++ str f))
(* Compile a vernac file *)
-let compile opts ~verbosely ~f_in ~f_out =
+let compile opts ~echo ~f_in ~f_out =
let open Vernac.State in
let check_pending_proofs () =
let pfs = Proof_global.get_all_proof_names () in
@@ -232,7 +230,7 @@ let compile opts ~verbosely ~f_in ~f_out =
Dumpglob.start_dump_glob ~vfile:long_f_dot_v ~vofile:long_f_dot_vo;
Dumpglob.dump_string ("F" ^ Names.DirPath.to_string ldir ^ "\n");
let wall_clock1 = Unix.gettimeofday () in
- let state = Vernac.load_vernac ~time:opts.time ~verbosely ~check:true ~interactive:false ~state long_f_dot_v in
+ let state = Vernac.load_vernac ~time:opts.time ~echo ~check:true ~interactive:false ~state long_f_dot_v in
let _doc = Stm.join ~doc:state.doc in
let wall_clock2 = Unix.gettimeofday () in
check_pending_proofs ();
@@ -273,7 +271,7 @@ let compile opts ~verbosely ~f_in ~f_out =
let state = { doc; sid; proof = None } in
let state = load_init_vernaculars opts ~state in
let ldir = Stm.get_ldir ~doc:state.doc in
- let state = Vernac.load_vernac ~time:opts.time ~verbosely ~check:false ~interactive:false ~state long_f_dot_v in
+ let state = Vernac.load_vernac ~time:opts.time ~echo ~check:false ~interactive:false ~state long_f_dot_v in
let doc = Stm.finish ~doc:state.doc in
check_pending_proofs ();
let _doc = Stm.snapshot_vio ~doc ldir long_f_dot_vio in
@@ -288,17 +286,17 @@ let compile opts ~verbosely ~f_in ~f_out =
let univs, proofs = Stm.finish_tasks lfdv univs disch proofs tasks in
Library.save_library_raw lfdv sum lib univs proofs
-let compile opts ~verbosely ~f_in ~f_out =
+let compile opts ~echo ~f_in ~f_out =
ignore(CoqworkmgrApi.get 1);
- compile opts ~verbosely ~f_in ~f_out;
+ compile opts ~echo ~f_in ~f_out;
CoqworkmgrApi.giveback 1
-let compile_file opts (f_in, verbosely) =
+let compile_file opts (f_in, echo) =
if !Flags.beautify then
Flags.with_option Flags.beautify_file
- (fun f_in -> compile opts ~verbosely ~f_in ~f_out:None) f_in
+ (fun f_in -> compile opts ~echo ~f_in ~f_out:None) f_in
else
- compile opts ~verbosely ~f_in ~f_out:None
+ compile opts ~echo ~f_in ~f_out:None
let compile_files opts =
let compile_list = List.rev opts.compile_list in
diff --git a/toplevel/coqtop.mli b/toplevel/coqtop.mli
index dedb298e2..510e10dd1 100644
--- a/toplevel/coqtop.mli
+++ b/toplevel/coqtop.mli
@@ -15,9 +15,6 @@ val init_toplevel : string list -> Vernac.State.t option * Coqargs.coq_cmdopts
val start : unit -> unit
-(* Last document seen after `Drop` *)
-val drop_last_doc : Vernac.State.t option ref
-
(* For other toploops *)
val toploop_init : (Coqargs.coq_cmdopts -> string list -> string list) ref
val toploop_run : (Coqargs.coq_cmdopts -> state:Vernac.State.t -> unit) ref
diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml
index 92dee84f3..5d3addfec 100644
--- a/toplevel/vernac.ml
+++ b/toplevel/vernac.ml
@@ -40,37 +40,6 @@ let vernac_echo ?loc in_chan = let open Loc in
Feedback.msg_notice @@ str @@ really_input_string in_chan len
) loc
-(* vernac parses the given stream, executes interpfun on the syntax tree it
- * parses, and is verbose on "primitives" commands if verbosely is true *)
-
-let beautify_suffix = ".beautified"
-
-let set_formatter_translator ch =
- let out s b e = output_substring ch s b e in
- let ft = Format.make_formatter out (fun () -> flush ch) in
- Format.pp_set_max_boxes ft max_int;
- ft
-
-let pr_new_syntax_in_context ?loc ft_beautify ocom =
- let loc = Option.cata Loc.unloc (0,0) loc in
- let fs = States.freeze ~marshallable:`No in
- (* Side-effect: order matters *)
- let before = comment (CLexer.extract_comments (fst loc)) in
- let com = match ocom with
- | Some com -> Ppvernac.pr_vernac com
- | None -> mt() in
- let after = comment (CLexer.extract_comments (snd loc)) in
- if !Flags.beautify_file then
- (Pp.pp_with ft_beautify (hov 0 (before ++ com ++ after));
- Format.pp_print_flush ft_beautify ())
- else
- Feedback.msg_info (hov 4 (str"New Syntax:" ++ fnl() ++ (hov 0 com)));
- States.unfreeze fs
-
-let pr_new_syntax ?loc po ft_beautify ocom =
- (* Reinstall the context of parsing which includes the bindings of comments to locations *)
- Pcoq.Gram.with_parsable po (pr_new_syntax_in_context ?loc ft_beautify) ocom
-
(* For coqtop -time, we display the position in the file,
and a glimpse of the executed command *)
@@ -120,16 +89,9 @@ module State = struct
end
-let rec interp_vernac ~time ~check ~interactive ~state (loc,com) =
+let interp_vernac ~time ~check ~interactive ~state (loc,com) =
let open State in
- let interp v =
- match under_control v with
- | VernacLoad (verbosely, fname) ->
- let fname = Envars.expand_path_macros ~warn:(fun x -> Feedback.msg_warning (str x)) fname in
- let fname = CUnix.make_suffix fname ".v" in
- let f = Loadpath.locate_file fname in
- load_vernac ~time ~verbosely ~check ~interactive ~state f
- | _ ->
+ try
(* XXX: We need to run this before add as the classification is
highly dynamic and depends on the structure of the
document. Hopefully this is fixed when VtMeta can be removed
@@ -139,13 +101,17 @@ let rec interp_vernac ~time ~check ~interactive ~state (loc,com) =
against that... *)
let wflags = CWarnings.get_flags () in
CWarnings.set_flags "none";
- let is_proof_step = match fst (Vernac_classifier.classify_vernac v) with
+ let is_proof_step = match fst (Vernac_classifier.classify_vernac com) with
| VtProofStep _ | VtMeta | VtStartProof _ -> true
| _ -> false
in
CWarnings.set_flags wflags;
- let doc, nsid, ntip = Stm.add ~doc:state.doc ~ontop:state.sid (not !Flags.quiet) (loc,v) in
+ (* The -time option is only supported from console-based
+ clients due to the way it prints. *)
+ if time then print_cmd_header ?loc com;
+ let com = if time then VernacTime(time,(CAst.make ?loc com)) else com in
+ let doc, nsid, ntip = Stm.add ~doc:state.doc ~ontop:state.sid (not !Flags.quiet) (loc,com) in
(* Main STM interaction *)
if ntip <> `NewTip then
@@ -165,13 +131,6 @@ let rec interp_vernac ~time ~check ~interactive ~state (loc,com) =
if print_goals then Feedback.msg_notice (pr_open_cur_subgoals ());
let new_proof = Proof_global.give_me_the_proof_opt () in
{ doc = ndoc; sid = nsid; proof = new_proof }
- in
- try
- (* The -time option is only supported from console-based
- clients due to the way it prints. *)
- if time then print_cmd_header ?loc com;
- let com = if time then VernacTime(time, CAst.make ?loc com) else com in
- interp com
with reraise ->
(* XXX: In non-interactive mode edit_at seems to do very weird
things, so we better avoid it while we investigate *)
@@ -184,18 +143,16 @@ let rec interp_vernac ~time ~check ~interactive ~state (loc,com) =
end in iraise (reraise, info)
(* Load a vernac file. CErrors are annotated with file and location *)
-and load_vernac ~time ~verbosely ~check ~interactive ~state file =
- let ft_beautify, close_beautify =
- if !Flags.beautify_file then
- let chan_beautify = open_out (file^beautify_suffix) in
- set_formatter_translator chan_beautify, fun () -> close_out chan_beautify;
- else
- !Topfmt.std_ft, fun () -> ()
- in
+let load_vernac_core ~time ~echo ~check ~interactive ~state file =
+ (* Keep in sync *)
let in_chan = open_utf8_file_in file in
- let in_echo = if verbosely then Some (open_utf8_file_in file) else None in
+ let in_echo = if echo then Some (open_utf8_file_in file) else None in
+ let input_cleanup () = close_in in_chan; Option.iter close_in in_echo in
+
let in_pa = Pcoq.Gram.parsable ~file:(Loc.InFile file) (Stream.of_channel in_chan) in
let rstate = ref state in
+ (* For beautify, list of parsed sids *)
+ let rids = ref [] in
let open State in
try
(* we go out of the following infinite loop when a End_of_input is
@@ -222,36 +179,78 @@ and load_vernac ~time ~verbosely ~check ~interactive ~state file =
*)
in
(* Printing of vernacs *)
- if !Flags.beautify then pr_new_syntax ?loc in_pa ft_beautify (Some ast);
Option.iter (vernac_echo ?loc) in_echo;
checknav_simple (loc, ast);
let state = Flags.silently (interp_vernac ~time ~check ~interactive ~state:!rstate) (loc, ast) in
+ rids := state.sid :: !rids;
rstate := state;
done;
- !rstate
+ input_cleanup ();
+ !rstate, !rids, Pcoq.Gram.comment_state in_pa
with any -> (* whatever the exception *)
let (e, info) = CErrors.push any in
- close_in in_chan;
- Option.iter close_in in_echo;
+ input_cleanup ();
match e with
- | Stm.End_of_input ->
- (* Is this called so comments at EOF are printed? *)
- if !Flags.beautify then
- pr_new_syntax ~loc:(Loc.make_loc (max_int,max_int)) in_pa ft_beautify None;
- if !Flags.beautify_file then close_beautify ();
- !rstate
- | reraise ->
- if !Flags.beautify_file then close_beautify ();
- iraise (disable_drop e, info)
-
-(** [eval_expr : ?preserving:bool -> Loc.t * Vernacexpr.vernac_expr -> unit]
- It executes one vernacular command. By default the command is
- 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. *)
+ | Stm.End_of_input -> !rstate, !rids, Pcoq.Gram.comment_state in_pa
+ | reraise -> iraise (disable_drop e, info)
let process_expr ~time ~state loc_ast =
checknav_deep loc_ast;
interp_vernac ~time ~interactive:true ~check:true ~state loc_ast
+
+(******************************************************************************)
+(* Beautify-specific code *)
+(******************************************************************************)
+
+(* vernac parses the given stream, executes interpfun on the syntax tree it
+ * parses, and is verbose on "primitives" commands if verbosely is true *)
+let beautify_suffix = ".beautified"
+
+let set_formatter_translator ch =
+ let out s b e = output_substring ch s b e in
+ let ft = Format.make_formatter out (fun () -> flush ch) in
+ Format.pp_set_max_boxes ft max_int;
+ ft
+
+let pr_new_syntax ?loc ft_beautify ocom =
+ let loc = Option.cata Loc.unloc (0,0) loc in
+ let before = comment (Pputils.extract_comments (fst loc)) in
+ let com = Option.cata Ppvernac.pr_vernac (mt ()) ocom in
+ let after = comment (Pputils.extract_comments (snd loc)) in
+ if !Flags.beautify_file then
+ (Pp.pp_with ft_beautify (hov 0 (before ++ com ++ after));
+ Format.pp_print_flush ft_beautify ())
+ else
+ Feedback.msg_info (hov 4 (str"New Syntax:" ++ fnl() ++ (hov 0 com)))
+
+(* load_vernac with beautify *)
+let beautify_pass ~doc ~comments ~ids ~filename =
+ let ft_beautify, close_beautify =
+ if !Flags.beautify_file then
+ let chan_beautify = open_out (filename^beautify_suffix) in
+ set_formatter_translator chan_beautify, fun () -> close_out chan_beautify;
+ else
+ !Topfmt.std_ft, fun () -> ()
+ in
+ (* The interface to the comment printer is imperative, so we first
+ set the comments, then we call print. This has to be done for
+ each file. *)
+ Pputils.beautify_comments := comments;
+ List.iter (fun id ->
+ Option.iter (fun (loc,ast) ->
+ pr_new_syntax ?loc ft_beautify (Some ast))
+ (Stm.get_ast ~doc id)) ids;
+
+ (* Is this called so comments at EOF are printed? *)
+ pr_new_syntax ~loc:(Loc.make_loc (max_int,max_int)) ft_beautify None;
+ close_beautify ()
+
+(* Main driver for file loading. For now, we only do one beautify
+ pass. *)
+let load_vernac ~time ~echo ~check ~interactive ~state filename =
+ let ostate, ids, comments = load_vernac_core ~time ~echo ~check ~interactive ~state filename in
+ (* Pass for beautify *)
+ if !Flags.beautify then beautify_pass ~doc:ostate.State.doc ~comments ~ids:List.(rev ids) ~filename;
+ (* End pass *)
+ ostate
diff --git a/toplevel/vernac.mli b/toplevel/vernac.mli
index e909ada1e..ca825197b 100644
--- a/toplevel/vernac.mli
+++ b/toplevel/vernac.mli
@@ -26,5 +26,5 @@ val process_expr : time:bool -> state:State.t -> Vernacexpr.vernac_control Loc.l
(** [load_vernac echo sid file] Loads [file] on top of [sid], will
echo the commands if [echo] is set. Callers are expected to handle
and print errors in form of exceptions. *)
-val load_vernac : time:bool -> verbosely:bool -> check:bool -> interactive:bool ->
+val load_vernac : time:bool -> echo:bool -> check:bool -> interactive:bool ->
state:State.t -> string -> State.t
diff --git a/vernac/classes.ml b/vernac/classes.ml
index 695be74bb..cd5eff4e7 100644
--- a/vernac/classes.ml
+++ b/vernac/classes.ml
@@ -52,7 +52,7 @@ let _ =
let open Vernacexpr in
{ info with hint_pattern =
Option.map
- (Constrintern.intern_constr_pattern (Global.env()))
+ (Constrintern.intern_constr_pattern (Global.env()) Evd.(from_env Global.(env())))
info.hint_pattern } in
Flags.silently (fun () ->
Hints.add_hints local [typeclasses_db]
@@ -334,7 +334,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance)
let init_refine =
Tacticals.New.tclTHENLIST [
Refine.refine ~typecheck:false (fun sigma -> (sigma,EConstr.of_constr (Option.get term)));
- Proofview.Unsafe.tclNEWGOALS gls;
+ Proofview.Unsafe.tclNEWGOALS (CList.map Proofview.with_empty_state gls);
Tactics.New.reduce_after_refine;
]
in
diff --git a/vernac/comAssumption.ml b/vernac/comAssumption.ml
index 7e5b941ad..9fd2153cb 100644
--- a/vernac/comAssumption.ml
+++ b/vernac/comAssumption.ml
@@ -10,7 +10,6 @@ open Pp
open CErrors
open Util
open Vars
-open Environ
open Declare
open Names
open Globnames
@@ -87,7 +86,6 @@ match local with
let interp_assumption sigma env impls bl c =
let c = mkCProdN ?loc:(local_binders_loc bl) bl c in
let sigma, (ty, impls) = interp_type_evars_impls env sigma ~impls c in
- let ty = EConstr.Unsafe.to_constr ty in
sigma, (ty, impls)
(* When monomorphic the universe constraints are declared with the first declaration only. *)
@@ -151,9 +149,9 @@ let do_assumptions kind nl l =
let (sigma,_,_),l = List.fold_left_map (fun (sigma,env,ienv) (is_coe,(idl,c)) ->
let sigma,(t,imps) = interp_assumption sigma env ienv [] c in
let env =
- push_named_context (List.map (fun {CAst.v=id} -> LocalAssum (id,t)) idl) env in
+ EConstr.push_named_context (List.map (fun {CAst.v=id} -> LocalAssum (id,t)) idl) env in
let ienv = List.fold_right (fun {CAst.v=id} ienv ->
- let impls = compute_internalization_data env Variable t imps in
+ let impls = compute_internalization_data env sigma Variable t imps in
Id.Map.add id impls ienv) idl ienv in
((sigma,env,ienv),((is_coe,idl),t,imps)))
(sigma,env,empty_internalization_env) l
@@ -161,7 +159,7 @@ let do_assumptions kind nl l =
let sigma = solve_remaining_evars all_and_fail_flags env sigma Evd.empty in
(* The universe constraints come from the whole telescope. *)
let sigma = Evd.nf_constraints sigma in
- let nf_evar c = EConstr.to_constr sigma (EConstr.of_constr c) in
+ let nf_evar c = EConstr.to_constr sigma c in
let uvars, l = List.fold_left_map (fun uvars (coe,t,imps) ->
let t = nf_evar t in
let uvars = Univ.LSet.union uvars (Univops.universes_of_constr env t) in
diff --git a/vernac/comFixpoint.ml b/vernac/comFixpoint.ml
index 489f299a2..edfe7aa81 100644
--- a/vernac/comFixpoint.ml
+++ b/vernac/comFixpoint.ml
@@ -212,8 +212,7 @@ let interp_recursive ~program_mode ~cofix fixl notations =
let env_rec = push_named_context rec_sign env in
(* Get interpretation metadatas *)
- let fixtypes = List.map EConstr.Unsafe.to_constr fixtypes in
- let impls = compute_internalization_env env Recursive fixnames fixtypes fiximps in
+ let impls = compute_internalization_env env sigma Recursive fixnames fixtypes fiximps in
(* Interp bodies with rollback because temp use of notations/implicit *)
let sigma, fixdefs =
@@ -226,10 +225,9 @@ let interp_recursive ~program_mode ~cofix fixl notations =
(* Instantiate evars and check all are resolved *)
let sigma = solve_unif_constraints_with_heuristics env_rec sigma in
- let sigma, nf = nf_evars_and_universes sigma in
- let fixdefs = List.map (fun c -> Option.map EConstr.Unsafe.to_constr c) fixdefs in
- let fixdefs = List.map (Option.map nf) fixdefs in
- let fixtypes = List.map nf fixtypes in
+ let sigma, _ = nf_evars_and_universes sigma in
+ let fixdefs = List.map (fun c -> Option.map EConstr.(to_constr sigma) c) fixdefs in
+ let fixtypes = List.map EConstr.(to_constr sigma) fixtypes in
let fixctxs = List.map (fun (_,ctx) -> ctx) fixctxs in
(* Build the fix declaration block *)
diff --git a/vernac/comInductive.ml b/vernac/comInductive.ml
index c650e9e40..cef5546c6 100644
--- a/vernac/comInductive.ml
+++ b/vernac/comInductive.ml
@@ -11,7 +11,6 @@ open CErrors
open Sorts
open Util
open Constr
-open Termops
open Environ
open Declare
open Names
@@ -51,7 +50,7 @@ let rec complete_conclusion a cs = CAst.map_with_loc (fun ?loc -> function
)
let push_types env idl tl =
- List.fold_left2 (fun env id t -> Environ.push_rel (LocalAssum (Name id,t)) env)
+ List.fold_left2 (fun env id t -> EConstr.push_rel (LocalAssum (Name id,t)) env)
env idl tl
type structured_one_inductive_expr = {
@@ -90,7 +89,7 @@ let check_all_names_different indl =
| _ -> raise (InductiveError (SameNamesOverlap l))
let mk_mltype_data sigma env assums arity indname =
- let is_ml_type = is_sort env sigma (EConstr.of_constr arity) in
+ let is_ml_type = is_sort env sigma arity in
(is_ml_type,indname,assums)
let prepare_param = function
@@ -130,14 +129,13 @@ let is_impredicative env u =
u = Prop Null || (is_impredicative_set env && u = Prop Pos)
let interp_ind_arity env sigma ind =
- let c = intern_gen IsType env ind.ind_arity in
+ let c = intern_gen IsType env sigma ind.ind_arity in
let impls = Implicit_quantifiers.implicits_of_glob_constr ~with_products:true c in
let sigma,t = understand_tcc env sigma ~expected_type:IsType c in
let pseudo_poly = check_anonymous_type c in
let () = if not (Reductionops.is_arity env sigma t) then
user_err ?loc:(constr_loc ind.ind_arity) (str "Not an arity")
in
- let t = EConstr.Unsafe.to_constr t in
sigma, (t, pseudo_poly, impls)
let interp_cstrs env sigma impls mldata arity ind =
@@ -272,7 +270,6 @@ let interp_mutual_inductive (paramsl,indl) notations cum poly prv finite =
let sigma, (impls, ((env_params, ctx_params), userimpls)) =
interp_context_evars env0 sigma paramsl
in
- let ctx_params = List.map (fun d -> map_rel_decl EConstr.Unsafe.to_constr d) ctx_params in
let indnames = List.map (fun ind -> ind.ind_name) indl in
(* Names of parameters as arguments of the inductive type (defs removed) *)
@@ -282,16 +279,16 @@ let interp_mutual_inductive (paramsl,indl) notations cum poly prv finite =
(* Interpret the arities *)
let sigma, arities = List.fold_left_map (fun sigma -> interp_ind_arity env_params sigma) sigma indl in
- let fullarities = List.map (fun (c, _, _) -> Term.it_mkProd_or_LetIn c ctx_params) arities in
+ let fullarities = List.map (fun (c, _, _) -> EConstr.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
+ let env_ar_params = EConstr.push_rel_context ctx_params env_ar in
(* Compute interpretation metadatas *)
let indimpls = List.map (fun (_, _, impls) -> userimpls @
lift_implicits (Context.Rel.nhyps ctx_params) impls) arities in
let arities = List.map pi1 arities and aritypoly = List.map pi2 arities in
- let impls = compute_internalization_env env0 ~impls (Inductive (params,true)) indnames fullarities indimpls in
- let ntn_impls = compute_internalization_env env0 (Inductive (params,true)) indnames fullarities indimpls in
+ let impls = compute_internalization_env env0 sigma ~impls (Inductive (params,true)) indnames fullarities indimpls in
+ let ntn_impls = compute_internalization_env env0 sigma (Inductive (params,true)) indnames fullarities indimpls in
let mldatas = List.map2 (mk_mltype_data sigma env_params params) arities indnames in
let sigma, constructors =
@@ -306,15 +303,14 @@ let interp_mutual_inductive (paramsl,indl) notations cum poly prv finite =
let sigma = solve_remaining_evars all_and_fail_flags env_params sigma Evd.empty in
(* Compute renewed arities *)
let sigma, nf = nf_evars_and_universes sigma in
- let arities = List.map nf arities in
let constructors = List.map (fun (idl,cl,impsl) -> (idl,List.map nf cl,impsl)) constructors in
+ let arities = List.map EConstr.(to_constr sigma) arities in
let sigma = List.fold_left2 (fun sigma ty poly -> make_conclusion_flexible sigma ty poly) sigma arities aritypoly in
let sigma, arities = inductive_levels env_ar_params sigma poly arities constructors in
let sigma, nf' = nf_evars_and_universes sigma in
- let nf x = nf' (nf x) in
let arities = List.map nf' arities in
let constructors = List.map (fun (idl,cl,impsl) -> (idl,List.map nf' cl,impsl)) constructors in
- let ctx_params = Context.Rel.map nf ctx_params in
+ let ctx_params = List.map Termops.(map_rel_decl (EConstr.to_constr sigma)) ctx_params in
let uctx = Evd.check_univ_decl ~poly sigma decl in
List.iter (fun c -> check_evars env_params Evd.empty sigma (EConstr.of_constr c)) arities;
Context.Rel.iter (fun c -> check_evars env0 Evd.empty sigma (EConstr.of_constr c)) ctx_params;
diff --git a/vernac/comProgramFixpoint.ml b/vernac/comProgramFixpoint.ml
index af34f8b29..bd7ee0978 100644
--- a/vernac/comProgramFixpoint.ml
+++ b/vernac/comProgramFixpoint.ml
@@ -171,8 +171,8 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
let sigma, intern_body =
let ctx = LocalAssum (Name recname, get_type curry_fun) :: binders_rel in
let (r, l, impls, scopes) =
- Constrintern.compute_internalization_data env
- Constrintern.Recursive (EConstr.Unsafe.to_constr full_arity) impls
+ Constrintern.compute_internalization_data env sigma
+ Constrintern.Recursive full_arity impls
in
let newimpls = Id.Map.singleton recname
(r, l, impls @ [(Some (Id.of_string "recproof", Impargs.Manual, (true, false)))],
diff --git a/vernac/obligations.ml b/vernac/obligations.ml
index e4bcbc4bb..6447fc350 100644
--- a/vernac/obligations.ml
+++ b/vernac/obligations.ml
@@ -816,13 +816,13 @@ let solve_by_tac name evi t poly ctx =
let id = name in
let concl = EConstr.of_constr evi.evar_concl in
(* spiwack: the status is dropped. *)
- let (entry,_,ctx') = Pfedit.build_constant_by_tactic
+ 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 = Safe_typing.inline_private_constants_in_definition_entry env entry in
let body, () = Future.force entry.const_entry_body in
let ctx' = Evd.merge_context_set ~sideff:true Evd.univ_rigid (Evd.from_ctx ctx') (snd body) in
- Inductiveops.control_only_guard (Global.env ()) (fst body);
+ Inductiveops.control_only_guard env ctx' (EConstr.of_constr (fst body));
(fst body), entry.const_entry_type, Evd.evar_universe_context ctx'
let obligation_terminator name num guard hook auto pf =
@@ -838,7 +838,7 @@ let obligation_terminator name num guard hook auto pf =
let (body, cstr), () = Future.force entry.Entries.const_entry_body in
let sigma = Evd.from_ctx uctx in
let sigma = Evd.merge_context_set ~sideff:true Evd.univ_rigid sigma cstr in
- Inductiveops.control_only_guard (Global.env ()) body;
+ Inductiveops.control_only_guard (Global.env ()) sigma (EConstr.of_constr body);
(** Declare the obligation ourselves and drop the hook *)
let prg = get_info (ProgMap.find name !from_prg) in
(** Ensure universes are substituted properly in body and type *)
diff --git a/vernac/record.ml b/vernac/record.ml
index 1140e3d37..44113bfad 100644
--- a/vernac/record.ml
+++ b/vernac/record.ml
@@ -70,7 +70,7 @@ let interp_fields_evars env sigma impls_env nots l =
let impls =
match i with
| Anonymous -> impls
- | Name id -> Id.Map.add id (compute_internalization_data env Constrintern.Method (EConstr.to_constr sigma t') impl) impls
+ | Name id -> Id.Map.add id (compute_internalization_data env sigma Constrintern.Method t' impl) impls
in
let d = match b' with
| None -> LocalAssum (i,t')
@@ -145,7 +145,7 @@ let typecheck_params_and_fields finite def id poly pl t ps nots fs =
let assums = List.filter is_local_assum newps in
let params = List.map (RelDecl.get_name %> Name.get_id) assums in
let ty = Inductive (params,(finite != Declarations.BiFinite)) in
- let impls_env = compute_internalization_env env0 ~impls:impls_env ty [id] [EConstr.to_constr sigma arity] [imps] in
+ let impls_env = compute_internalization_env env0 sigma ~impls:impls_env ty [id] [arity] [imps] in
let env2,sigma,impls,newfs,data =
interp_fields_evars env_ar sigma impls_env nots (binders_of_decls fs)
in
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
index 4613100fc..6b7746db4 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -1025,7 +1025,9 @@ let vernac_arguments ~atts reference args more_implicits nargs_for_red flags =
let sr = smart_global reference in
let inf_names =
let ty, _ = Global.type_of_global_in_context (Global.env ()) sr in
- Impargs.compute_implicits_names (Global.env ()) ty
+ let env = Global.env () in
+ let sigma = Evd.from_env env in
+ Impargs.compute_implicits_names env sigma (EConstr.of_constr ty)
in
let prev_names =
try Arguments_renaming.arguments_names sr with Not_found -> inf_names
@@ -1253,7 +1255,7 @@ let vernac_reserve bl =
let env = Global.env() in
let sigma = Evd.from_env env in
let t,ctx = Constrintern.interp_type env sigma c in
- let t = Detyping.detype Detyping.Now false Id.Set.empty env (Evd.from_ctx ctx) (EConstr.of_constr t) in
+ let t = Detyping.detype Detyping.Now false Id.Set.empty env (Evd.from_ctx ctx) 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
@@ -1627,6 +1629,7 @@ let vernac_global_check c =
let senv = Global.safe_env() in
let uctx = UState.context_set uctx in
let senv = Safe_typing.push_context_set false uctx senv in
+ let c = EConstr.to_constr sigma c in
let j = Safe_typing.typing senv c in
let env = Safe_typing.env_of_safe_env senv in
print_safe_judgment env sigma j ++
@@ -1747,10 +1750,10 @@ let interp_search_restriction = function
open Search
-let interp_search_about_item env =
+let interp_search_about_item env sigma =
function
| SearchSubPattern pat ->
- let _,pat = intern_constr_pattern env pat in
+ let _,pat = intern_constr_pattern env sigma pat in
GlobSearchSubPattern pat
| SearchString (s,None) when Id.is_valid s ->
GlobSearchString s
@@ -1796,13 +1799,13 @@ let vernac_search ~atts s gopt r =
(* 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
+ let get_pattern c = snd (intern_constr_pattern env Evd.(from_env env) c) in
let pr_search ref env c =
let pr = pr_global ref in
let pp = if !search_output_name_only
then pr
else begin
- let pc = pr_lconstr_env env Evd.empty c in
+ let pc = pr_lconstr_env env Evd.(from_env env) c in
hov 2 (pr ++ str":" ++ spc () ++ pc)
end
in Feedback.msg_notice pp
@@ -1815,7 +1818,8 @@ let vernac_search ~atts s gopt r =
| SearchHead c ->
(Search.search_by_head gopt (get_pattern c) r |> Search.prioritize_search) pr_search
| SearchAbout sl ->
- (Search.search_about gopt (List.map (on_snd (interp_search_about_item env)) sl) r |> Search.prioritize_search) pr_search
+ (Search.search_about gopt (List.map (on_snd (interp_search_about_item env Evd.(from_env env))) sl) r |>
+ Search.prioritize_search) pr_search
let vernac_locate = function
| LocateAny (AN qid) -> print_located_qualid qid
@@ -1910,8 +1914,7 @@ let vernac_check_guard () =
let message =
try
let { Evd.it=gl ; sigma=sigma } = Proof.V82.top_goal pts in
- Inductiveops.control_only_guard (Goal.V82.env sigma gl)
- (EConstr.Unsafe.to_constr pfterm);
+ Inductiveops.control_only_guard (Goal.V82.env sigma gl) sigma pfterm;
(str "The condition holds up to here")
with UserError(_,s) ->
(str ("Condition violated: ") ++s)
@@ -1926,6 +1929,8 @@ exception End_of_input
without a considerable amount of refactoring.
*)
let vernac_load interp fname =
+ if Proof_global.there_are_pending_proofs () then
+ CErrors.user_err Pp.(str "Load is not supported inside proofs.");
let interp x =
let proof_mode = Proof_global.get_default_proof_mode_name () [@ocaml.warning "-3"] in
Proof_global.activate_proof_mode proof_mode [@ocaml.warning "-3"];
@@ -1942,8 +1947,13 @@ let vernac_load interp fname =
let longfname = Loadpath.locate_file fname in
let in_chan = open_utf8_file_in longfname in
Pcoq.Gram.parsable ~file:(Loc.InFile longfname) (Stream.of_channel in_chan) in
- try while true do interp (snd (parse_sentence input)) done
- with End_of_input -> ()
+ begin
+ try while true do interp (snd (parse_sentence input)) done
+ with End_of_input -> ()
+ end;
+ (* If Load left a proof open, we fail too. *)
+ if Proof_global.there_are_pending_proofs () then
+ CErrors.user_err Pp.(str "Files processed by Load cannot leave open proofs.")
(* "locality" is the prefix "Local" attribute, while the "local" component
* is the outdated/deprecated "Local" attribute of some vernacular commands
@@ -1954,6 +1964,7 @@ let interp ?proof ~atts ~st c =
vernac_pperr_endline (fun () -> str "interpreting: " ++ Ppvernac.pr_vernac_expr c);
match c with
+ (* Loading a file requires access to the control interpreter *)
| VernacLoad _ -> assert false
(* The STM should handle that, but LOAD bypasses the STM... *)