aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorMaxime Dénès2019-03-28 13:36:52 +0100
committerMaxime Dénès2019-03-28 13:36:52 +0100
commit688e20c432d2639050a62703e1c566ddfbe42b2a (patch)
tree3b34d3bd3b73a42a8eb730a3bb6c0e6a5cb00a5f
parent6d0ffe795f6f29730d59c379285201fd46023935 (diff)
parent91dfe5163fd4405977ad8fc8fe178ba5bcd73c88 (diff)
Merge PR #9129: [proof] Removal of imperative state ; interpretation layers only.
Ack-by: SkySkimmer Reviewed-by: aspiwack Ack-by: ejgallego Ack-by: gares Ack-by: herbelin Reviewed-by: mattam82 Ack-by: maximedenes
-rw-r--r--coqpp/coqpp_ast.mli1
-rw-r--r--coqpp/coqpp_lex.mll1
-rw-r--r--coqpp/coqpp_main.ml13
-rw-r--r--coqpp/coqpp_parse.mly21
-rw-r--r--dev/ci/user-overlays/09129-ejgallego-proof+no_global_partial.sh30
-rw-r--r--dev/top_printers.ml12
-rw-r--r--doc/plugin_tutorial/tuto1/src/g_tuto1.mlg14
-rw-r--r--doc/plugin_tutorial/tuto1/src/simple_declare.ml6
-rw-r--r--doc/sphinx/proof-engine/ltac.rst10
-rw-r--r--dune2
-rw-r--r--ide/idetop.ml21
-rw-r--r--plugins/cc/ccalgo.mli2
-rw-r--r--plugins/derive/derive.ml7
-rw-r--r--plugins/derive/derive.mli2
-rw-r--r--plugins/derive/g_derive.mlg4
-rw-r--r--plugins/extraction/extract_env.ml13
-rw-r--r--plugins/extraction/extract_env.mli2
-rw-r--r--plugins/extraction/g_extraction.mlg4
-rw-r--r--plugins/funind/functional_principles_proofs.ml53
-rw-r--r--plugins/funind/functional_principles_types.ml81
-rw-r--r--plugins/funind/g_indfun.mlg54
-rw-r--r--plugins/funind/glob_term_to_relation.ml2
-rw-r--r--plugins/funind/indfun.ml54
-rw-r--r--plugins/funind/indfun.mli14
-rw-r--r--plugins/funind/indfun_common.ml3
-rw-r--r--plugins/funind/indfun_common.mli6
-rw-r--r--plugins/funind/invfun.ml22
-rw-r--r--plugins/funind/recdef.ml349
-rw-r--r--plugins/funind/recdef.mli2
-rw-r--r--plugins/ltac/extratactics.mlg18
-rw-r--r--plugins/ltac/g_auto.mlg1
-rw-r--r--plugins/ltac/g_ltac.mlg35
-rw-r--r--plugins/ltac/g_obligations.mlg16
-rw-r--r--plugins/ltac/g_rewrite.mlg63
-rw-r--r--plugins/ltac/pptactic.ml1
-rw-r--r--plugins/ltac/rewrite.ml125
-rw-r--r--plugins/ltac/rewrite.mli14
-rw-r--r--plugins/ltac/tacinterp.ml75
-rw-r--r--plugins/ltac/tacinterp.mli7
-rw-r--r--plugins/ltac/tactic_debug.ml5
-rw-r--r--plugins/micromega/coq_micromega.ml19
-rw-r--r--plugins/setoid_ring/g_newring.mlg22
-rw-r--r--plugins/ssr/ssrcommon.ml5
-rw-r--r--plugins/ssr/ssrelim.ml14
-rw-r--r--plugins/ssr/ssrequality.ml25
-rw-r--r--plugins/ssr/ssrfwd.ml2
-rw-r--r--plugins/ssr/ssrvernac.mlg12
-rw-r--r--plugins/ssrmatching/ssrmatching.ml23
-rw-r--r--plugins/ssrmatching/ssrmatching.mli1
-rw-r--r--plugins/syntax/g_numeral.mlg21
-rw-r--r--plugins/syntax/g_string.mlg20
-rw-r--r--pretyping/geninterp.ml7
-rw-r--r--pretyping/geninterp.mli7
-rw-r--r--pretyping/globEnv.ml6
-rw-r--r--pretyping/globEnv.mli4
-rw-r--r--pretyping/pretyping.ml28
-rw-r--r--pretyping/pretyping.mli1
-rw-r--r--printing/proof_diffs.ml9
-rw-r--r--proofs/evar_refiner.ml1
-rw-r--r--proofs/pfedit.ml67
-rw-r--r--proofs/pfedit.mli21
-rw-r--r--proofs/proof_global.ml241
-rw-r--r--proofs/proof_global.mli65
-rw-r--r--stm/proofBlockDelimiter.ml13
-rw-r--r--stm/stm.ml77
-rw-r--r--tactics/auto.ml7
-rw-r--r--tactics/autorewrite.ml8
-rw-r--r--tactics/class_tactics.ml11
-rw-r--r--tactics/hints.ml4
-rw-r--r--tactics/hints.mli2
-rw-r--r--tactics/tactics.ml1
-rw-r--r--toplevel/ccompile.ml2
-rw-r--r--toplevel/coqc.ml3
-rw-r--r--toplevel/coqloop.ml6
-rw-r--r--toplevel/vernac.ml2
-rw-r--r--vernac/classes.ml71
-rw-r--r--vernac/classes.mli10
-rw-r--r--vernac/comAssumption.ml12
-rw-r--r--vernac/comAssumption.mli25
-rw-r--r--vernac/comDefinition.ml4
-rw-r--r--vernac/comDefinition.mli17
-rw-r--r--vernac/comFixpoint.ml47
-rw-r--r--vernac/comFixpoint.mli17
-rw-r--r--vernac/declareDef.ml8
-rw-r--r--vernac/declareDef.mli6
-rw-r--r--vernac/lemmas.ml145
-rw-r--r--vernac/lemmas.mli34
-rw-r--r--vernac/obligations.ml45
-rw-r--r--vernac/obligations.mli15
-rw-r--r--vernac/search.ml33
-rw-r--r--vernac/search.mli12
-rw-r--r--vernac/vernacentries.ml800
-rw-r--r--vernac/vernacentries.mli8
-rw-r--r--vernac/vernacstate.ml84
-rw-r--r--vernac/vernacstate.mli58
95 files changed, 1927 insertions, 1446 deletions
diff --git a/coqpp/coqpp_ast.mli b/coqpp/coqpp_ast.mli
index 8e10ec49ce..81109887ba 100644
--- a/coqpp/coqpp_ast.mli
+++ b/coqpp/coqpp_ast.mli
@@ -103,6 +103,7 @@ type classification =
type vernac_rule = {
vernac_atts : (string * string) list option;
+ vernac_state: string option;
vernac_toks : ext_token list;
vernac_class : code option;
vernac_depr : bool;
diff --git a/coqpp/coqpp_lex.mll b/coqpp/coqpp_lex.mll
index c38755943a..81ba8ad98c 100644
--- a/coqpp/coqpp_lex.mll
+++ b/coqpp/coqpp_lex.mll
@@ -130,6 +130,7 @@ rule extend = parse
| space { extend lexbuf }
| '\"' { string lexbuf }
| '\n' { newline lexbuf; extend lexbuf }
+| "![" { BANGBRACKET }
| "#[" { HASHBRACKET }
| '[' { LBRACKET }
| ']' { RBRACKET }
diff --git a/coqpp/coqpp_main.ml b/coqpp/coqpp_main.ml
index d33eef135f..90158c3aa3 100644
--- a/coqpp/coqpp_main.ml
+++ b/coqpp/coqpp_main.ml
@@ -347,9 +347,18 @@ let print_atts_right fmt = function
let nota = match atts with [_] -> "" | _ -> "Attributes.Notations." in
fprintf fmt "(Attributes.parse %s%a atts)" nota aux atts
+let print_body_wrapper fmt r =
+ match r.vernac_state with
+ | Some "proof" ->
+ fprintf fmt "let proof = (%a) ~pstate:st.Vernacstate.proof in { st with Vernacstate.proof }" print_code r.vernac_body
+ | None ->
+ fprintf fmt "let () = %a in st" print_code r.vernac_body
+ | Some x ->
+ fatal ("unsupported state specifier: " ^ x)
+
let print_body_fun fmt r =
- fprintf fmt "let coqpp_body %a%a ~st = let () = %a in st in "
- print_binders r.vernac_toks print_atts_left r.vernac_atts print_code r.vernac_body
+ fprintf fmt "let coqpp_body %a%a ~st = @[%a@] in "
+ print_binders r.vernac_toks print_atts_left r.vernac_atts print_body_wrapper r
let print_body fmt r =
fprintf fmt "@[(%afun %a~atts@ ~st@ -> coqpp_body %a%a ~st)@]"
diff --git a/coqpp/coqpp_parse.mly b/coqpp/coqpp_parse.mly
index abe52ab46b..43ba990f6a 100644
--- a/coqpp/coqpp_parse.mly
+++ b/coqpp/coqpp_parse.mly
@@ -65,7 +65,7 @@ let parse_user_entry s sep =
%token VERNAC TACTIC GRAMMAR EXTEND END DECLARE PLUGIN DEPRECATED ARGUMENT
%token RAW_PRINTED GLOB_PRINTED
%token COMMAND CLASSIFIED PRINTED TYPED INTERPRETED GLOBALIZED SUBSTITUTED BY AS
-%token HASHBRACKET LBRACKET RBRACKET PIPE ARROW FUN COMMA EQUAL STAR
+%token BANGBRACKET HASHBRACKET LBRACKET RBRACKET PIPE ARROW FUN COMMA EQUAL STAR
%token LPAREN RPAREN COLON SEMICOLON
%token GLOBAL FIRST LAST BEFORE AFTER LEVEL LEFTA RIGHTA NONA
%token EOF
@@ -209,13 +209,14 @@ vernac_rules:
;
vernac_rule:
-| PIPE vernac_attributes_opt LBRACKET ext_tokens RBRACKET rule_deprecation rule_classifier ARROW CODE
+| PIPE vernac_attributes_opt vernac_state_opt LBRACKET ext_tokens RBRACKET rule_deprecation rule_classifier ARROW CODE
{ {
vernac_atts = $2;
- vernac_toks = $4;
- vernac_depr = $6;
- vernac_class= $7;
- vernac_body = $9;
+ vernac_state= $3;
+ vernac_toks = $5;
+ vernac_depr = $7;
+ vernac_class= $8;
+ vernac_body = $10;
} }
;
@@ -235,6 +236,14 @@ vernac_attribute:
| qualid_or_ident { ($1, $1) }
;
+vernac_state_opt:
+| { None }
+| BANGBRACKET vernac_state RBRACKET { Some $2 }
+;
+
+vernac_state:
+| qualid_or_ident { $1 }
+
rule_deprecation:
| { false }
| DEPRECATED { true }
diff --git a/dev/ci/user-overlays/09129-ejgallego-proof+no_global_partial.sh b/dev/ci/user-overlays/09129-ejgallego-proof+no_global_partial.sh
new file mode 100644
index 0000000000..c09d1b8929
--- /dev/null
+++ b/dev/ci/user-overlays/09129-ejgallego-proof+no_global_partial.sh
@@ -0,0 +1,30 @@
+if [ "$CI_PULL_REQUEST" = "9129" ] || [ "$CI_BRANCH" = "proof+no_global_partial" ]; then
+
+ aac_tactics_CI_REF=proof+no_global_partial
+ aac_tactics_CI_GITURL=https://github.com/ejgallego/aac-tactics
+
+ # coqhammer_CI_REF=proof+no_global_partial
+ # coqhammer_CI_GITURL=https://github.com/ejgallego/coqhammer
+
+ elpi_CI_REF=proof+no_global_partial
+ elpi_CI_GITURL=https://github.com/ejgallego/coq-elpi
+
+ equations_CI_REF=proof+no_global_partial
+ equations_CI_GITURL=https://github.com/ejgallego/Coq-Equations
+
+ ltac2_CI_REF=proof+no_global_partial
+ ltac2_CI_GITURL=https://github.com/ejgallego/ltac2
+
+ # unicoq_CI_REF=proof+no_global_partial
+ # unicoq_CI_GITURL=https://github.com/ejgallego/unicoq
+
+ mtac2_CI_REF=proof+no_global_partial
+ mtac2_CI_GITURL=https://github.com/ejgallego/Mtac2
+
+ paramcoq_CI_REF=proof+no_global_partial
+ paramcoq_CI_GITURL=https://github.com/ejgallego/paramcoq
+
+ quickchick_CI_REF=proof+no_global_partial
+ quickchick_CI_GITURL=https://github.com/ejgallego/QuickChick
+
+fi
diff --git a/dev/top_printers.ml b/dev/top_printers.ml
index 499bbba37e..1f4f2246be 100644
--- a/dev/top_printers.ml
+++ b/dev/top_printers.ml
@@ -60,19 +60,21 @@ let prrecarg = function
str "Imbr[" ++ MutInd.print mind ++ pr_comma () ++ int i ++ str "]"
let ppwf_paths x = pp (Rtree.pp_tree prrecarg x)
+let get_current_context = Vernacstate.Proof_global.get_current_context
+
(* term printers *)
-let envpp pp = let sigma,env = Pfedit.get_current_context () in pp env sigma
+let envpp pp = let sigma,env = get_current_context () in pp env sigma
let rawdebug = ref false
let ppevar evk = pp (Evar.print evk)
let pr_constr t =
- let sigma, env = Pfedit.get_current_context () in
+ let sigma, env = get_current_context () in
Printer.pr_constr_env env sigma t
let pr_econstr t =
- let sigma, env = Pfedit.get_current_context () in
+ let sigma, env = get_current_context () in
Printer.pr_econstr_env env sigma t
let ppconstr x = pp (pr_constr x)
let ppeconstr x = pp (pr_econstr x)
-let ppconstr_expr x = let sigma,env = Pfedit.get_current_context () in pp (Ppconstr.pr_constr_expr env sigma x)
+let ppconstr_expr x = let sigma,env = get_current_context () in pp (Ppconstr.pr_constr_expr env sigma x)
let ppsconstr x = ppconstr (Mod_subst.force_constr x)
let ppconstr_univ x = Constrextern.with_universes ppconstr x
let ppglob_constr = (fun x -> pp(pr_lglob_constr_env (Global.env()) x))
@@ -500,7 +502,7 @@ let ppist ist =
(* Vernac-level debugging commands *)
let in_current_context f c =
- let (evmap,sign) = Pfedit.get_current_context () in
+ let (evmap,sign) = get_current_context () in
f (fst (Constrintern.interp_constr sign evmap c))(*FIXME*)
(* We expand the result of preprocessing to be independent of camlp5
diff --git a/doc/plugin_tutorial/tuto1/src/g_tuto1.mlg b/doc/plugin_tutorial/tuto1/src/g_tuto1.mlg
index 4df284d2d9..1d0aca1caf 100644
--- a/doc/plugin_tutorial/tuto1/src/g_tuto1.mlg
+++ b/doc/plugin_tutorial/tuto1/src/g_tuto1.mlg
@@ -145,10 +145,12 @@ END
it gives an error message that is basically impossible to understand. *)
VERNAC COMMAND EXTEND ExploreProof CLASSIFIED AS QUERY
-| [ "Cmd9" ] ->
- { let p = Proof_global.give_me_the_proof () in
- let sigma, env = Pfedit.get_current_context () in
- let pprf = Proof.partial_proof p in
- Feedback.msg_notice
- (Pp.prlist_with_sep Pp.fnl (Printer.pr_econstr_env env sigma) pprf) }
+| ![ proof ] [ "Cmd9" ] ->
+ { fun ~pstate ->
+ Option.iter (fun (pstate : Proof_global.t) ->
+ let sigma, env = Pfedit.get_current_context pstate in
+ let pprf = Proof.partial_proof Proof_global.(give_me_the_proof pstate) in
+ Feedback.msg_notice
+ (Pp.prlist_with_sep Pp.fnl (Printer.pr_econstr_env env sigma) pprf)) pstate;
+ pstate }
END
diff --git a/doc/plugin_tutorial/tuto1/src/simple_declare.ml b/doc/plugin_tutorial/tuto1/src/simple_declare.ml
index e370d37fc4..23f8fbe888 100644
--- a/doc/plugin_tutorial/tuto1/src/simple_declare.ml
+++ b/doc/plugin_tutorial/tuto1/src/simple_declare.ml
@@ -1,5 +1,5 @@
(* Ideally coq/coq#8811 would get merged and then this function could be much simpler. *)
-let edeclare ?hook ident (_, poly, _ as k) ~opaque sigma udecl body tyopt imps =
+let edeclare ?hook ~ontop ident (_, poly, _ as k) ~opaque sigma udecl body tyopt imps =
let sigma = Evd.minimize_universes sigma in
let body = EConstr.to_constr sigma body in
let tyopt = Option.map (EConstr.to_constr sigma) tyopt in
@@ -13,13 +13,13 @@ let edeclare ?hook ident (_, poly, _ as k) ~opaque sigma udecl body tyopt imps =
let ubinders = Evd.universe_binders sigma in
let ce = Declare.definition_entry ?types:tyopt ~univs body in
let hook_data = Option.map (fun hook -> hook, uctx, []) hook in
- DeclareDef.declare_definition ident k ce ubinders imps ?hook_data
+ DeclareDef.declare_definition ~ontop ident k ce ubinders imps ?hook_data
let packed_declare_definition ~poly ident value_with_constraints =
let body, ctx = value_with_constraints in
let sigma = Evd.from_ctx ctx in
let k = (Decl_kinds.Global, poly, Decl_kinds.Definition) in
let udecl = UState.default_univ_decl in
- ignore (edeclare ident k ~opaque:false sigma udecl body None [])
+ ignore (edeclare ~ontop:None ident k ~opaque:false sigma udecl body None [])
(* But this definition cannot be undone by Reset ident *)
diff --git a/doc/sphinx/proof-engine/ltac.rst b/doc/sphinx/proof-engine/ltac.rst
index 52e3029b8f..0322b43694 100644
--- a/doc/sphinx/proof-engine/ltac.rst
+++ b/doc/sphinx/proof-engine/ltac.rst
@@ -1071,6 +1071,16 @@ Proving a subgoal as a separate lemma
It may be useful to generate lemmas minimal w.r.t. the assumptions they
depend on. This can be obtained thanks to the option below.
+ .. warning::
+
+ The abstract tactic, while very useful, still has some known
+ limitations, see https://github.com/coq/coq/issues/9146 for more
+ details. Thus we recommend using it caution in some
+ "non-standard" contexts. In particular, ``abstract`` won't
+ properly work when used inside quotations ``ltac:(...)``, or
+ if used as part of typeclass resolution, it may produce wrong
+ terms when in universe polymorphic mode.
+
.. tacv:: abstract @expr using @ident
Give explicitly the name of the auxiliary lemma.
diff --git a/dune b/dune
index f1f966b7fd..787c3c3674 100644
--- a/dune
+++ b/dune
@@ -42,3 +42,5 @@
(name runtest)
(package coqide-server)
(deps test-suite/summary.log))
+
+; (dirs (:standard _build_ci))
diff --git a/ide/idetop.ml b/ide/idetop.ml
index 608577b297..f744ce2ee3 100644
--- a/ide/idetop.ml
+++ b/ide/idetop.ml
@@ -231,30 +231,30 @@ let goals () =
let doc = get_doc () in
set_doc @@ Stm.finish ~doc;
try
- let newp = Proof_global.give_me_the_proof () in
+ let newp = Vernacstate.Proof_global.give_me_the_proof () in
if Proof_diffs.show_diffs () then begin
let oldp = Stm.get_prev_proof ~doc (Stm.get_current_state ~doc) in
let diff_goal_map = Proof_diffs.make_goal_map oldp newp in
Some (export_pre_goals Proof.(data newp) (process_goal_diffs diff_goal_map oldp))
end else
Some (export_pre_goals Proof.(data newp) process_goal)
- with Proof_global.NoCurrentProof -> None;;
+ with Vernacstate.Proof_global.NoCurrentProof -> None;;
let evars () =
try
let doc = get_doc () in
set_doc @@ Stm.finish ~doc;
- let pfts = Proof_global.give_me_the_proof () in
+ let pfts = Vernacstate.Proof_global.give_me_the_proof () in
let Proof.{ sigma } = Proof.data pfts in
let exl = Evar.Map.bindings (Evd.undefined_map sigma) in
let map_evar ev = { Interface.evar_info = string_of_ppcmds (pr_evar sigma ev); } in
let el = List.map map_evar exl in
Some el
- with Proof_global.NoCurrentProof -> None
+ with Vernacstate.Proof_global.NoCurrentProof -> None
let hints () =
try
- let pfts = Proof_global.give_me_the_proof () in
+ let pfts = Vernacstate.Proof_global.give_me_the_proof () in
let Proof.{ goals; sigma } = Proof.data pfts in
match goals with
| [] -> None
@@ -263,7 +263,7 @@ let hints () =
let get_hint_hyp env d accu = hyp_next_tac sigma env d :: accu in
let hint_hyps = List.rev (Environ.fold_named_context get_hint_hyp env ~init: []) in
Some (hint_hyps, concl_next_tac)
- with Proof_global.NoCurrentProof -> None
+ with Vernacstate.Proof_global.NoCurrentProof -> None
(** Other API calls *)
@@ -284,11 +284,11 @@ let status force =
List.rev_map Names.Id.to_string l
in
let proof =
- try Some (Names.Id.to_string (Proof_global.get_current_proof_name ()))
- with Proof_global.NoCurrentProof -> None
+ try Some (Names.Id.to_string (Vernacstate.Proof_global.get_current_proof_name ()))
+ with Vernacstate.Proof_global.NoCurrentProof -> None
in
let allproofs =
- let l = Proof_global.get_all_proof_names () in
+ let l = Vernacstate.Proof_global.get_all_proof_names () in
List.map Names.Id.to_string l
in
{
@@ -336,7 +336,8 @@ let import_search_constraint = function
| Interface.Include_Blacklist -> Search.Include_Blacklist
let search flags =
- List.map export_coq_object (Search.interface_search (
+ let pstate = Vernacstate.Proof_global.get () in
+ List.map export_coq_object (Search.interface_search ?pstate (
List.map (fun (c, b) -> (import_search_constraint c, b)) flags)
)
diff --git a/plugins/cc/ccalgo.mli b/plugins/cc/ccalgo.mli
index 978969bf59..5066c3931d 100644
--- a/plugins/cc/ccalgo.mli
+++ b/plugins/cc/ccalgo.mli
@@ -255,5 +255,3 @@ val find_contradiction : UF.t ->
(Names.Id.t * (int * int)) list ->
(Names.Id.t * (int * int))
*)
-
-
diff --git a/plugins/derive/derive.ml b/plugins/derive/derive.ml
index afdbfa1999..4425e41652 100644
--- a/plugins/derive/derive.ml
+++ b/plugins/derive/derive.ml
@@ -101,8 +101,7 @@ let start_deriving f suchthat lemma =
in
let terminator = Proof_global.make_terminator terminator in
- let () = Proof_global.start_dependent_proof lemma kind goals terminator in
- let _ = Proof_global.with_current_proof begin fun _ p ->
+ let pstate = Proof_global.start_dependent_proof ~ontop:None lemma kind goals terminator in
+ fst @@ Proof_global.with_current_proof begin fun _ p ->
Proof.run_tactic env Proofview.(tclFOCUS 1 2 shelve) p
- end in
- ()
+ end pstate
diff --git a/plugins/derive/derive.mli b/plugins/derive/derive.mli
index 06ff9c48cf..6bb923118e 100644
--- a/plugins/derive/derive.mli
+++ b/plugins/derive/derive.mli
@@ -12,4 +12,4 @@
(which can contain references to [f]) in the context extended by
[f:=?x]. When the proof ends, [f] is defined as the value of [?x]
and [lemma] as the proof. *)
-val start_deriving : Names.Id.t -> Constrexpr.constr_expr -> Names.Id.t -> unit
+val start_deriving : Names.Id.t -> Constrexpr.constr_expr -> Names.Id.t -> Proof_global.t
diff --git a/plugins/derive/g_derive.mlg b/plugins/derive/g_derive.mlg
index 0cdf8fb5d8..214a9d8bb5 100644
--- a/plugins/derive/g_derive.mlg
+++ b/plugins/derive/g_derive.mlg
@@ -23,6 +23,6 @@ let classify_derive_command _ = Vernacextend.(VtStartProof (Doesn'tGuaranteeOpac
}
VERNAC COMMAND EXTEND Derive CLASSIFIED BY { classify_derive_command }
-| [ "Derive" ident(f) "SuchThat" constr(suchthat) "As" ident(lemma) ] ->
- { Derive.start_deriving f suchthat lemma }
+| ![ proof ] [ "Derive" ident(f) "SuchThat" constr(suchthat) "As" ident(lemma) ] ->
+ { fun ~pstate -> Some Derive.(start_deriving f suchthat lemma) }
END
diff --git a/plugins/extraction/extract_env.ml b/plugins/extraction/extract_env.ml
index 0fa9be21c9..8f17f7b2dd 100644
--- a/plugins/extraction/extract_env.ml
+++ b/plugins/extraction/extract_env.ml
@@ -750,16 +750,19 @@ let extract_and_compile l =
Feedback.msg_notice (str "Extracted code successfully compiled")
(* Show the extraction of the current ongoing proof *)
-
-let show_extraction () =
+let show_extraction ~pstate =
+ let pstate = match pstate with
+ | None -> CErrors.user_err Pp.(str "No ongoing proof")
+ | Some pstate -> pstate
+ in
init ~inner:true false false;
- let prf = Proof_global.give_me_the_proof () in
- let sigma, env = Pfedit.get_current_context () in
+ let prf = Proof_global.give_me_the_proof pstate in
+ let sigma, env = Pfedit.get_current_context pstate in
let trms = Proof.partial_proof prf in
let extr_term t =
let ast, ty = extract_constr env sigma t in
let mp = Lib.current_mp () in
- let l = Label.of_id (Proof_global.get_current_proof_name ()) in
+ let l = Label.of_id (Proof_global.get_current_proof_name pstate) in
let fake_ref = ConstRef (Constant.make2 mp l) in
let decl = Dterm (fake_ref, ast, ty) in
print_one_decl [] mp decl
diff --git a/plugins/extraction/extract_env.mli b/plugins/extraction/extract_env.mli
index 54fde2ca46..7ba7e05019 100644
--- a/plugins/extraction/extract_env.mli
+++ b/plugins/extraction/extract_env.mli
@@ -40,4 +40,4 @@ val structure_for_compute :
(* Show the extraction of the current ongoing proof *)
-val show_extraction : unit -> unit
+val show_extraction : pstate:Proof_global.t option -> unit
diff --git a/plugins/extraction/g_extraction.mlg b/plugins/extraction/g_extraction.mlg
index 1445dffefa..d7bb27f121 100644
--- a/plugins/extraction/g_extraction.mlg
+++ b/plugins/extraction/g_extraction.mlg
@@ -178,6 +178,6 @@ END
(* Show the extraction of the current proof *)
VERNAC COMMAND EXTEND ShowExtraction CLASSIFIED AS QUERY
-| [ "Show" "Extraction" ]
- -> { show_extraction () }
+| ![ proof ] [ "Show" "Extraction" ]
+ -> { fun ~pstate -> let () = show_extraction ~pstate in pstate }
END
diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml
index 16f376931e..287a374ab1 100644
--- a/plugins/funind/functional_principles_proofs.ml
+++ b/plugins/funind/functional_principles_proofs.ml
@@ -722,7 +722,7 @@ let build_proof
(treat_new_case
ptes_infos
nb_instantiate_partial
- (build_proof env sigma do_finalize)
+ (build_proof do_finalize)
t
dyn_infos)
g'
@@ -733,7 +733,7 @@ let build_proof
]
g
in
- build_proof env sigma do_finalize_t {dyn_infos with info = t} g
+ build_proof do_finalize_t {dyn_infos with info = t} g
| Lambda(n,t,b) ->
begin
match EConstr.kind sigma (pf_concl g) with
@@ -749,7 +749,7 @@ let build_proof
in
let new_infos = {dyn_infos with info = new_term} in
let do_prove new_hyps =
- build_proof env sigma do_finalize
+ build_proof do_finalize
{new_infos with
rec_hyps = new_hyps;
nb_rec_hyps = List.length new_hyps
@@ -762,7 +762,7 @@ let build_proof
do_finalize dyn_infos g
end
| Cast(t,_,_) ->
- build_proof env sigma do_finalize {dyn_infos with info = t} g
+ build_proof do_finalize {dyn_infos with info = t} g
| Const _ | Var _ | Meta _ | Evar _ | Sort _ | Construct _ | Ind _ | Int _ ->
do_finalize dyn_infos g
| App(_,_) ->
@@ -792,7 +792,7 @@ let build_proof
| Lambda _ ->
let new_term =
Reductionops.nf_beta env sigma dyn_infos.info in
- build_proof env sigma do_finalize {dyn_infos with info = new_term}
+ build_proof do_finalize {dyn_infos with info = new_term}
g
| LetIn _ ->
let new_infos =
@@ -805,11 +805,11 @@ let build_proof
h_reduce_with_zeta (Locusops.onHyp hyp_id))
dyn_infos.rec_hyps;
h_reduce_with_zeta Locusops.onConcl;
- build_proof env sigma do_finalize new_infos
+ build_proof do_finalize new_infos
]
g
| Cast(b,_,_) ->
- build_proof env sigma do_finalize {dyn_infos with info = b } g
+ build_proof do_finalize {dyn_infos with info = b } g
| Case _ | Fix _ | CoFix _ ->
let new_finalize dyn_infos =
let new_infos =
@@ -819,7 +819,7 @@ let build_proof
in
build_proof_args env sigma do_finalize new_infos
in
- build_proof env sigma new_finalize {dyn_infos with info = f } g
+ build_proof new_finalize {dyn_infos with info = f } g
end
| Fix _ | CoFix _ ->
user_err Pp.(str ( "Anonymous local (co)fixpoints are not handled yet"))
@@ -839,12 +839,12 @@ let build_proof
(fun hyp_id -> h_reduce_with_zeta (Locusops.onHyp hyp_id))
dyn_infos.rec_hyps;
h_reduce_with_zeta Locusops.onConcl;
- build_proof env sigma do_finalize new_infos
+ build_proof do_finalize new_infos
] g
| Rel _ -> anomaly (Pp.str "Free var in goal conclusion!")
- and build_proof env sigma do_finalize dyn_infos g =
+ and build_proof do_finalize dyn_infos g =
(* observe (str "proving with "++Printer.pr_lconstr dyn_infos.info++ str " on goal " ++ pr_gls g); *)
- observe_tac_stream (str "build_proof with " ++ pr_leconstr_env env sigma dyn_infos.info ) (build_proof_aux do_finalize dyn_infos) g
+ observe_tac_stream (str "build_proof with " ++ pr_leconstr_env (pf_env g) (project g) dyn_infos.info ) (build_proof_aux do_finalize dyn_infos) g
and build_proof_args env sigma do_finalize dyn_infos (* f_args' args *) :tactic =
fun g ->
let (f_args',args) = dyn_infos.info in
@@ -866,7 +866,7 @@ let build_proof
{dyn_infos with info = (mkApp(f_args',[|new_arg|])), args}
)
in
- build_proof env sigma do_finalize
+ build_proof do_finalize
{dyn_infos with info = arg }
g
in
@@ -879,19 +879,7 @@ let build_proof
in
(* observe_tac "build_proof" *)
fun g ->
- let env = pf_env g in
- let sigma = project g in
- build_proof env sigma (clean_goal_with_heq ptes_infos do_finish_proof) dyn_infos g
-
-
-
-
-
-
-
-
-
-
+ build_proof (clean_goal_with_heq ptes_infos do_finish_proof) dyn_infos g
(* Proof of principles from structural functions *)
@@ -1002,19 +990,18 @@ let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num
]
in
(* Pp.msgnl (str "lemma type (2) " ++ Printer.pr_lconstr_env (Global.env ()) evd lemma_type); *)
- Lemmas.start_proof
+ let pstate = Lemmas.start_proof ~ontop:None
(*i The next call to mk_equation_id is valid since we are constructing the lemma
Ensures by: obvious
i*)
(mk_equation_id f_id)
(Decl_kinds.Global, false, (Decl_kinds.Proof Decl_kinds.Theorem))
evd
- lemma_type;
- ignore (Pfedit.by (Proofview.V82.tactic prove_replacement));
- Lemmas.save_proof (Vernacexpr.(Proved(Proof_global.Transparent,None)));
- evd
-
-
+ lemma_type
+ in
+ let pstate,_ = Pfedit.by (Proofview.V82.tactic prove_replacement) pstate in
+ let pstate = Lemmas.save_proof_proved ?proof:None ~pstate ~opaque:Proof_global.Transparent ~idopt:None in
+ pstate, evd
let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num all_funs g =
@@ -1028,7 +1015,7 @@ let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num a
Ensures by: obvious
i*)
let equation_lemma_id = (mk_equation_id f_id) in
- evd := generate_equation_lemma !evd all_funs f fun_num (List.length params) (List.length rev_args_id) rec_arg_num;
+ evd := snd @@ generate_equation_lemma !evd all_funs f fun_num (List.length params) (List.length rev_args_id) rec_arg_num;
let _ =
match e with
| Option.IsNone ->
diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml
index 1217ba0eba..e9a2c285d0 100644
--- a/plugins/funind/functional_principles_types.ml
+++ b/plugins/funind/functional_principles_types.ml
@@ -308,31 +308,30 @@ let build_functional_principle (evd:Evd.evar_map ref) interactive_proof old_prin
let sigma, _ = Typing.type_of ~refresh:true (Global.env ()) !evd (EConstr.of_constr new_principle_type) in
evd := sigma;
let hook = Lemmas.mk_hook (hook new_principle_type) in
- begin
- Lemmas.start_proof
+ let pstate =
+ Lemmas.start_proof ~ontop:None
new_princ_name
(Decl_kinds.Global,false,(Decl_kinds.Proof Decl_kinds.Theorem))
!evd
(EConstr.of_constr new_principle_type)
- ;
- (* let _tim1 = System.get_time () in *)
- let map (c, u) = EConstr.mkConstU (c, EConstr.EInstance.make u) in
- ignore (Pfedit.by (Proofview.V82.tactic (proof_tac (Array.map map funs) mutr_nparams)));
- (* let _tim2 = System.get_time () in *)
- (* begin *)
- (* let dur1 = System.time_difference tim1 tim2 in *)
- (* Pp.msgnl (str ("Time to compute proof: ") ++ str (string_of_float dur1)); *)
- (* end; *)
+ in
+ (* let _tim1 = System.get_time () in *)
+ let map (c, u) = EConstr.mkConstU (c, EConstr.EInstance.make u) in
+ let pstate,_ = Pfedit.by (Proofview.V82.tactic (proof_tac (Array.map map funs) mutr_nparams)) pstate in
+ (* let _tim2 = System.get_time () in *)
+ (* begin *)
+ (* let dur1 = System.time_difference tim1 tim2 in *)
+ (* Pp.msgnl (str ("Time to compute proof: ") ++ str (string_of_float dur1)); *)
+ (* end; *)
- let open Proof_global in
- let { id; entries; persistence } = fst @@ close_proof ~opaque:Transparent ~keep_body_ucst_separate:false (fun x -> x) in
- match entries with
- | [entry] ->
- discard_current ();
- (id,(entry,persistence)), hook
- | _ ->
- CErrors.anomaly Pp.(str "[build_functional_principle] close_proof returned more than one proof term")
- end
+ let open Proof_global in
+ let { id; entries; persistence } = fst @@ close_proof ~opaque:Transparent ~keep_body_ucst_separate:false (fun x -> x) pstate in
+ match entries with
+ | [entry] ->
+ let pstate = discard_current pstate in
+ (id,(entry,persistence)), hook, pstate
+ | _ ->
+ CErrors.anomaly Pp.(str "[build_functional_principle] close_proof returned more than one proof term")
let generate_functional_principle (evd: Evd.evar_map ref)
interactive_proof
@@ -382,7 +381,7 @@ let generate_functional_principle (evd: Evd.evar_map ref)
register_with_sort InProp;
register_with_sort InSet
in
- let ((id,(entry,g_kind)),hook) =
+ let ((id,(entry,g_kind)),hook,pstate) =
build_functional_principle evd interactive_proof old_princ_type new_sorts funs i
proof_tac hook
in
@@ -390,25 +389,9 @@ let generate_functional_principle (evd: Evd.evar_map ref)
Don't forget to close the goal if an error is raised !!!!
*)
let uctx = Evd.evar_universe_context sigma in
- save false new_princ_name entry ~hook uctx g_kind
+ save new_princ_name entry ~hook uctx g_kind
with e when CErrors.noncritical e ->
- begin
- begin
- try
- let id = Proof_global.get_current_proof_name () in
- let s = Id.to_string id in
- let n = String.length "___________princ_________" in
- if String.length s >= n
- then if String.equal (String.sub s 0 n) "___________princ_________"
- then Proof_global.discard_current ()
- else ()
- else ()
- with e when CErrors.noncritical e -> ()
- end;
- raise (Defining_principle e)
- end
-(* defined () *)
-
+ raise (Defining_principle e)
exception Not_Rec
@@ -537,7 +520,7 @@ let make_scheme evd (fas : (pconstant*Sorts.family) list) : Safe_typing.private_
s::l_schemes -> s,l_schemes
| _ -> anomaly (Pp.str "")
in
- let ((_,(const,_)),_) =
+ let ((_,(const,_)),_,pstate) =
try
build_functional_principle evd false
first_type
@@ -547,21 +530,7 @@ let make_scheme evd (fas : (pconstant*Sorts.family) list) : Safe_typing.private_
(prove_princ_for_struct evd false 0 (Array.of_list (List.map fst funs)))
(fun _ _ _ _ _ -> ())
with e when CErrors.noncritical e ->
- begin
- begin
- try
- let id = Proof_global.get_current_proof_name () in
- let s = Id.to_string id in
- let n = String.length "___________princ_________" in
- if String.length s >= n
- then if String.equal (String.sub s 0 n) "___________princ_________"
- then Proof_global.discard_current ()
- else ()
- else ()
- with e when CErrors.noncritical e -> ()
- end;
- raise (Defining_principle e)
- end
+ raise (Defining_principle e)
in
incr i;
@@ -611,7 +580,7 @@ let make_scheme evd (fas : (pconstant*Sorts.family) list) : Safe_typing.private_
(* If we reach this point, the two principle are not mutually recursive
We fall back to the previous method
*)
- let ((_,(const,_)),_) =
+ let ((_,(const,_)),_,pstate) =
build_functional_principle
evd
false
diff --git a/plugins/funind/g_indfun.mlg b/plugins/funind/g_indfun.mlg
index 6f67ab4d8b..4e8cf80ed2 100644
--- a/plugins/funind/g_indfun.mlg
+++ b/plugins/funind/g_indfun.mlg
@@ -177,7 +177,7 @@ let () =
(* TASSI: n'importe quoi ! *)
VERNAC COMMAND EXTEND Function
-| ["Function" ne_function_rec_definition_loc_list_sep(recsl,"with")]
+| ![ proof ] ["Function" ne_function_rec_definition_loc_list_sep(recsl,"with")]
=> { let hard = List.exists (function
| _,((_,(_,(CMeasureRec _|CWfRec _)),_,_,_),_) -> true
| _,((_,(_,CStructRec),_,_,_),_) -> false) recsl in
@@ -223,37 +223,34 @@ let warning_error names e =
}
VERNAC COMMAND EXTEND NewFunctionalScheme
-| ["Functional" "Scheme" ne_fun_scheme_arg_list_sep(fas,"with") ]
+| ![ proof ] ["Functional" "Scheme" ne_fun_scheme_arg_list_sep(fas,"with") ]
=> { Vernacextend.(VtSideff(List.map pi1 fas), VtLater) }
->
- {
+ { fun ~pstate ->
begin
- try
- Functional_principles_types.build_scheme fas
- with Functional_principles_types.No_graph_found ->
- begin
- match fas with
- | (_,fun_name,_)::_ ->
- begin
- begin
- make_graph (Smartlocate.global_with_alias fun_name)
- end
- ;
- try Functional_principles_types.build_scheme fas
- with Functional_principles_types.No_graph_found ->
- CErrors.user_err Pp.(str "Cannot generate induction principle(s)")
- | e when CErrors.noncritical e ->
- let names = List.map (fun (_,na,_) -> na) fas in
- warning_error names e
-
- end
+ try
+ Functional_principles_types.build_scheme fas; pstate
+ with
+ | Functional_principles_types.No_graph_found ->
+ begin
+ match fas with
+ | (_,fun_name,_)::_ ->
+ begin
+ let pstate = make_graph ~pstate (Smartlocate.global_with_alias fun_name) in
+ try Functional_principles_types.build_scheme fas; pstate
+ with
+ | Functional_principles_types.No_graph_found ->
+ CErrors.user_err Pp.(str "Cannot generate induction principle(s)")
+ | e when CErrors.noncritical e ->
+ let names = List.map (fun (_,na,_) -> na) fas in
+ warning_error names e; pstate
+ end
| _ -> assert false (* we can only have non empty list *)
- end
- | e when CErrors.noncritical e ->
- let names = List.map (fun (_,na,_) -> na) fas in
- warning_error names e
+ end
+ | e when CErrors.noncritical e ->
+ let names = List.map (fun (_,na,_) -> na) fas in
+ warning_error names e; pstate
end
-
}
END
(***** debug only ***)
@@ -266,5 +263,6 @@ END
(***** debug only ***)
VERNAC COMMAND EXTEND GenerateGraph CLASSIFIED AS QUERY
-| ["Generate" "graph" "for" reference(c)] -> { make_graph (Smartlocate.global_with_alias c) }
+| ![ proof ] ["Generate" "graph" "for" reference(c)] ->
+ { make_graph (Smartlocate.global_with_alias c) }
END
diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml
index f4807954a7..275b58f0aa 100644
--- a/plugins/funind/glob_term_to_relation.ml
+++ b/plugins/funind/glob_term_to_relation.ml
@@ -369,7 +369,7 @@ let add_pat_variables sigma pat typ env : Environ.env =
let cs_args_types :types list = List.map RelDecl.get_type constructor.Inductiveops.cs_args in
List.fold_left2 add_pat_variables env patl (List.rev cs_args_types)
in
- let new_env = add_pat_variables env pat typ in
+ let new_env = add_pat_variables env pat typ in
let res =
fst (
Context.Rel.fold_outside
diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml
index b69ca7080c..a5c19f3217 100644
--- a/plugins/funind/indfun.ml
+++ b/plugins/funind/indfun.ml
@@ -410,11 +410,11 @@ let generate_principle (evd:Evd.evar_map ref) pconstants on_error
with e when CErrors.noncritical e ->
on_error names e
-let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list) =
+let register_struct ~pstate is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list) =
match fixpoint_exprl with
| [(({CAst.v=fname},pl),_,bl,ret_type,body),_] when not is_rec ->
let body = match body with | Some body -> body | None -> user_err ~hdr:"Function" (str "Body of Function must be given") in
- ComDefinition.do_definition
+ ComDefinition.do_definition ~ontop:pstate
~program_mode:false
fname
(Decl_kinds.Global,false,Decl_kinds.Definition) pl
@@ -432,9 +432,9 @@ let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexp
(Evd.from_env (Global.env ()),[])
fixpoint_exprl
in
- evd,List.rev rev_pconstants
+ pstate, evd,List.rev rev_pconstants
| _ ->
- ComFixpoint.do_fixpoint Global false fixpoint_exprl;
+ let pstate = ComFixpoint.do_fixpoint ~ontop:pstate Global false fixpoint_exprl in
let evd,rev_pconstants =
List.fold_left
(fun (evd,l) ((({CAst.v=fname},_),_,_,_,_),_) ->
@@ -448,8 +448,8 @@ let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexp
(Evd.from_env (Global.env ()),[])
fixpoint_exprl
in
- evd,List.rev rev_pconstants
-
+ pstate,evd,List.rev rev_pconstants
+
let generate_correction_proof_wf f_ref tcc_lemma_ref
is_mes functional_ref eq_ref rec_arg_num rec_arg_type nb_args relation
@@ -638,10 +638,10 @@ let recompute_binder_list (fixpoint_exprl : (Vernacexpr.fixpoint_expr * Vernacex
fixpoint_exprl_with_new_bl
-let do_generate_principle pconstants on_error register_built interactive_proof
- (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list) :unit =
+let do_generate_principle ~pstate pconstants on_error register_built interactive_proof
+ (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list) : Proof_global.t option =
List.iter (fun (_,l) -> if not (List.is_empty l) then error "Function does not support notations for now") fixpoint_exprl;
- let _is_struct =
+ let pstate, _is_struct =
match fixpoint_exprl with
| [((_,(wf_x,Constrexpr.CWfRec wf_rel),_,_,_),_) as fixpoint_expr] ->
let (((({CAst.v=name},pl),_,args,types,body)),_) as fixpoint_expr =
@@ -665,8 +665,8 @@ let do_generate_principle pconstants on_error register_built interactive_proof
true
in
if register_built
- then register_wf name rec_impls wf_rel (map_option (fun x -> x.CAst.v) wf_x) using_lemmas args types body pre_hook;
- false
+ then register_wf name rec_impls wf_rel (map_option (fun x -> x.CAst.v) wf_x) using_lemmas args types body pre_hook, false
+ else pstate, false
|[((_,(wf_x,Constrexpr.CMeasureRec(wf_mes,wf_rel_opt)),_,_,_),_) as fixpoint_expr] ->
let (((({CAst.v=name},_),_,args,types,body)),_) as fixpoint_expr =
match recompute_binder_list [fixpoint_expr] with
@@ -689,8 +689,8 @@ let do_generate_principle pconstants on_error register_built interactive_proof
true
in
if register_built
- then register_mes name rec_impls wf_mes wf_rel_opt (map_option (fun x -> x.CAst.v) wf_x) using_lemmas args types body pre_hook;
- true
+ then register_mes name rec_impls wf_mes wf_rel_opt (map_option (fun x -> x.CAst.v) wf_x) using_lemmas args types body pre_hook, true
+ else pstate, true
| _ ->
List.iter (function ((_na,(_,ord),_args,_body,_type),_not) ->
match ord with
@@ -707,10 +707,10 @@ let do_generate_principle pconstants on_error register_built interactive_proof
(* ok all the expressions are structural *)
let recdefs,rec_impls = build_newrecursive fixpoint_exprl in
let is_rec = List.exists (is_rec fix_names) recdefs in
- let evd,pconstants =
+ let pstate,evd,pconstants =
if register_built
- then register_struct is_rec fixpoint_exprl
- else (Evd.from_env (Global.env ()),pconstants)
+ then register_struct ~pstate is_rec fixpoint_exprl
+ else pstate, Evd.from_env (Global.env ()), pconstants
in
let evd = ref evd in
generate_principle
@@ -723,10 +723,11 @@ let do_generate_principle pconstants on_error register_built interactive_proof
recdefs
interactive_proof
(Functional_principles_proofs.prove_princ_for_struct evd interactive_proof);
- if register_built then begin derive_inversion fix_names; end;
- true;
+ if register_built then
+ begin derive_inversion fix_names; end;
+ pstate, true
in
- ()
+ pstate
let rec add_args id new_args = CAst.map (function
| CRef (qid,_) as b ->
@@ -843,13 +844,14 @@ let rec get_args b t : Constrexpr.local_binder_expr list *
| _ -> [],b,t
-let make_graph (f_ref : GlobRef.t) =
+let make_graph ~pstate (f_ref : GlobRef.t) =
+ let sigma, env = Option.cata Pfedit.get_current_context
+ (let e = Global.env () in Evd.from_env e, e) pstate in
let c,c_body =
match f_ref with
| ConstRef c ->
begin try c,Global.lookup_constant c
with Not_found ->
- let sigma, env = Pfedit.get_current_context () in
raise (UserError (None,str "Cannot find " ++ Printer.pr_leconstr_env env sigma (mkConst c)) )
end
| _ -> raise (UserError (None, str "Not a function reference") )
@@ -857,8 +859,7 @@ let make_graph (f_ref : GlobRef.t) =
(match Global.body_of_constant_body c_body with
| None -> error "Cannot build a graph over an axiom!"
| Some (body, _) ->
- let env = Global.env () in
- let sigma = Evd.from_env env in
+ let env = Global.env () in
let extern_body,extern_type =
with_full_print (fun () ->
(Constrextern.extern_constr false env sigma (EConstr.of_constr body),
@@ -902,12 +903,11 @@ let make_graph (f_ref : GlobRef.t) =
[((CAst.make id,None),(None,Constrexpr.CStructRec),nal_tas,t,Some b),[]]
in
let mp = Constant.modpath c in
- do_generate_principle [c,Univ.Instance.empty] error_error false false expr_list;
+ let pstate = do_generate_principle ~pstate [c,Univ.Instance.empty] error_error false false expr_list in
(* We register the infos *)
List.iter
(fun ((({CAst.v=id},_),_,_,_,_),_) -> add_Function false (Constant.make2 mp (Label.of_id id)))
- expr_list)
+ expr_list;
+ pstate)
let do_generate_principle = do_generate_principle [] warning_error true
-
-
diff --git a/plugins/funind/indfun.mli b/plugins/funind/indfun.mli
index f209fb19fd..acf85f539e 100644
--- a/plugins/funind/indfun.mli
+++ b/plugins/funind/indfun.mli
@@ -5,18 +5,16 @@ val warn_cannot_define_graph : ?loc:Loc.t -> Pp.t * Pp.t -> unit
val warn_cannot_define_principle : ?loc:Loc.t -> Pp.t * Pp.t -> unit
-val do_generate_principle :
- bool ->
- (Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list ->
- unit
-
+val do_generate_principle : pstate:Proof_global.t option ->
+ bool ->
+ (Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list ->
+ Proof_global.t option
-val functional_induction :
+val functional_induction :
bool ->
EConstr.constr ->
(EConstr.constr * EConstr.constr bindings) option ->
Ltac_plugin.Tacexpr.or_and_intro_pattern option ->
Goal.goal Evd.sigma -> Goal.goal list Evd.sigma
-
-val make_graph : GlobRef.t -> unit
+val make_graph : pstate:Proof_global.t option -> GlobRef.t -> Proof_global.t option
diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml
index e34323abf4..40f66ce5eb 100644
--- a/plugins/funind/indfun_common.ml
+++ b/plugins/funind/indfun_common.ml
@@ -129,7 +129,7 @@ let get_locality = function
| Local -> true
| Global -> false
-let save with_clean id const ?hook uctx (locality,_,kind) =
+let save id const ?hook uctx (locality,_,kind) =
let fix_exn = Future.fix_exn_of const.const_entry_body in
let l,r = match locality with
| Discharge when Lib.sections_are_opened () ->
@@ -143,7 +143,6 @@ let save with_clean id const ?hook uctx (locality,_,kind) =
let kn = declare_constant id ~local (DefinitionEntry const, k) in
(locality, ConstRef kn)
in
- if with_clean then Proof_global.discard_current ();
Lemmas.call_hook ?hook ~fix_exn uctx [] l r;
definition_message id
diff --git a/plugins/funind/indfun_common.mli b/plugins/funind/indfun_common.mli
index 12facc5744..9670cf1fa7 100644
--- a/plugins/funind/indfun_common.mli
+++ b/plugins/funind/indfun_common.mli
@@ -43,8 +43,7 @@ val jmeq : unit -> EConstr.constr
val jmeq_refl : unit -> EConstr.constr
val save
- : bool
- -> Id.t
+ : Id.t
-> Safe_typing.private_constants Entries.definition_entry
-> ?hook:Lemmas.declaration_hook
-> UState.t
@@ -78,15 +77,12 @@ val find_Function_infos : Constant.t -> function_info
val find_Function_of_graph : inductive -> function_info
(* WARNING: To be used just after the graph definition !!! *)
val add_Function : bool -> Constant.t -> unit
-
val update_Function : function_info -> unit
-
(** debugging *)
val pr_info : Environ.env -> Evd.evar_map -> function_info -> Pp.t
val pr_table : Environ.env -> Evd.evar_map -> Pp.t
-
(* val function_debug : bool ref *)
val do_observe : unit -> bool
val do_rewrite_dependent : unit -> bool
diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml
index 37dbfec4c9..edb698280f 100644
--- a/plugins/funind/invfun.ml
+++ b/plugins/funind/invfun.ml
@@ -802,16 +802,16 @@ let derive_correctness make_scheme (funs: pconstant list) (graphs:inductive list
Ensures by: obvious
i*)
let lem_id = mk_correct_id f_id in
- let (typ,_) = lemmas_types_infos.(i) in
- Lemmas.start_proof
+ let (typ,_) = lemmas_types_infos.(i) in
+ let pstate = Lemmas.start_proof ~ontop:None
lem_id
(Decl_kinds.Global,false,((Decl_kinds.Proof Decl_kinds.Theorem)))
!evd
- typ;
- ignore (Pfedit.by
+ typ in
+ let pstate = fst @@ Pfedit.by
(Proofview.V82.tactic (observe_tac ("prove correctness ("^(Id.to_string f_id)^")")
- (proving_tac i))));
- (Lemmas.save_proof (Vernacexpr.(Proved(Proof_global.Transparent,None))));
+ (proving_tac i))) pstate in
+ let _ = Lemmas.save_proof_proved ?proof:None ~pstate ~opaque:Proof_global.Transparent ~idopt:None in
let finfo = find_Function_infos (fst f_as_constant) in
(* let lem_cst = fst (destConst (Constrintern.global_reference lem_id)) in *)
let _,lem_cst_constr = Evd.fresh_global
@@ -865,13 +865,13 @@ let derive_correctness make_scheme (funs: pconstant list) (graphs:inductive list
Ensures by: obvious
i*)
let lem_id = mk_complete_id f_id in
- Lemmas.start_proof lem_id
+ let pstate = Lemmas.start_proof ~ontop:None lem_id
(Decl_kinds.Global,false,(Decl_kinds.Proof Decl_kinds.Theorem)) sigma
- (fst lemmas_types_infos.(i));
- ignore (Pfedit.by
+ (fst lemmas_types_infos.(i)) in
+ let pstate = fst (Pfedit.by
(Proofview.V82.tactic (observe_tac ("prove completeness ("^(Id.to_string f_id)^")")
- (proving_tac i)))) ;
- (Lemmas.save_proof (Vernacexpr.(Proved(Proof_global.Transparent,None))));
+ (proving_tac i))) pstate) in
+ let _pstate = Lemmas.save_proof_proved ?proof:None ~pstate ~opaque:Proof_global.Transparent ~idopt:None in
let finfo = find_Function_infos (fst f_as_constant) in
let _,lem_cst_constr = Evd.fresh_global
(Global.env ()) !evd (Constrintern.locate_reference (Libnames.qualid_of_ident lem_id)) in
diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml
index e19741a4e9..3c2b03dfe0 100644
--- a/plugins/funind/recdef.ml
+++ b/plugins/funind/recdef.ml
@@ -72,7 +72,7 @@ let declare_fun f_id kind ?univs value =
let ce = definition_entry ?univs value (*FIXME *) in
ConstRef(declare_constant f_id (DefinitionEntry ce, kind));;
-let defined () = Lemmas.save_proof (Vernacexpr.(Proved (Proof_global.Transparent,None)))
+let defined pstate = Lemmas.save_proof_proved ?proof:None ~pstate ~opaque:Proof_global.Transparent ~idopt:None
let def_of_const t =
match (Constr.kind t) with
@@ -228,6 +228,7 @@ let observe strm =
let do_observe_tac s tac g =
let goal = Printer.pr_goal g in
+ let s = s (pf_env g) (project g) in
let lmsg = (str "recdef : ") ++ s in
observe (s++fnl());
Stack.push (lmsg,goal) debug_queue;
@@ -252,8 +253,8 @@ let observe_tclTHENLIST s tacl =
then
let rec aux n = function
| [] -> tclIDTAC
- | [tac] -> observe_tac (s ++ spc () ++ int n) tac
- | tac::tacl -> observe_tac (s ++ spc () ++ int n) (tclTHEN tac (aux (succ n) tacl))
+ | [tac] -> observe_tac (fun env sigma -> s env sigma ++ spc () ++ int n) tac
+ | tac::tacl -> observe_tac (fun env sigma -> s env sigma ++ spc () ++ int n) (tclTHEN tac (aux (succ n) tacl))
in
aux 0 tacl
else tclTHENLIST tacl
@@ -268,11 +269,11 @@ let tclUSER tac is_mes l g =
| None -> tclIDTAC
| Some l -> tclMAP (fun id -> tclTRY (Proofview.V82.of_tactic (clear [id]))) (List.rev l)
in
- observe_tclTHENLIST (str "tclUSER1")
+ observe_tclTHENLIST (fun _ _ -> str "tclUSER1")
[
clear_tac;
if is_mes
- then observe_tclTHENLIST (str "tclUSER2")
+ then observe_tclTHENLIST (fun _ _ -> str "tclUSER2")
[
Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, evaluable_of_global_reference
(delayed_force Indfun_common.ltof_ref))]);
@@ -394,12 +395,12 @@ let treat_case forbid_new_ids to_intros finalize_tac nb_lam e infos : tactic =
) [] rev_context in
let rev_ids = pf_get_new_ids (List.rev ids) g in
let new_b = substl (List.map mkVar rev_ids) b in
- observe_tclTHENLIST (str "treat_case1")
+ observe_tclTHENLIST (fun _ _ -> str "treat_case1")
[
h_intros (List.rev rev_ids);
Proofview.V82.of_tactic (intro_using teq_id);
onLastHypId (fun heq ->
- observe_tclTHENLIST (str "treat_case2")[
+ observe_tclTHENLIST (fun _ _ -> str "treat_case2")[
Proofview.V82.of_tactic (clear to_intros);
h_intros to_intros;
(fun g' ->
@@ -426,6 +427,7 @@ let treat_case forbid_new_ids to_intros finalize_tac nb_lam e infos : tactic =
let rec travel_aux jinfo continuation_tac (expr_info:constr infos) g =
let sigma = project g in
+ let env = pf_env g in
match EConstr.kind sigma expr_info.info with
| CoFix _ | Fix _ -> user_err Pp.(str "Function cannot treat local fixpoint or cofixpoint")
| Proj _ -> user_err Pp.(str "Function cannot treat projections")
@@ -441,18 +443,18 @@ let rec travel_aux jinfo continuation_tac (expr_info:constr infos) g =
| Prod _ ->
begin
try
- check_not_nested (pf_env g) sigma (expr_info.f_id::expr_info.forbidden_ids) expr_info.info;
+ check_not_nested env sigma (expr_info.f_id::expr_info.forbidden_ids) expr_info.info;
jinfo.otherS () expr_info continuation_tac expr_info g
with e when CErrors.noncritical e ->
- user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr_env (pf_env g) sigma expr_info.info ++ str " can not contain a recursive call to " ++ Id.print expr_info.f_id)
+ user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr_env env sigma expr_info.info ++ str " can not contain a recursive call to " ++ Id.print expr_info.f_id)
end
| Lambda(n,t,b) ->
begin
try
- check_not_nested (pf_env g) sigma (expr_info.f_id::expr_info.forbidden_ids) expr_info.info;
+ check_not_nested env sigma (expr_info.f_id::expr_info.forbidden_ids) expr_info.info;
jinfo.otherS () expr_info continuation_tac expr_info g
with e when CErrors.noncritical e ->
- user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr_env (pf_env g) sigma expr_info.info ++ str " can not contain a recursive call to " ++ Id.print expr_info.f_id)
+ user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr_env env sigma expr_info.info ++ str " can not contain a recursive call to " ++ Id.print expr_info.f_id)
end
| Case(ci,t,a,l) ->
begin
@@ -480,8 +482,8 @@ let rec travel_aux jinfo continuation_tac (expr_info:constr infos) g =
jinfo.apP (f,args) expr_info continuation_tac in
travel_args jinfo
expr_info.is_main_branch new_continuation_tac new_infos g
- | Case _ -> user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr_env (pf_env g) sigma expr_info.info ++ str " can not contain an applied match (See Limitation in Section 2.3 of refman)")
- | _ -> anomaly (Pp.str "travel_aux : unexpected "++ Printer.pr_leconstr_env (pf_env g) sigma expr_info.info ++ Pp.str ".")
+ | Case _ -> user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr_env env sigma expr_info.info ++ str " can not contain an applied match (See Limitation in Section 2.3 of refman)")
+ | _ -> anomaly (Pp.str "travel_aux : unexpected "++ Printer.pr_leconstr_env env sigma expr_info.info ++ Pp.str ".")
end
| Cast(t,_,_) -> travel jinfo continuation_tac {expr_info with info=t} g
| Const _ | Var _ | Meta _ | Evar _ | Sort _ | Construct _ | Ind _ | Int _ ->
@@ -503,10 +505,9 @@ and travel_args jinfo is_final continuation_tac infos =
travel jinfo new_continuation_tac
{infos with info=arg;is_final=false}
and travel jinfo continuation_tac expr_info =
- fun g ->
observe_tac
- (str jinfo.message ++ Printer.pr_leconstr_env (pf_env g) (project g) expr_info.info)
- (travel_aux jinfo continuation_tac expr_info) g
+ (fun env sigma -> str jinfo.message ++ Printer.pr_leconstr_env env sigma expr_info.info)
+ (travel_aux jinfo continuation_tac expr_info)
(* Termination proof *)
@@ -527,16 +528,16 @@ let rec prove_lt hyple g =
in
let y =
List.hd (List.tl (snd (decompose_app sigma (pf_unsafe_type_of g (mkVar h))))) in
- observe_tclTHENLIST (str "prove_lt1")[
+ observe_tclTHENLIST (fun _ _ -> str "prove_lt1")[
Proofview.V82.of_tactic (apply (mkApp(le_lt_trans (),[|varx;y;varz;mkVar h|])));
- observe_tac (str "prove_lt") (prove_lt hyple)
+ observe_tac (fun _ _ -> str "prove_lt") (prove_lt hyple)
]
with Not_found ->
(
(
- observe_tclTHENLIST (str "prove_lt2")[
+ observe_tclTHENLIST (fun _ _ -> str "prove_lt2")[
Proofview.V82.of_tactic (apply (delayed_force lt_S_n));
- (observe_tac (str "assumption: " ++ Printer.pr_goal g) (Proofview.V82.of_tactic assumption))
+ (observe_tac (fun _ _ -> str "assumption: " ++ Printer.pr_goal g) (Proofview.V82.of_tactic assumption))
])
)
end
@@ -552,26 +553,26 @@ let rec destruct_bounds_aux infos (bound,hyple,rechyps) lbounds g =
let h' = next_ident_away_in_goal (h'_id) ids in
let ids = h'::ids in
let def = next_ident_away_in_goal def_id ids in
- observe_tclTHENLIST (str "destruct_bounds_aux1")[
+ observe_tclTHENLIST (fun _ _ -> str "destruct_bounds_aux1")[
Proofview.V82.of_tactic (split (ImplicitBindings [s_max]));
Proofview.V82.of_tactic (intro_then
(fun id ->
Proofview.V82.tactic begin
- observe_tac (str "destruct_bounds_aux")
+ observe_tac (fun _ _ -> str "destruct_bounds_aux")
(tclTHENS (Proofview.V82.of_tactic (simplest_case (mkVar id)))
[
- observe_tclTHENLIST (str "")[Proofview.V82.of_tactic (intro_using h_id);
+ observe_tclTHENLIST (fun _ _ -> str "")[Proofview.V82.of_tactic (intro_using h_id);
Proofview.V82.of_tactic (simplest_elim(mkApp(delayed_force lt_n_O,[|s_max|])));
Proofview.V82.of_tactic default_full_auto];
- observe_tclTHENLIST (str "destruct_bounds_aux2")[
- observe_tac (str "clearing k ") (Proofview.V82.of_tactic (clear [id]));
+ observe_tclTHENLIST (fun _ _ -> str "destruct_bounds_aux2")[
+ observe_tac (fun _ _ -> str "clearing k ") (Proofview.V82.of_tactic (clear [id]));
h_intros [k;h';def];
- observe_tac (str "simple_iter") (Proofview.V82.of_tactic (simpl_iter Locusops.onConcl));
- observe_tac (str "unfold functional")
+ observe_tac (fun _ _ -> str "simple_iter") (Proofview.V82.of_tactic (simpl_iter Locusops.onConcl));
+ observe_tac (fun _ _ -> str "unfold functional")
(Proofview.V82.of_tactic (unfold_in_concl[(Locus.OnlyOccurrences [1],
evaluable_of_global_reference infos.func)]));
(
- observe_tclTHENLIST (str "test")[
+ observe_tclTHENLIST (fun _ _ -> str "test")[
list_rewrite true
(List.fold_right
(fun e acc -> (mkVar e,true)::acc)
@@ -582,16 +583,16 @@ let rec destruct_bounds_aux infos (bound,hyple,rechyps) lbounds g =
(* (List.map (fun e -> (mkVar e,true)) infos.eqs) *)
(* ; *)
- (observe_tac (str "finishing")
+ (observe_tac (fun _ _ -> str "finishing")
(tclORELSE
(Proofview.V82.of_tactic intros_reflexivity)
- (observe_tac (str "calling prove_lt") (prove_lt hyple))))])
+ (observe_tac (fun _ _ -> str "calling prove_lt") (prove_lt hyple))))])
]
]
)end))
] g
| (_,v_bound)::l ->
- observe_tclTHENLIST (str "destruct_bounds_aux3")[
+ observe_tclTHENLIST (fun _ _ -> str "destruct_bounds_aux3")[
Proofview.V82.of_tactic (simplest_elim (mkVar v_bound));
Proofview.V82.of_tactic (clear [v_bound]);
tclDO 2 (Proofview.V82.of_tactic intro);
@@ -599,7 +600,7 @@ let rec destruct_bounds_aux infos (bound,hyple,rechyps) lbounds g =
(fun p_hyp ->
(onNthHypId 2
(fun p ->
- observe_tclTHENLIST (str "destruct_bounds_aux4")[
+ observe_tclTHENLIST (fun _ _ -> str "destruct_bounds_aux4")[
Proofview.V82.of_tactic (simplest_elim
(mkApp(delayed_force max_constr, [| bound; mkVar p|])));
tclDO 3 (Proofview.V82.of_tactic intro);
@@ -623,32 +624,33 @@ let destruct_bounds infos =
let terminate_app f_and_args expr_info continuation_tac infos =
if expr_info.is_final && expr_info.is_main_branch
then
- observe_tclTHENLIST (str "terminate_app1")[
+ observe_tclTHENLIST (fun _ _ -> str "terminate_app1")[
continuation_tac infos;
- observe_tac (str "first split")
+ observe_tac (fun _ _ -> str "first split")
(Proofview.V82.of_tactic (split (ImplicitBindings [infos.info])));
- observe_tac (str "destruct_bounds (1)") (destruct_bounds infos)
+ observe_tac (fun _ _ -> str "destruct_bounds (1)") (destruct_bounds infos)
]
else continuation_tac infos
let terminate_others _ expr_info continuation_tac infos =
if expr_info.is_final && expr_info.is_main_branch
then
- observe_tclTHENLIST (str "terminate_others")[
+ observe_tclTHENLIST (fun _ _ -> str "terminate_others")[
continuation_tac infos;
- observe_tac (str "first split")
+ observe_tac (fun _ _ -> str "first split")
(Proofview.V82.of_tactic (split (ImplicitBindings [infos.info])));
- observe_tac (str "destruct_bounds") (destruct_bounds infos)
+ observe_tac (fun _ _ -> str "destruct_bounds") (destruct_bounds infos)
]
else continuation_tac infos
let terminate_letin (na,b,t,e) expr_info continuation_tac info g =
let sigma = project g in
+ let env = pf_env g in
let new_e = subst1 info.info e in
let new_forbidden =
let forbid =
try
- check_not_nested (pf_env g) sigma (expr_info.f_id::expr_info.forbidden_ids) b;
+ check_not_nested env sigma (expr_info.f_id::expr_info.forbidden_ids) b;
true
with e when CErrors.noncritical e -> false
in
@@ -693,7 +695,7 @@ let mkDestructEq :
let new_hyps = mkApp(Lazy.force refl_equal, [|type_of_expr; expr|])::
to_revert_constr in
pf_typel new_hyps (fun _ ->
- observe_tclTHENLIST (str "mkDestructEq")
+ observe_tclTHENLIST (fun _ _ -> str "mkDestructEq")
[Proofview.V82.of_tactic (generalize new_hyps);
(fun g2 ->
let changefun patvars env sigma =
@@ -705,9 +707,10 @@ let mkDestructEq :
let terminate_case next_step (ci,a,t,l) expr_info continuation_tac infos g =
let sigma = project g in
+ let env = pf_env g in
let f_is_present =
try
- check_not_nested (pf_env g) sigma (expr_info.f_id::expr_info.forbidden_ids) a;
+ check_not_nested env sigma (expr_info.f_id::expr_info.forbidden_ids) a;
false
with e when CErrors.noncritical e ->
true
@@ -721,45 +724,46 @@ let terminate_case next_step (ci,a,t,l) expr_info continuation_tac infos g =
let destruct_tac,rev_to_thin_intro =
mkDestructEq [expr_info.rec_arg_id] a' g in
let to_thin_intro = List.rev rev_to_thin_intro in
- observe_tac (str "treating cases (" ++ int (Array.length l) ++ str")" ++ spc () ++ Printer.pr_leconstr_env (pf_env g) sigma a')
+ observe_tac (fun _ _ -> str "treating cases (" ++ int (Array.length l) ++ str")" ++ spc () ++ Printer.pr_leconstr_env (pf_env g) sigma a')
(try
(tclTHENS
destruct_tac
- (List.map_i (fun i e -> observe_tac (str "do treat case") (treat_case f_is_present to_thin_intro (next_step continuation_tac) ci.ci_cstr_ndecls.(i) e new_info)) 0 (Array.to_list l)
+ (List.map_i (fun i e -> observe_tac (fun _ _ -> str "do treat case") (treat_case f_is_present to_thin_intro (next_step continuation_tac) ci.ci_cstr_ndecls.(i) e new_info)) 0 (Array.to_list l)
))
with
| UserError(Some "Refiner.thensn_tac3",_)
| UserError(Some "Refiner.tclFAIL_s",_) ->
- (observe_tac (str "is computable " ++ Printer.pr_leconstr_env (pf_env g) sigma new_info.info) (next_step continuation_tac {new_info with info = Reductionops.nf_betaiotazeta (pf_env g) sigma new_info.info} )
+ (observe_tac (fun _ _ -> str "is computable " ++ Printer.pr_leconstr_env env sigma new_info.info) (next_step continuation_tac {new_info with info = Reductionops.nf_betaiotazeta (pf_env g) sigma new_info.info} )
))
g
let terminate_app_rec (f,args) expr_info continuation_tac _ g =
let sigma = project g in
- List.iter (check_not_nested (pf_env g) sigma (expr_info.f_id::expr_info.forbidden_ids))
+ let env = pf_env g in
+ List.iter (check_not_nested env sigma (expr_info.f_id::expr_info.forbidden_ids))
args;
begin
try
let v = List.assoc_f (List.equal (EConstr.eq_constr sigma)) args expr_info.args_assoc in
let new_infos = {expr_info with info = v} in
- observe_tclTHENLIST (str "terminate_app_rec")[
+ observe_tclTHENLIST (fun _ _ -> str "terminate_app_rec")[
continuation_tac new_infos;
if expr_info.is_final && expr_info.is_main_branch
then
- observe_tclTHENLIST (str "terminate_app_rec1")[
- observe_tac (str "first split")
+ observe_tclTHENLIST (fun _ _ -> str "terminate_app_rec1")[
+ observe_tac (fun _ _ -> str "first split")
(Proofview.V82.of_tactic (split (ImplicitBindings [new_infos.info])));
- observe_tac (str "destruct_bounds (3)")
+ observe_tac (fun _ _ -> str "destruct_bounds (3)")
(destruct_bounds new_infos)
]
else
tclIDTAC
] g
with Not_found ->
- observe_tac (str "terminate_app_rec not found") (tclTHENS
+ observe_tac (fun _ _ -> str "terminate_app_rec not found") (tclTHENS
(Proofview.V82.of_tactic (simplest_elim (mkApp(mkVar expr_info.ih,Array.of_list args))))
[
- observe_tclTHENLIST (str "terminate_app_rec2")[
+ observe_tclTHENLIST (fun _ _ -> str "terminate_app_rec2")[
Proofview.V82.of_tactic (intro_using rec_res_id);
Proofview.V82.of_tactic intro;
onNthHypId 1
@@ -772,14 +776,14 @@ let terminate_app_rec (f,args) expr_info continuation_tac _ g =
(v,v_bound)::expr_info.values_and_bounds;
args_assoc=(args,mkVar v)::expr_info.args_assoc
} in
- observe_tclTHENLIST (str "terminate_app_rec3")[
+ observe_tclTHENLIST (fun _ _ -> str "terminate_app_rec3")[
continuation_tac new_infos;
if expr_info.is_final && expr_info.is_main_branch
then
- observe_tclTHENLIST (str "terminate_app_rec4")[
- observe_tac (str "first split")
+ observe_tclTHENLIST (fun _ _ -> str "terminate_app_rec4")[
+ observe_tac (fun _ _ -> str "first split")
(Proofview.V82.of_tactic (split (ImplicitBindings [new_infos.info])));
- observe_tac (str "destruct_bounds (2)")
+ observe_tac (fun _ _ -> str "destruct_bounds (2)")
(destruct_bounds new_infos)
]
else
@@ -789,12 +793,12 @@ let terminate_app_rec (f,args) expr_info continuation_tac _ g =
)
)
];
- observe_tac (str "proving decreasing") (
+ observe_tac (fun _ _ -> str "proving decreasing") (
tclTHENS (* proof of args < formal args *)
(Proofview.V82.of_tactic (apply (Lazy.force expr_info.acc_inv)))
[
- observe_tac (str "assumption") (Proofview.V82.of_tactic assumption);
- observe_tclTHENLIST (str "terminate_app_rec5")
+ observe_tac (fun _ _ -> str "assumption") (Proofview.V82.of_tactic assumption);
+ observe_tclTHENLIST (fun _ _ -> str "terminate_app_rec5")
[
tclTRY(list_rewrite true
(List.map
@@ -830,7 +834,7 @@ let prove_terminate = travel terminate_info
(* Equation proof *)
let equation_case next_step (ci,a,t,l) expr_info continuation_tac infos =
- observe_tac (str "equation case") (terminate_case next_step (ci,a,t,l) expr_info continuation_tac infos)
+ observe_tac (fun _ _ -> str "equation case") (terminate_case next_step (ci,a,t,l) expr_info continuation_tac infos)
let rec prove_le g =
let sigma = project g in
@@ -856,9 +860,9 @@ let rec prove_le g =
let _,args = decompose_app sigma t in
List.hd (List.tl args)
in
- observe_tclTHENLIST (str "prove_le")[
+ observe_tclTHENLIST (fun _ _ -> str "prove_le")[
Proofview.V82.of_tactic (apply(mkApp(le_trans (),[|x;y;z;mkVar h|])));
- observe_tac (str "prove_le (rec)") (prove_le)
+ observe_tac (fun _ _ -> str "prove_le (rec)") (prove_le)
]
with Not_found -> tclFAIL 0 (mt())
end;
@@ -868,8 +872,8 @@ let rec prove_le g =
let rec make_rewrite_list expr_info max = function
| [] -> tclIDTAC
| (_,p,hp)::l ->
- observe_tac (str "make_rewrite_list") (tclTHENS
- (observe_tac (str "rewrite heq on " ++ Id.print p ) (
+ observe_tac (fun _ _ -> str "make_rewrite_list") (tclTHENS
+ (observe_tac (fun _ _ -> str "rewrite heq on " ++ Id.print p ) (
(fun g ->
let sigma = project g in
let t_eq = compute_renamed_type g (mkVar hp) in
@@ -886,16 +890,16 @@ let rec make_rewrite_list expr_info max = function
CAst.make @@ (NamedHyp k, f_S max)]) false) g) )
)
[make_rewrite_list expr_info max l;
- observe_tclTHENLIST (str "make_rewrite_list")[ (* x < S max proof *)
+ observe_tclTHENLIST (fun _ _ -> str "make_rewrite_list")[ (* x < S max proof *)
Proofview.V82.of_tactic (apply (delayed_force le_lt_n_Sm));
- observe_tac (str "prove_le(2)") prove_le
+ observe_tac (fun _ _ -> str "prove_le(2)") prove_le
]
] )
let make_rewrite expr_info l hp max =
tclTHENFIRST
- (observe_tac (str "make_rewrite") (make_rewrite_list expr_info max l))
- (observe_tac (str "make_rewrite") (tclTHENS
+ (observe_tac (fun _ _ -> str "make_rewrite") (make_rewrite_list expr_info max l))
+ (observe_tac (fun _ _ -> str "make_rewrite") (tclTHENS
(fun g ->
let sigma = project g in
let t_eq = compute_renamed_type g (mkVar hp) in
@@ -905,30 +909,30 @@ let make_rewrite expr_info l hp max =
let def_na,_,_ = destProd sigma t in
Nameops.Name.get_id k_na.binder_name,Nameops.Name.get_id def_na.binder_name
in
- observe_tac (str "general_rewrite_bindings")
+ observe_tac (fun _ _ -> str "general_rewrite_bindings")
(Proofview.V82.of_tactic (general_rewrite_bindings false Locus.AllOccurrences
true (* dep proofs also: *) true
(mkVar hp,
ExplicitBindings[CAst.make @@ (NamedHyp def, expr_info.f_constr);
CAst.make @@ (NamedHyp k, f_S (f_S max))]) false)) g)
- [observe_tac(str "make_rewrite finalize") (
+ [observe_tac(fun _ _ -> str "make_rewrite finalize") (
(* tclORELSE( h_reflexivity) *)
- (observe_tclTHENLIST (str "make_rewrite")[
+ (observe_tclTHENLIST (fun _ _ -> str "make_rewrite")[
Proofview.V82.of_tactic (simpl_iter Locusops.onConcl);
- observe_tac (str "unfold functional")
+ observe_tac (fun _ _ -> str "unfold functional")
(Proofview.V82.of_tactic (unfold_in_concl[(Locus.OnlyOccurrences [1],
evaluable_of_global_reference expr_info.func)]));
(list_rewrite true
(List.map (fun e -> mkVar e,true) expr_info.eqs));
- (observe_tac (str "h_reflexivity")
+ (observe_tac (fun _ _ -> str "h_reflexivity")
(Proofview.V82.of_tactic intros_reflexivity)
)
]))
;
- observe_tclTHENLIST (str "make_rewrite1")[ (* x < S (S max) proof *)
+ observe_tclTHENLIST (fun _ _ -> str "make_rewrite1")[ (* x < S (S max) proof *)
Proofview.V82.of_tactic (apply (EConstr.of_constr (delayed_force le_lt_SS)));
- observe_tac (str "prove_le (3)") prove_le
+ observe_tac (fun _ _ -> str "prove_le (3)") prove_le
]
])
)
@@ -937,7 +941,7 @@ let rec compute_max rew_tac max l =
match l with
| [] -> rew_tac max
| (_,p,_)::l ->
- observe_tclTHENLIST (str "compute_max")[
+ observe_tclTHENLIST (fun _ _ -> str "compute_max")[
Proofview.V82.of_tactic (simplest_elim
(mkApp(delayed_force max_constr, [| max; mkVar p|])));
tclDO 3 (Proofview.V82.of_tactic intro);
@@ -954,17 +958,17 @@ let rec destruct_hex expr_info acc l =
match List.rev acc with
| [] -> tclIDTAC
| (_,p,hp)::tl ->
- observe_tac (str "compute max ") (compute_max (make_rewrite expr_info tl hp) (mkVar p) tl)
+ observe_tac (fun _ _ -> str "compute max ") (compute_max (make_rewrite expr_info tl hp) (mkVar p) tl)
end
| (v,hex)::l ->
- observe_tclTHENLIST (str "destruct_hex")[
+ observe_tclTHENLIST (fun _ _ -> str "destruct_hex")[
Proofview.V82.of_tactic (simplest_case (mkVar hex));
Proofview.V82.of_tactic (clear [hex]);
tclDO 2 (Proofview.V82.of_tactic intro);
onNthHypId 1 (fun hp ->
onNthHypId 2 (fun p ->
observe_tac
- (str "destruct_hex after " ++ Id.print hp ++ spc () ++ Id.print p)
+ (fun _ _ -> str "destruct_hex after " ++ Id.print hp ++ spc () ++ Id.print p)
(destruct_hex expr_info ((v,p,hp)::acc) l)
)
)
@@ -972,7 +976,7 @@ let rec destruct_hex expr_info acc l =
let rec intros_values_eq expr_info acc =
tclORELSE(
- observe_tclTHENLIST (str "intros_values_eq")[
+ observe_tclTHENLIST (fun _ _ -> str "intros_values_eq")[
tclDO 2 (Proofview.V82.of_tactic intro);
onNthHypId 1 (fun hex ->
(onNthHypId 2 (fun v -> intros_values_eq expr_info ((v,hex)::acc)))
@@ -983,23 +987,17 @@ let rec intros_values_eq expr_info acc =
))
let equation_others _ expr_info continuation_tac infos =
- fun g ->
- let env = pf_env g in
- let sigma = project g in
if expr_info.is_final && expr_info.is_main_branch
then
- observe_tac (str "equation_others (cont_tac +intros) " ++ Printer.pr_leconstr_env env sigma expr_info.info)
+ observe_tac (fun env sigma -> str "equation_others (cont_tac +intros) " ++ Printer.pr_leconstr_env env sigma expr_info.info)
(tclTHEN
(continuation_tac infos)
- (fun g ->
- let env = pf_env g in
- let sigma = project g in
- observe_tac (str "intros_values_eq equation_others " ++ Printer.pr_leconstr_env env sigma expr_info.info) (intros_values_eq expr_info []) g)) g
- else observe_tac (str "equation_others (cont_tac) " ++ Printer.pr_leconstr_env env sigma expr_info.info) (continuation_tac infos) g
+ (observe_tac (fun env sigma -> str "intros_values_eq equation_others " ++ Printer.pr_leconstr_env env sigma expr_info.info) (intros_values_eq expr_info [])))
+ else observe_tac (fun env sigma -> str "equation_others (cont_tac) " ++ Printer.pr_leconstr_env env sigma expr_info.info) (continuation_tac infos)
let equation_app f_and_args expr_info continuation_tac infos =
if expr_info.is_final && expr_info.is_main_branch
- then ((observe_tac (str "intros_values_eq equation_app") (intros_values_eq expr_info [])))
+ then ((observe_tac (fun _ _ -> str "intros_values_eq equation_app") (intros_values_eq expr_info [])))
else continuation_tac infos
let equation_app_rec (f,args) expr_info continuation_tac info g =
@@ -1008,19 +1006,19 @@ let equation_app_rec (f,args) expr_info continuation_tac info g =
try
let v = List.assoc_f (List.equal (EConstr.eq_constr sigma)) args expr_info.args_assoc in
let new_infos = {expr_info with info = v} in
- observe_tac (str "app_rec found") (continuation_tac new_infos) g
+ observe_tac (fun _ _ -> str "app_rec found") (continuation_tac new_infos) g
with Not_found ->
if expr_info.is_final && expr_info.is_main_branch
then
- observe_tclTHENLIST (str "equation_app_rec")
+ observe_tclTHENLIST (fun _ _ -> str "equation_app_rec")
[ Proofview.V82.of_tactic (simplest_case (mkApp (expr_info.f_terminate,Array.of_list args)));
continuation_tac {expr_info with args_assoc = (args,delayed_force coq_O)::expr_info.args_assoc};
- observe_tac (str "app_rec intros_values_eq") (intros_values_eq expr_info [])
+ observe_tac (fun _ _ -> str "app_rec intros_values_eq") (intros_values_eq expr_info [])
] g
else
- observe_tclTHENLIST (str "equation_app_rec1")[
+ observe_tclTHENLIST (fun _ _ -> str "equation_app_rec1")[
Proofview.V82.of_tactic (simplest_case (mkApp (expr_info.f_terminate,Array.of_list args)));
- observe_tac (str "app_rec not_found") (continuation_tac {expr_info with args_assoc = (args,delayed_force coq_O)::expr_info.args_assoc})
+ observe_tac (fun _ _ -> str "app_rec not_found") (continuation_tac {expr_info with args_assoc = (args,delayed_force coq_O)::expr_info.args_assoc})
] g
end
@@ -1104,7 +1102,7 @@ let termination_proof_header is_mes input_type ids args_id relation
(h_intros args_id)
(tclTHENS
(observe_tac
- (str "first assert")
+ (fun _ _ -> str "first assert")
(Proofview.V82.of_tactic (assert_before
(Name wf_rec_arg)
(mkApp (delayed_force acc_rel,
@@ -1116,7 +1114,7 @@ let termination_proof_header is_mes input_type ids args_id relation
(* accesibility proof *)
tclTHENS
(observe_tac
- (str "second assert")
+ (fun _ _ -> str "second assert")
(Proofview.V82.of_tactic (assert_before
(Name wf_thm)
(mkApp (delayed_force well_founded,[|input_type;relation|]))
@@ -1124,26 +1122,26 @@ let termination_proof_header is_mes input_type ids args_id relation
)
[
(* interactive proof that the relation is well_founded *)
- observe_tac (str "wf_tac") (wf_tac is_mes (Some args_id));
+ observe_tac (fun _ _ -> str "wf_tac") (wf_tac is_mes (Some args_id));
(* this gives the accessibility argument *)
observe_tac
- (str "apply wf_thm")
+ (fun _ _ -> str "apply wf_thm")
(Proofview.V82.of_tactic (Simple.apply (mkApp(mkVar wf_thm,[|mkVar rec_arg_id|])))
)
]
;
(* rest of the proof *)
- observe_tclTHENLIST (str "rest of proof")
- [observe_tac (str "generalize")
+ observe_tclTHENLIST (fun _ _ -> str "rest of proof")
+ [observe_tac (fun _ _ -> str "generalize")
(onNLastHypsId (nargs+1)
(tclMAP (fun id ->
tclTHEN (Proofview.V82.of_tactic (Tactics.generalize [mkVar id])) (Proofview.V82.of_tactic (clear [id])))
))
;
- observe_tac (str "fix") (Proofview.V82.of_tactic (fix hrec (nargs+1)));
+ observe_tac (fun _ _ -> str "fix") (Proofview.V82.of_tactic (fix hrec (nargs+1)));
h_intros args_id;
Proofview.V82.of_tactic (Simple.intro wf_rec_arg);
- observe_tac (str "tac") (tac wf_rec_arg hrec wf_rec_arg acc_inv)
+ observe_tac (fun _ _ -> str "tac") (tac wf_rec_arg hrec wf_rec_arg acc_inv)
]
]
) g
@@ -1222,8 +1220,8 @@ let whole_start (concl_tac:tactic) nb_args is_mes func input_type relation rec_a
g
end
-let get_current_subgoals_types () =
- let p = Proof_global.give_me_the_proof () in
+let get_current_subgoals_types pstate =
+ let p = Proof_global.give_me_the_proof pstate in
let sgs,_,_,_,sigma = Proof.proof p in
sigma, List.map (Goal.V82.abstract_type sigma) sgs
@@ -1283,8 +1281,8 @@ let clear_goals sigma =
List.map clear_goal
-let build_new_goal_type () =
- let sigma, sub_gls_types = get_current_subgoals_types () in
+let build_new_goal_type pstate =
+ let sigma, sub_gls_types = get_current_subgoals_types pstate in
(* Pp.msgnl (str "sub_gls_types1 := " ++ Util.prlist_with_sep (fun () -> Pp.fnl () ++ Pp.fnl ()) Printer.pr_lconstr sub_gls_types); *)
let sub_gls_types = clear_goals sigma sub_gls_types in
(* Pp.msgnl (str "sub_gls_types2 := " ++ Pp.prlist_with_sep (fun () -> Pp.fnl () ++ Pp.fnl ()) Printer.pr_lconstr sub_gls_types); *)
@@ -1299,9 +1297,9 @@ let is_opaque_constant c =
| Declarations.Def _ -> Proof_global.Transparent
| Declarations.Primitive _ -> Proof_global.Opaque
-let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decompose_and_tac,nb_goal) =
+let open_new_goal pstate build_proof sigma using_lemmas ref_ goal_name (gls_type,decompose_and_tac,nb_goal) =
(* Pp.msgnl (str "gls_type := " ++ Printer.pr_lconstr gls_type); *)
- let current_proof_name = Proof_global.get_current_proof_name () in
+ let current_proof_name = Proof_global.get_current_proof_name pstate in
let name = match goal_name with
| Some s -> s
| None ->
@@ -1325,11 +1323,10 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp
let lid = ref [] in
let h_num = ref (-1) in
let env = Global.env () in
- Proof_global.discard_all ();
- build_proof (Evd.from_env env)
+ let pstate = build_proof env (Evd.from_env env)
( fun gls ->
let hid = next_ident_away_in_goal h_id (pf_ids_of_hyps gls) in
- observe_tclTHENLIST (str "")
+ observe_tclTHENLIST (fun _ _ -> str "")
[
Proofview.V82.of_tactic (generalize [lemma]);
Proofview.V82.of_tactic (Simple.intro hid);
@@ -1353,7 +1350,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp
Proofview.V82.of_tactic (Auto.h_auto None [] (Some [])) g
| _ ->
incr h_num;
- (observe_tac (str "finishing using")
+ (observe_tac (fun _ _ -> str "finishing using")
(
tclCOMPLETE(
tclFIRST[
@@ -1369,20 +1366,19 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp
)
)
g)
-;
- Lemmas.save_proof (Vernacexpr.Proved(opacity,None));
+ in
+ let _pstate = Lemmas.save_proof_proved ?proof:None ~pstate ~opaque:opacity ~idopt:None in
+ ()
in
- Lemmas.start_proof
+ let pstate = Lemmas.start_proof ~ontop:(Some pstate)
na
(Decl_kinds.Global, false (* FIXME *), Decl_kinds.Proof Decl_kinds.Lemma)
- sigma gls_type
- ~hook:(Lemmas.mk_hook hook);
- if Indfun_common.is_strict_tcc ()
+ sigma gls_type ~hook:(Lemmas.mk_hook hook) in
+ let pstate = if Indfun_common.is_strict_tcc ()
then
- ignore (by (Proofview.V82.tactic (tclIDTAC)))
+ fst @@ by (Proofview.V82.tactic (tclIDTAC)) pstate
else
- begin
- ignore (by (Proofview.V82.tactic begin
+ fst @@ by (Proofview.V82.tactic begin
fun g ->
tclTHEN
(decompose_and_tac)
@@ -1398,14 +1394,12 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp
)
using_lemmas)
) tclIDTAC)
- g end))
- end;
+ g end) pstate
+ in
try
- ignore (by (Proofview.V82.tactic tclIDTAC)); (* raises UserError _ if the proof is complete *)
+ Some (fst @@ by (Proofview.V82.tactic tclIDTAC) pstate) (* raises UserError _ if the proof is complete *)
with UserError _ ->
- defined ()
-
-
+ defined pstate
let com_terminate
tcc_lemma_name
@@ -1418,32 +1412,26 @@ let com_terminate
thm_name using_lemmas
nb_args ctx
hook =
- let start_proof ctx (tac_start:tactic) (tac_end:tactic) =
- let evd, env = Pfedit.get_current_context () in (* XXX *)
- Lemmas.start_proof thm_name
+ let start_proof env ctx (tac_start:tactic) (tac_end:tactic) =
+ let pstate = Lemmas.start_proof ~ontop:None thm_name
(Global, false (* FIXME *), Proof Lemma) ~sign:(Environ.named_context_val env)
- ctx (EConstr.of_constr (compute_terminate_type nb_args fonctional_ref)) ~hook;
-
- ignore (by (Proofview.V82.tactic (observe_tac (str "starting_tac") tac_start)));
- ignore (by (Proofview.V82.tactic (observe_tac (str "whole_start") (whole_start tac_end nb_args is_mes fonctional_ref
- input_type relation rec_arg_num ))))
+ ctx (EConstr.of_constr (compute_terminate_type nb_args fonctional_ref)) ~hook in
+ let pstate = fst @@ by (Proofview.V82.tactic (observe_tac (fun _ _ -> str "starting_tac") tac_start)) pstate in
+ fst @@ by (Proofview.V82.tactic (observe_tac (fun _ _ -> str "whole_start") (whole_start tac_end nb_args is_mes fonctional_ref
+ input_type relation rec_arg_num ))) pstate
in
- start_proof ctx tclIDTAC tclIDTAC;
+ let pstate = start_proof Global.(env ()) ctx tclIDTAC tclIDTAC in
try
- let sigma, new_goal_type = build_new_goal_type () in
+ let sigma, new_goal_type = build_new_goal_type pstate in
let sigma = Evd.from_ctx (Evd.evar_universe_context sigma) in
- open_new_goal start_proof sigma
+ open_new_goal pstate start_proof sigma
using_lemmas tcc_lemma_ref
(Some tcc_lemma_name)
- (new_goal_type);
+ (new_goal_type)
with EmptySubgoals ->
(* a non recursive function declared with measure ! *)
tcc_lemma_ref := Not_needed;
- defined ()
-
-
-
-
+ defined pstate
let start_equation (f:GlobRef.t) (term_f:GlobRef.t)
(cont_tactic:Id.t list -> tactic) g =
@@ -1453,33 +1441,27 @@ let start_equation (f:GlobRef.t) (term_f:GlobRef.t)
let terminate_constr = EConstr.of_constr terminate_constr in
let nargs = nb_prod (project g) (EConstr.of_constr (type_of_const sigma terminate_constr)) in
let x = n_x_id ids nargs in
- observe_tac (str "start_equation") (observe_tclTHENLIST (str "start_equation") [
+ observe_tac (fun _ _ -> str "start_equation") (observe_tclTHENLIST (fun _ _ -> str "start_equation") [
h_intros x;
Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, evaluable_of_global_reference f)]);
- observe_tac (str "simplest_case")
+ observe_tac (fun _ _ -> str "simplest_case")
(Proofview.V82.of_tactic (simplest_case (mkApp (terminate_constr,
Array.of_list (List.map mkVar x)))));
- observe_tac (str "prove_eq") (cont_tactic x)]) g;;
+ observe_tac (fun _ _ -> str "prove_eq") (cont_tactic x)]) g;;
-let (com_eqn : int -> Id.t ->
- GlobRef.t -> GlobRef.t -> GlobRef.t
- -> Constr.t -> unit) =
- fun nb_arg eq_name functional_ref f_ref terminate_ref equation_lemma_type ->
+let com_eqn sign uctx nb_arg eq_name functional_ref f_ref terminate_ref equation_lemma_type =
let open CVars in
let opacity =
match terminate_ref with
| ConstRef c -> is_opaque_constant c
| _ -> anomaly ~label:"terminate_lemma" (Pp.str "not a constant.")
in
- let evd, env = Pfedit.get_current_context () in (* XXX *)
- let evd = Evd.from_ctx (Evd.evar_universe_context evd) in
+ let evd = Evd.from_ctx uctx in
let f_constr = constr_of_global f_ref in
let equation_lemma_type = subst1 f_constr equation_lemma_type in
- (Lemmas.start_proof eq_name (Global, false, Proof Lemma)
- ~sign:(Environ.named_context_val env)
- evd
- (EConstr.of_constr equation_lemma_type);
- ignore (by
+ let pstate = Lemmas.start_proof ~ontop:None eq_name (Global, false, Proof Lemma) ~sign evd
+ (EConstr.of_constr equation_lemma_type) in
+ let pstate = fst @@ by
(Proofview.V82.tactic (start_equation f_ref terminate_ref
(fun x ->
prove_eq (fun _ -> tclIDTAC)
@@ -1506,15 +1488,16 @@ let (com_eqn : int -> Id.t ->
ih = Id.of_string "______";
}
)
- )));
+ )) pstate in
(* (try Vernacentries.interp (Vernacexpr.VernacShow Vernacexpr.ShowProof) with _ -> ()); *)
(* Vernacentries.interp (Vernacexpr.VernacShow Vernacexpr.ShowScript); *)
- Flags.silently (fun () -> Lemmas.save_proof (Vernacexpr.Proved(opacity,None))) () ;
-(* Pp.msgnl (str "eqn finished"); *)
- );;
+ let _ = Flags.silently (fun () -> Lemmas.save_proof_proved ?proof:None ~pstate ~opaque:opacity ~idopt:None) () in
+ ()
+(* Pp.msgnl (fun _ _ -> str "eqn finished"); *)
+
let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num eq
- generate_induction_principle using_lemmas : unit =
+ generate_induction_principle using_lemmas : Proof_global.t option =
let open Term in
let open Constr in
let open CVars in
@@ -1529,15 +1512,15 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
let equation_lemma_type = Reductionops.nf_betaiotazeta env evd (Evarutil.nf_evar evd ty) in
let function_type = EConstr.to_constr ~abort_on_undefined_evars:false evd function_type in
let equation_lemma_type = EConstr.Unsafe.to_constr equation_lemma_type in
- (* Pp.msgnl (str "lemma type := " ++ Printer.pr_lconstr equation_lemma_type ++ fnl ()); *)
+ (* Pp.msgnl (fun _ _ -> str "lemma type := " ++ Printer.pr_lconstr equation_lemma_type ++ fnl ()); *)
let res_vars,eq' = decompose_prod equation_lemma_type in
let env_eq' = Environ.push_rel_context (List.map (fun (x,y) -> LocalAssum (x,y)) res_vars) env in
let eq' = Reductionops.nf_zeta env_eq' evd (EConstr.of_constr eq') in
let eq' = EConstr.Unsafe.to_constr eq' in
let res =
-(* Pp.msgnl (str "res_var :=" ++ Printer.pr_lconstr_env (push_rel_context (List.map (function (x,t) -> (x,None,t)) res_vars) env) eq'); *)
-(* Pp.msgnl (str "rec_arg_num := " ++ str (string_of_int rec_arg_num)); *)
-(* Pp.msgnl (str "eq' := " ++ str (string_of_int rec_arg_num)); *)
+(* Pp.msgnl (fun _ _ -> str "res_var :=" ++ Printer.pr_lconstr_env (push_rel_context (List.map (function (x,t) -> (x,None,t)) res_vars) env) eq'); *)
+(* Pp.msgnl (fun _ _ -> str "rec_arg_num := " ++ str (fun _ _ -> string_of_int rec_arg_num)); *)
+(* Pp.msgnl (fun _ _ -> str "eq' := " ++ str (fun _ _ -> string_of_int rec_arg_num)); *)
match Constr.kind eq' with
| App(e,[|_;_;eq_fix|]) ->
mkLambda (make_annot (Name function_name) Sorts.Relevant,function_type,subst_var function_name (compose_lam res_vars eq_fix))
@@ -1562,14 +1545,16 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
let evd = Evd.from_ctx evuctx in
let tcc_lemma_name = add_suffix function_name "_tcc" in
let tcc_lemma_constr = ref Undefined in
- (* let _ = Pp.msgnl (str "relation := " ++ Printer.pr_lconstr_env env_with_pre_rec_args relation) in *)
- let hook _ _ _ _ =
+ (* let _ = Pp.msgnl (fun _ _ -> str "relation := " ++ Printer.pr_lconstr_env env_with_pre_rec_args relation) in *)
+ let hook uctx _ _ _ =
let term_ref = Nametab.locate (qualid_of_ident term_id) in
let f_ref = declare_f function_name (IsProof Lemma) arg_types term_ref in
let _ = Extraction_plugin.Table.extraction_inline true [qualid_of_ident term_id] in
(* message "start second proof"; *)
- let stop =
- try com_eqn (List.length res_vars) equation_id functional_ref f_ref term_ref (subst_var function_name equation_lemma_type);
+ let stop =
+ (* XXX: What is the correct way to get sign at hook time *)
+ let sign = Environ.named_context_val Global.(env ()) in
+ try com_eqn sign uctx (List.length res_vars) equation_id functional_ref f_ref term_ref (subst_var function_name equation_lemma_type);
false
with e when CErrors.noncritical e ->
begin
@@ -1601,14 +1586,14 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
in
(* XXX STATE Why do we need this... why is the toplevel protection not enought *)
funind_purify (fun () ->
- com_terminate
- tcc_lemma_name
- tcc_lemma_constr
- is_mes functional_ref
- (EConstr.of_constr rec_arg_type)
- relation rec_arg_num
- term_id
- using_lemmas
- (List.length res_vars)
- evd (Lemmas.mk_hook hook))
- ()
+ let pstate = com_terminate
+ tcc_lemma_name
+ tcc_lemma_constr
+ is_mes functional_ref
+ (EConstr.of_constr rec_arg_type)
+ relation rec_arg_num
+ term_id
+ using_lemmas
+ (List.length res_vars)
+ evd (Lemmas.mk_hook hook)
+ in pstate) ()
diff --git a/plugins/funind/recdef.mli b/plugins/funind/recdef.mli
index 549f1fc0e4..a006c2c354 100644
--- a/plugins/funind/recdef.mli
+++ b/plugins/funind/recdef.mli
@@ -14,6 +14,6 @@ bool ->
int -> Constrexpr.constr_expr -> (pconstant ->
Indfun_common.tcc_lemma_value ref ->
pconstant ->
- pconstant -> int -> EConstr.types -> int -> EConstr.constr -> unit) -> Constrexpr.constr_expr list -> unit
+ pconstant -> int -> EConstr.types -> int -> EConstr.constr -> unit) -> Constrexpr.constr_expr list -> Proof_global.t option
diff --git a/plugins/ltac/extratactics.mlg b/plugins/ltac/extratactics.mlg
index 0428f08138..f5098d2a34 100644
--- a/plugins/ltac/extratactics.mlg
+++ b/plugins/ltac/extratactics.mlg
@@ -53,6 +53,7 @@ let with_delayed_uconstr ist c tac =
fail_evar = false;
expand_evars = true;
program_mode = false;
+ polymorphic = false;
} in
let c = Tacinterp.type_uconstr ~flags ist c in
Tacticals.New.tclDELAYEDWITHHOLES false c tac
@@ -348,6 +349,7 @@ let constr_flags () = {
Pretyping.fail_evar = false;
Pretyping.expand_evars = true;
Pretyping.program_mode = false;
+ Pretyping.polymorphic = false;
}
let refine_tac ist simple with_classes c =
@@ -813,9 +815,9 @@ END
TACTIC EXTEND transparent_abstract
| [ "transparent_abstract" tactic3(t) ] -> { Proofview.Goal.enter begin fun gl ->
- Abstract.tclABSTRACT ~opaque:false None (Tacinterp.tactic_of_value ist t) end }
+ Abstract.tclABSTRACT ~opaque:false None (Tacinterp.tactic_of_value ist t) end; }
| [ "transparent_abstract" tactic3(t) "using" ident(id) ] -> { Proofview.Goal.enter begin fun gl ->
- Abstract.tclABSTRACT ~opaque:false (Some id) (Tacinterp.tactic_of_value ist t) end }
+ Abstract.tclABSTRACT ~opaque:false (Some id) (Tacinterp.tactic_of_value ist t) end; }
END
(* ********************************************************************* *)
@@ -913,9 +915,9 @@ END
the semantics of the LCF-style tactics, hence with the classic tactic
mode. *)
VERNAC COMMAND EXTEND GrabEvars
-| [ "Grab" "Existential" "Variables" ]
+| ![ proof ] [ "Grab" "Existential" "Variables" ]
=> { classify_as_proofstep }
- -> { Proof_global.simple_with_current_proof (fun _ p -> Proof.V82.grab_evars p) }
+ -> { fun ~pstate -> Option.map (Proof_global.simple_with_current_proof (fun _ p -> Proof.V82.grab_evars p)) pstate }
END
(* Shelves all the goals under focus. *)
@@ -945,9 +947,9 @@ END
(* Command to add every unshelved variables to the focus *)
VERNAC COMMAND EXTEND Unshelve
-| [ "Unshelve" ]
+| ![ proof ] [ "Unshelve" ]
=> { classify_as_proofstep }
- -> { Proof_global.simple_with_current_proof (fun _ p -> Proof.unshelve p) }
+ -> { fun ~pstate -> Option.map (Proof_global.simple_with_current_proof (fun _ p -> Proof.unshelve p)) pstate }
END
(* Gives up on the goals under focus: the goals are considered solved,
@@ -1098,8 +1100,8 @@ END
VERNAC COMMAND EXTEND OptimizeProof
-| [ "Optimize" "Proof" ] => { classify_as_proofstep } ->
- { Proof_global.compact_the_proof () }
+| ![ proof ] [ "Optimize" "Proof" ] => { classify_as_proofstep } ->
+ { fun ~pstate -> Option.map Proof_global.compact_the_proof pstate }
| [ "Optimize" "Heap" ] => { classify_as_proofstep } ->
{ Gc.compact () }
END
diff --git a/plugins/ltac/g_auto.mlg b/plugins/ltac/g_auto.mlg
index 3a4b0571d4..523c7c8305 100644
--- a/plugins/ltac/g_auto.mlg
+++ b/plugins/ltac/g_auto.mlg
@@ -58,6 +58,7 @@ let eval_uconstrs ist cs =
fail_evar = false;
expand_evars = true;
program_mode = false;
+ polymorphic = false;
} in
let map c env sigma = c env sigma in
List.map (fun c -> map (Tacinterp.type_uconstr ~flags ist c)) cs
diff --git a/plugins/ltac/g_ltac.mlg b/plugins/ltac/g_ltac.mlg
index a348e2cea4..7eb34158e8 100644
--- a/plugins/ltac/g_ltac.mlg
+++ b/plugins/ltac/g_ltac.mlg
@@ -374,20 +374,21 @@ let () = declare_int_option {
optwrite = fun n -> print_info_trace := n;
}
-let vernac_solve n info tcom b =
+let vernac_solve ~pstate n info tcom b =
let open Goal_select in
- let status = Proof_global.with_current_proof (fun etac p ->
- let with_end_tac = if b then Some etac else None in
- let global = match n with SelectAll | SelectList _ -> true | _ -> false in
- let info = Option.append info !print_info_trace in
- let (p,status) =
- Pfedit.solve n info (Tacinterp.hide_interp global tcom None) ?with_end_tac p
- in
- (* in case a strict subtree was completed,
- go back to the top of the prooftree *)
- let p = Proof.maximal_unfocus Vernacentries.command_focus p in
- p,status) in
- if not status then Feedback.feedback Feedback.AddedAxiom
+ let pstate, status = Proof_global.with_current_proof (fun etac p ->
+ let with_end_tac = if b then Some etac else None in
+ let global = match n with SelectAll | SelectList _ -> true | _ -> false in
+ let info = Option.append info !print_info_trace in
+ let (p,status) =
+ Pfedit.solve n info (Tacinterp.hide_interp global tcom None) ?with_end_tac p
+ in
+ (* in case a strict subtree was completed,
+ go back to the top of the prooftree *)
+ let p = Proof.maximal_unfocus Vernacentries.command_focus p in
+ p,status) pstate in
+ if not status then Feedback.feedback Feedback.AddedAxiom;
+ Some pstate
let pr_ltac_selector s = Pptactic.pr_goal_selector ~toplevel:true s
@@ -434,12 +435,12 @@ let is_explicit_terminator = function TacSolve _ -> true | _ -> false
}
VERNAC { tactic_mode } EXTEND VernacSolve
-| [ ltac_selector_opt(g) ltac_info_opt(n) tactic(t) ltac_use_default(def) ] =>
+| ![ proof ] [ ltac_selector_opt(g) ltac_info_opt(n) tactic(t) ltac_use_default(def) ] =>
{ classify_as_proofstep } -> {
let g = Option.default (Goal_select.get_default_goal_selector ()) g in
- vernac_solve g n t def
+ Vernacentries.vernac_require_open_proof vernac_solve g n t def
}
-| [ "par" ":" ltac_info_opt(n) tactic(t) ltac_use_default(def) ] =>
+| ![ proof ] [ "par" ":" ltac_info_opt(n) tactic(t) ltac_use_default(def) ] =>
{
let anon_abstracting_tac = is_anonymous_abstract t in
let solving_tac = is_explicit_terminator t in
@@ -449,7 +450,7 @@ VERNAC { tactic_mode } EXTEND VernacSolve
VtLater
} -> {
let t = rm_abstract t in
- vernac_solve Goal_select.SelectAll n t def
+ Vernacentries.vernac_require_open_proof vernac_solve Goal_select.SelectAll n t def
}
END
diff --git a/plugins/ltac/g_obligations.mlg b/plugins/ltac/g_obligations.mlg
index a12dee48a8..de3a9c9fa9 100644
--- a/plugins/ltac/g_obligations.mlg
+++ b/plugins/ltac/g_obligations.mlg
@@ -80,25 +80,25 @@ GRAMMAR EXTEND Gram
open Obligations
-let obligation obl tac = with_tac (fun t -> Obligations.obligation obl t) tac
-let next_obligation obl tac = with_tac (fun t -> Obligations.next_obligation obl t) tac
+let obligation ~pstate obl tac = Some (with_tac (fun t -> Obligations.obligation ~ontop:pstate obl t) tac)
+let next_obligation ~pstate obl tac = Some (with_tac (fun t -> Obligations.next_obligation ~ontop:pstate obl t) tac)
let classify_obbl _ = Vernacextend.(VtStartProof (Doesn'tGuaranteeOpacity,[]), VtLater)
}
VERNAC COMMAND EXTEND Obligations CLASSIFIED BY { classify_obbl }
-| [ "Obligation" integer(num) "of" ident(name) ":" lglob(t) withtac(tac) ] ->
+| ![ proof ] [ "Obligation" integer(num) "of" ident(name) ":" lglob(t) withtac(tac) ] ->
{ obligation (num, Some name, Some t) tac }
-| [ "Obligation" integer(num) "of" ident(name) withtac(tac) ] ->
+| ![ proof ] [ "Obligation" integer(num) "of" ident(name) withtac(tac) ] ->
{ obligation (num, Some name, None) tac }
-| [ "Obligation" integer(num) ":" lglob(t) withtac(tac) ] ->
+| ![ proof ] [ "Obligation" integer(num) ":" lglob(t) withtac(tac) ] ->
{ obligation (num, None, Some t) tac }
-| [ "Obligation" integer(num) withtac(tac) ] ->
+| ![ proof ] [ "Obligation" integer(num) withtac(tac) ] ->
{ obligation (num, None, None) tac }
-| [ "Next" "Obligation" "of" ident(name) withtac(tac) ] ->
+| ![ proof ] [ "Next" "Obligation" "of" ident(name) withtac(tac) ] ->
{ next_obligation (Some name) tac }
-| [ "Next" "Obligation" withtac(tac) ] -> { next_obligation None tac }
+| ![ proof ] [ "Next" "Obligation" withtac(tac) ] -> { next_obligation None tac }
END
VERNAC COMMAND EXTEND Solve_Obligation CLASSIFIED AS SIDEFF
diff --git a/plugins/ltac/g_rewrite.mlg b/plugins/ltac/g_rewrite.mlg
index 86a227415a..469551809c 100644
--- a/plugins/ltac/g_rewrite.mlg
+++ b/plugins/ltac/g_rewrite.mlg
@@ -180,34 +180,34 @@ TACTIC EXTEND setoid_rewrite
END
VERNAC COMMAND EXTEND AddRelation CLASSIFIED AS SIDEFF
- | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
"symmetry" "proved" "by" constr(lemma2) "as" ident(n) ] ->
{ declare_relation atts a aeq n (Some lemma1) (Some lemma2) None }
- | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
"as" ident(n) ] ->
{ declare_relation atts a aeq n (Some lemma1) None None }
- | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "as" ident(n) ] ->
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Relation" constr(a) constr(aeq) "as" ident(n) ] ->
{ declare_relation atts a aeq n None None None }
END
VERNAC COMMAND EXTEND AddRelation2 CLASSIFIED AS SIDEFF
- | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2)
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Relation" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2)
"as" ident(n) ] ->
{ declare_relation atts a aeq n None (Some lemma2) None }
- | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] ->
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Relation" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] ->
{ declare_relation atts a aeq n None (Some lemma2) (Some lemma3) }
END
VERNAC COMMAND EXTEND AddRelation3 CLASSIFIED AS SIDEFF
- | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
"transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] ->
{ declare_relation atts a aeq n (Some lemma1) None (Some lemma3) }
- | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
"symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3)
"as" ident(n) ] ->
{ declare_relation atts a aeq n (Some lemma1) (Some lemma2) (Some lemma3) }
- | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "transitivity" "proved" "by" constr(lemma3)
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Relation" constr(a) constr(aeq) "transitivity" "proved" "by" constr(lemma3)
"as" ident(n) ] ->
{ declare_relation atts a aeq n None None (Some lemma3) }
END
@@ -234,64 +234,64 @@ GRAMMAR EXTEND Gram
END
VERNAC COMMAND EXTEND AddParametricRelation CLASSIFIED AS SIDEFF
- | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq)
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq)
"reflexivity" "proved" "by" constr(lemma1)
"symmetry" "proved" "by" constr(lemma2) "as" ident(n) ] ->
{ declare_relation atts ~binders:b a aeq n (Some lemma1) (Some lemma2) None }
- | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq)
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq)
"reflexivity" "proved" "by" constr(lemma1)
"as" ident(n) ] ->
{ declare_relation atts ~binders:b a aeq n (Some lemma1) None None }
- | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "as" ident(n) ] ->
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "as" ident(n) ] ->
{ declare_relation atts ~binders:b a aeq n None None None }
END
VERNAC COMMAND EXTEND AddParametricRelation2 CLASSIFIED AS SIDEFF
- | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2)
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2)
"as" ident(n) ] ->
{ declare_relation atts ~binders:b a aeq n None (Some lemma2) None }
- | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] ->
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] ->
{ declare_relation atts ~binders:b a aeq n None (Some lemma2) (Some lemma3) }
END
VERNAC COMMAND EXTEND AddParametricRelation3 CLASSIFIED AS SIDEFF
- | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
"transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] ->
{ declare_relation atts ~binders:b a aeq n (Some lemma1) None (Some lemma3) }
- | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
"symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3)
"as" ident(n) ] ->
{ declare_relation atts ~binders:b a aeq n (Some lemma1) (Some lemma2) (Some lemma3) }
- | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "transitivity" "proved" "by" constr(lemma3)
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "transitivity" "proved" "by" constr(lemma3)
"as" ident(n) ] ->
{ declare_relation atts ~binders:b a aeq n None None (Some lemma3) }
END
VERNAC COMMAND EXTEND AddSetoid1 CLASSIFIED AS SIDEFF
- | #[ atts = rewrite_attributes; ] [ "Add" "Setoid" constr(a) constr(aeq) constr(t) "as" ident(n) ] ->
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Setoid" constr(a) constr(aeq) constr(t) "as" ident(n) ] ->
{
- add_setoid atts [] a aeq t n;
+ add_setoid atts [] a aeq t n
}
- | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Setoid" binders(binders) ":" constr(a) constr(aeq) constr(t) "as" ident(n) ] ->
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Parametric" "Setoid" binders(binders) ":" constr(a) constr(aeq) constr(t) "as" ident(n) ] ->
{
- add_setoid atts binders a aeq t n;
+ add_setoid atts binders a aeq t n
}
- | #[ atts = rewrite_attributes; ] [ "Add" "Morphism" constr(m) ":" ident(n) ]
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Morphism" constr(m) ":" ident(n) ]
(* This command may or may not open a goal *)
=> { VtUnknown, VtNow }
-> {
- add_morphism_infer atts m n;
+ add_morphism_infer atts m n
}
- | #[ atts = rewrite_attributes; ] [ "Add" "Morphism" constr(m) "with" "signature" lconstr(s) "as" ident(n) ]
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Morphism" constr(m) "with" "signature" lconstr(s) "as" ident(n) ]
=> { VtStartProof(GuaranteesOpacity,[n]), VtLater }
-> {
- add_morphism atts [] m s n;
+ add_morphism atts [] m s n
}
- | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Morphism" binders(binders) ":" constr(m)
+ | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Parametric" "Morphism" binders(binders) ":" constr(m)
"with" "signature" lconstr(s) "as" ident(n) ]
=> { VtStartProof(GuaranteesOpacity,[n]), VtLater }
-> {
- add_morphism atts binders m s n;
+ add_morphism atts binders m s n
}
END
@@ -310,7 +310,12 @@ TACTIC EXTEND setoid_transitivity
END
VERNAC COMMAND EXTEND PrintRewriteHintDb CLASSIFIED AS QUERY
-| [ "Print" "Rewrite" "HintDb" preident(s) ] ->
- { let sigma, env = Pfedit.get_current_context () in
- Feedback.msg_notice (Autorewrite.print_rewrite_hintdb env sigma s) }
+| ![ proof ] [ "Print" "Rewrite" "HintDb" preident(s) ] ->
+ { (* This command should not use the proof env, keeping previous
+ behavior as requested in review. *)
+ fun ~pstate ->
+ let sigma, env = Option.cata Pfedit.get_current_context
+ (let e = Global.env () in Evd.from_env e, e) pstate in
+ Feedback.msg_notice (Autorewrite.print_rewrite_hintdb env sigma s);
+ pstate }
END
diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml
index 1bdba699f7..80070a7493 100644
--- a/plugins/ltac/pptactic.ml
+++ b/plugins/ltac/pptactic.ml
@@ -1307,7 +1307,6 @@ let lift_top f a = Genprint.TopPrinterBasic (fun () -> f a)
let register_basic_print0 wit f g h =
Genprint.register_print0 wit (lift f) (lift g) (lift_top h)
-
let pr_glob_constr_pptac env sigma c =
pr_glob_constr_env env c
diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml
index b1d5c0252f..75565c1a34 100644
--- a/plugins/ltac/rewrite.ml
+++ b/plugins/ltac/rewrite.ml
@@ -618,7 +618,9 @@ let solve_remaining_by env sigma holes by =
in
(* Only solve independent holes *)
let indep = List.map_filter map holes in
- let ist = { Geninterp.lfun = Id.Map.empty; extra = Geninterp.TacStore.empty } in
+ let ist = { Geninterp.lfun = Id.Map.empty
+ ; poly = false
+ ; extra = Geninterp.TacStore.empty } in
let solve_tac = match tac with
| Genarg.GenArg (Genarg.Glbwit tag, tac) ->
Ftactic.run (Geninterp.interp tag ist tac) (fun _ -> Proofview.tclUNIT ())
@@ -1790,15 +1792,15 @@ let declare_an_instance n s args =
let declare_instance a aeq n s = declare_an_instance n s [a;aeq]
-let anew_instance atts binders instance fields =
+let anew_instance ~pstate atts binders instance fields =
let program_mode = atts.program in
- new_instance ~program_mode atts.polymorphic
+ new_instance ~pstate ~program_mode atts.polymorphic
binders instance (Some (true, CAst.make @@ CRecord (fields)))
~global:atts.global ~generalize:false ~refine:false Hints.empty_hint_info
-let declare_instance_refl atts binders a aeq n lemma =
+let declare_instance_refl ~pstate atts binders a aeq n lemma =
let instance = declare_instance a aeq (add_suffix n "_Reflexive") "Coq.Classes.RelationClasses.Reflexive"
- in anew_instance atts binders instance
+ in anew_instance ~pstate atts binders instance
[(qualid_of_ident (Id.of_string "reflexivity"),lemma)]
let declare_instance_sym atts binders a aeq n lemma =
@@ -1811,47 +1813,44 @@ let declare_instance_trans atts binders a aeq n lemma =
in anew_instance atts binders instance
[(qualid_of_ident (Id.of_string "transitivity"),lemma)]
-let declare_relation atts ?(binders=[]) a aeq n refl symm trans =
+let declare_relation ~pstate atts ?(binders=[]) a aeq n refl symm trans =
init_setoid ();
- let instance = declare_instance a aeq (add_suffix n "_relation") "Coq.Classes.RelationClasses.RewriteRelation"
- in ignore(anew_instance atts binders instance []);
+ let instance = declare_instance a aeq (add_suffix n "_relation") "Coq.Classes.RelationClasses.RewriteRelation" in
+ let _, pstate = anew_instance ~pstate atts binders instance [] in
match (refl,symm,trans) with
- (None, None, None) -> ()
+ (None, None, None) -> pstate
| (Some lemma1, None, None) ->
- ignore (declare_instance_refl atts binders a aeq n lemma1)
+ snd @@ declare_instance_refl ~pstate atts binders a aeq n lemma1
| (None, Some lemma2, None) ->
- ignore (declare_instance_sym atts binders a aeq n lemma2)
+ snd @@ declare_instance_sym ~pstate atts binders a aeq n lemma2
| (None, None, Some lemma3) ->
- ignore (declare_instance_trans atts binders a aeq n lemma3)
+ snd @@ declare_instance_trans ~pstate atts binders a aeq n lemma3
| (Some lemma1, Some lemma2, None) ->
- ignore (declare_instance_refl atts binders a aeq n lemma1);
- ignore (declare_instance_sym atts binders a aeq n lemma2)
+ let _lemma_refl, pstate = declare_instance_refl ~pstate atts binders a aeq n lemma1 in
+ snd @@ declare_instance_sym ~pstate atts binders a aeq n lemma2
| (Some lemma1, None, Some lemma3) ->
- let _lemma_refl = declare_instance_refl atts binders a aeq n lemma1 in
- let _lemma_trans = declare_instance_trans atts binders a aeq n lemma3 in
- let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.PreOrder"
- in ignore(
- anew_instance atts binders instance
+ let _lemma_refl, pstate = declare_instance_refl ~pstate atts binders a aeq n lemma1 in
+ let _lemma_trans, pstate = declare_instance_trans ~pstate atts binders a aeq n lemma3 in
+ let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.PreOrder" in
+ snd @@ anew_instance ~pstate atts binders instance
[(qualid_of_ident (Id.of_string "PreOrder_Reflexive"), lemma1);
- (qualid_of_ident (Id.of_string "PreOrder_Transitive"),lemma3)])
+ (qualid_of_ident (Id.of_string "PreOrder_Transitive"),lemma3)]
| (None, Some lemma2, Some lemma3) ->
- let _lemma_sym = declare_instance_sym atts binders a aeq n lemma2 in
- let _lemma_trans = declare_instance_trans atts binders a aeq n lemma3 in
- let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.PER"
- in ignore(
- anew_instance atts binders instance
+ let _lemma_sym, pstate = declare_instance_sym ~pstate atts binders a aeq n lemma2 in
+ let _lemma_trans, pstate = declare_instance_trans ~pstate atts binders a aeq n lemma3 in
+ let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.PER" in
+ snd @@ anew_instance ~pstate atts binders instance
[(qualid_of_ident (Id.of_string "PER_Symmetric"), lemma2);
- (qualid_of_ident (Id.of_string "PER_Transitive"),lemma3)])
+ (qualid_of_ident (Id.of_string "PER_Transitive"),lemma3)]
| (Some lemma1, Some lemma2, Some lemma3) ->
- let _lemma_refl = declare_instance_refl atts binders a aeq n lemma1 in
- let _lemma_sym = declare_instance_sym atts binders a aeq n lemma2 in
- let _lemma_trans = declare_instance_trans atts binders a aeq n lemma3 in
- let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.Equivalence"
- in ignore(
- anew_instance atts binders instance
+ let _lemma_refl, pstate = declare_instance_refl ~pstate atts binders a aeq n lemma1 in
+ let _lemma_sym, pstate = declare_instance_sym ~pstate atts binders a aeq n lemma2 in
+ let _lemma_trans, pstate = declare_instance_trans ~pstate atts binders a aeq n lemma3 in
+ let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.Equivalence" in
+ snd @@ anew_instance ~pstate atts binders instance
[(qualid_of_ident (Id.of_string "Equivalence_Reflexive"), lemma1);
(qualid_of_ident (Id.of_string "Equivalence_Symmetric"), lemma2);
- (qualid_of_ident (Id.of_string "Equivalence_Transitive"), lemma3)])
+ (qualid_of_ident (Id.of_string "Equivalence_Transitive"), lemma3)]
let cHole = CAst.make @@ CHole (None, Namegen.IntroAnonymous, None)
@@ -1947,18 +1946,18 @@ let warn_add_setoid_deprecated =
CWarnings.create ~name:"add-setoid" ~category:"deprecated" (fun () ->
Pp.(str "Add Setoid is deprecated, please use Add Parametric Relation."))
-let add_setoid atts binders a aeq t n =
+let add_setoid ~pstate atts binders a aeq t n =
warn_add_setoid_deprecated ?loc:a.CAst.loc ();
init_setoid ();
- let _lemma_refl = declare_instance_refl atts binders a aeq n (mkappc "Seq_refl" [a;aeq;t]) in
- let _lemma_sym = declare_instance_sym atts binders a aeq n (mkappc "Seq_sym" [a;aeq;t]) in
- let _lemma_trans = declare_instance_trans atts binders a aeq n (mkappc "Seq_trans" [a;aeq;t]) in
+ let _lemma_refl, pstate = declare_instance_refl ~pstate atts binders a aeq n (mkappc "Seq_refl" [a;aeq;t]) in
+ let _lemma_sym, pstate = declare_instance_sym ~pstate atts binders a aeq n (mkappc "Seq_sym" [a;aeq;t]) in
+ let _lemma_trans, pstate = declare_instance_trans ~pstate atts binders a aeq n (mkappc "Seq_trans" [a;aeq;t]) in
let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.Equivalence"
- in ignore(
- anew_instance atts binders instance
+ in
+ snd @@ anew_instance ~pstate atts binders instance
[(qualid_of_ident (Id.of_string "Equivalence_Reflexive"), mkappc "Seq_refl" [a;aeq;t]);
(qualid_of_ident (Id.of_string "Equivalence_Symmetric"), mkappc "Seq_sym" [a;aeq;t]);
- (qualid_of_ident (Id.of_string "Equivalence_Transitive"), mkappc "Seq_trans" [a;aeq;t])])
+ (qualid_of_ident (Id.of_string "Equivalence_Transitive"), mkappc "Seq_trans" [a;aeq;t])]
let make_tactic name =
@@ -1970,7 +1969,7 @@ let warn_add_morphism_deprecated =
CWarnings.create ~name:"add-morphism" ~category:"deprecated" (fun () ->
Pp.(str "Add Morphism f : id is deprecated, please use Add Morphism f with signature (...) as id"))
-let add_morphism_infer atts m n =
+let add_morphism_infer ~pstate atts m n : Proof_global.t option =
warn_add_morphism_deprecated ?loc:m.CAst.loc ();
init_setoid ();
(* NB: atts.program is ignored, program mode automatically set by vernacentries *)
@@ -1981,45 +1980,47 @@ let add_morphism_infer atts m n =
if Lib.is_modtype () then
let uctx = UState.univ_entry ~poly:atts.polymorphic uctx in
let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest instance_id
- (Entries.ParameterEntry
- (None,(instance,uctx),None),
- Decl_kinds.IsAssumption Decl_kinds.Logical)
+ (Entries.ParameterEntry
+ (None,(instance,uctx),None),
+ Decl_kinds.IsAssumption Decl_kinds.Logical)
in
- add_instance (Typeclasses.new_instance
- (Lazy.force PropGlobal.proper_class) Hints.empty_hint_info atts.global (ConstRef cst));
- declare_projection n instance_id (ConstRef cst)
+ add_instance (Typeclasses.new_instance
+ (Lazy.force PropGlobal.proper_class) Hints.empty_hint_info atts.global (ConstRef cst));
+ declare_projection n instance_id (ConstRef cst);
+ pstate
else
let kind = Decl_kinds.Global, atts.polymorphic,
- Decl_kinds.DefinitionBody Decl_kinds.Instance
+ Decl_kinds.DefinitionBody Decl_kinds.Instance
in
let tac = make_tactic "Coq.Classes.SetoidTactics.add_morphism_tactic" in
let hook _ _ _ = function
- | Globnames.ConstRef cst ->
- add_instance (Typeclasses.new_instance
- (Lazy.force PropGlobal.proper_class) Hints.empty_hint_info
+ | Globnames.ConstRef cst ->
+ add_instance (Typeclasses.new_instance
+ (Lazy.force PropGlobal.proper_class) Hints.empty_hint_info
atts.global (ConstRef cst));
- declare_projection n instance_id (ConstRef cst)
- | _ -> assert false
+ declare_projection n instance_id (ConstRef cst)
+ | _ -> assert false
in
let hook = Lemmas.mk_hook hook in
- Flags.silently
- (fun () ->
- Lemmas.start_proof ~hook instance_id kind (Evd.from_ctx uctx) (EConstr.of_constr instance);
- ignore (Pfedit.by (Tacinterp.interp tac))) ()
+ Flags.silently
+ (fun () ->
+ let pstate = Lemmas.start_proof ~ontop:pstate ~hook instance_id kind (Evd.from_ctx uctx) (EConstr.of_constr instance) in
+ Some (fst Pfedit.(by (Tacinterp.interp tac) pstate))) ()
-let add_morphism atts binders m s n =
+let add_morphism ~pstate atts binders m s n =
init_setoid ();
let instance_id = add_suffix n "_Proper" in
let instance =
(((CAst.make @@ Name instance_id),None), Explicit,
CAst.make @@ CAppExpl (
(None, Libnames.qualid_of_string "Coq.Classes.Morphisms.Proper",None),
- [cHole; s; m]))
+ [cHole; s; m]))
in
let tac = Tacinterp.interp (make_tactic "add_morphism_tactic") in
- ignore(new_instance ~program_mode:atts.program ~global:atts.global atts.polymorphic binders instance
- None
- ~generalize:false ~tac ~hook:(declare_projection n instance_id) Hints.empty_hint_info)
+ let _, pstate = new_instance ~pstate ~program_mode:atts.program ~global:atts.global atts.polymorphic binders instance
+ None
+ ~generalize:false ~tac ~hook:(declare_projection n instance_id) Hints.empty_hint_info in
+ pstate
(** Bind to "rewrite" too *)
diff --git a/plugins/ltac/rewrite.mli b/plugins/ltac/rewrite.mli
index 2457b265f0..a200cb5ced 100644
--- a/plugins/ltac/rewrite.mli
+++ b/plugins/ltac/rewrite.mli
@@ -81,18 +81,18 @@ val cl_rewrite_clause :
val is_applied_rewrite_relation :
env -> evar_map -> rel_context -> constr -> types option
-val declare_relation : rewrite_attributes ->
+val declare_relation : pstate:Proof_global.t option -> rewrite_attributes ->
?binders:local_binder_expr list -> constr_expr -> constr_expr -> Id.t ->
- constr_expr option -> constr_expr option -> constr_expr option -> unit
+ constr_expr option -> constr_expr option -> constr_expr option -> Proof_global.t option
-val add_setoid :
+val add_setoid : pstate:Proof_global.t option ->
rewrite_attributes -> local_binder_expr list -> constr_expr -> constr_expr -> constr_expr ->
- Id.t -> unit
+ Id.t -> Proof_global.t option
-val add_morphism_infer : rewrite_attributes -> constr_expr -> Id.t -> unit
+val add_morphism_infer : pstate:Proof_global.t option -> rewrite_attributes -> constr_expr -> Id.t -> Proof_global.t option
-val add_morphism :
- rewrite_attributes -> local_binder_expr list -> constr_expr -> constr_expr -> Id.t -> unit
+val add_morphism : pstate:Proof_global.t option ->
+ rewrite_attributes -> local_binder_expr list -> constr_expr -> constr_expr -> Id.t -> Proof_global.t option
val get_reflexive_proof : env -> evar_map -> constr -> constr -> evar_map * constr
diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml
index eac84f0543..4398fb14ab 100644
--- a/plugins/ltac/tacinterp.ml
+++ b/plugins/ltac/tacinterp.ml
@@ -138,9 +138,10 @@ let f_debug : debug_info TacStore.field = TacStore.field ()
let f_trace : ltac_trace TacStore.field = TacStore.field ()
(* Signature for interpretation: val_interp and interpretation functions *)
-type interp_sign = Geninterp.interp_sign = {
- lfun : value Id.Map.t;
- extra : TacStore.t }
+type interp_sign = Geninterp.interp_sign =
+ { lfun : value Id.Map.t
+ ; poly : bool
+ ; extra : TacStore.t }
let extract_trace ist =
if is_traced () then match TacStore.get ist.extra f_trace with
@@ -544,12 +545,7 @@ let interp_gen kind ist pattern_mode flags env sigma c =
let (_, dummy_proofview) = Proofview.init sigma [] in
(* Again this is called at times with no open proof! *)
- let name, poly =
- try
- let Proof.{ name; poly } = Proof.data Proof_global.(give_me_the_proof ()) in
- name, poly
- with | Proof_global.NoCurrentProof -> Id.of_string "tacinterp", false
- in
+ let name, poly = Id.of_string "tacinterp", ist.poly in
let (trace,_,_,_) = Proofview.apply ~name ~poly env (push_trace (loc_of_glob_constr term,LtacConstrInterp (term,vars)) ist) dummy_proofview in
let (evd,c) =
catch_error trace (understand_ltac flags env sigma vars kind) term
@@ -566,11 +562,13 @@ let constr_flags () = {
fail_evar = true;
expand_evars = true;
program_mode = false;
+ polymorphic = false;
}
(* Interprets a constr; expects evars to be solved *)
let interp_constr_gen kind ist env sigma c =
- interp_gen kind ist false (constr_flags ()) env sigma c
+ let flags = { (constr_flags ()) with polymorphic = ist.Geninterp.poly } in
+ interp_gen kind ist false flags env sigma c
let interp_constr = interp_constr_gen WithoutTypeConstraint
@@ -582,6 +580,7 @@ let open_constr_use_classes_flags () = {
fail_evar = false;
expand_evars = true;
program_mode = false;
+ polymorphic = false;
}
let open_constr_no_classes_flags () = {
@@ -590,6 +589,7 @@ let open_constr_no_classes_flags () = {
fail_evar = false;
expand_evars = true;
program_mode = false;
+ polymorphic = false;
}
let pure_open_constr_flags = {
@@ -598,6 +598,7 @@ let pure_open_constr_flags = {
fail_evar = false;
expand_evars = false;
program_mode = false;
+ polymorphic = false;
}
(* Interprets an open constr *)
@@ -1021,6 +1022,7 @@ let type_uconstr ?(flags = (constr_flags ()))
ltac_idents = closure.idents;
ltac_genargs = Id.Map.empty;
} in
+ let flags = { flags with polymorphic = ist.Geninterp.poly } in
understand_ltac flags env sigma vars expected_type term
end
@@ -1146,6 +1148,7 @@ and eval_tactic ist tac : unit Proofview.tactic = match tac with
(* For extensions *)
| TacAlias {loc; v=(s,l)} ->
let alias = Tacenv.interp_alias s in
+ Proofview.tclProofInfo [@ocaml.warning "-3"] >>= fun (_name, poly) ->
let (>>=) = Ftactic.bind in
let interp_vars = Ftactic.List.map (fun v -> interp_tacarg ist v) l in
let tac l =
@@ -1153,8 +1156,9 @@ and eval_tactic ist tac : unit Proofview.tactic = match tac with
let lfun = List.fold_right2 addvar alias.Tacenv.alias_args l ist.lfun in
Ftactic.lift (push_trace (loc,LtacNotationCall s) ist) >>= fun trace ->
let ist = {
- lfun = lfun;
- extra = TacStore.set ist.extra f_trace trace; } in
+ lfun
+ ; poly
+ ; extra = TacStore.set ist.extra f_trace trace } in
val_interp ist alias.Tacenv.alias_body >>= fun v ->
Ftactic.lift (tactic_of_value ist v)
in
@@ -1207,12 +1211,13 @@ and interp_ltac_reference ?loc' mustbetac ist r : Val.t Ftactic.t =
if mustbetac then Ftactic.return (coerce_to_tactic loc id v) else Ftactic.return v
end
| ArgArg (loc,r) ->
+ Proofview.tclProofInfo [@ocaml.warning "-3"] >>= fun (_name, poly) ->
let ids = extract_ids [] ist.lfun Id.Set.empty in
let loc_info = (Option.default loc loc',LtacNameCall r) in
let extra = TacStore.set ist.extra f_avoid_ids ids in
push_trace loc_info ist >>= fun trace ->
let extra = TacStore.set extra f_trace trace in
- let ist = { lfun = Id.Map.empty; extra = extra; } in
+ let ist = { lfun = Id.Map.empty; poly; extra } in
let appl = GlbAppl[r,[]] in
Profile_ltac.do_profile "interp_ltac_reference" trace ~count_call:false
(val_interp ~appl ist (Tacenv.interp_ltac r))
@@ -1260,6 +1265,7 @@ and interp_tacarg ist arg : Val.t Ftactic.t =
(* Interprets an application node *)
and interp_app loc ist fv largs : Val.t Ftactic.t =
+ Proofview.tclProofInfo [@ocaml.warning "-3"] >>= fun (_name, poly) ->
let (>>=) = Ftactic.bind in
let fail = Tacticals.New.tclZEROMSG (str "Illegal tactic application.") in
if has_type fv (topwit wit_tacvalue) then
@@ -1277,9 +1283,11 @@ and interp_app loc ist fv largs : Val.t Ftactic.t =
if List.is_empty lvar then
begin wrap_error
begin
- let ist = {
- lfun = newlfun;
- extra = TacStore.set ist.extra f_trace []; } in
+ let ist =
+ { lfun = newlfun
+ ; poly
+ ; extra = TacStore.set ist.extra f_trace []
+ } in
Profile_ltac.do_profile "interp_app" trace ~count_call:false
(catch_error_tac trace (val_interp ist body)) >>= fun v ->
Ftactic.return (name_vfun (push_appl appl largs) v)
@@ -1317,8 +1325,10 @@ and tactic_of_value ist vle =
if has_type vle (topwit wit_tacvalue) then
match to_tacvalue vle with
| VFun (appl,trace,lfun,[],t) ->
+ Proofview.tclProofInfo [@ocaml.warning "-3"] >>= fun (_name, poly) ->
let ist = {
lfun = lfun;
+ poly;
extra = TacStore.set ist.extra f_trace []; } in
let tac = name_if_glob appl (eval_tactic ist t) in
Profile_ltac.do_profile "tactic_of_value" trace (catch_error_tac trace tac)
@@ -1388,6 +1398,7 @@ and interp_letin ist llc u =
(** [interp_match_success lz ist succ] interprets a single matching success
(of type {!Tactic_matching.t}). *)
and interp_match_success ist { Tactic_matching.subst ; context ; terms ; lhs } =
+ Proofview.tclProofInfo [@ocaml.warning "-3"] >>= fun (_name, poly) ->
let (>>=) = Ftactic.bind in
let lctxt = Id.Map.map interp_context context in
let hyp_subst = Id.Map.map Value.of_constr terms in
@@ -1396,9 +1407,11 @@ and interp_match_success ist { Tactic_matching.subst ; context ; terms ; lhs } =
val_interp ist lhs >>= fun v ->
if has_type v (topwit wit_tacvalue) then match to_tacvalue v with
| VFun (appl,trace,lfun,[],t) ->
- let ist = {
- lfun = lfun;
- extra = TacStore.set ist.extra f_trace trace; } in
+ let ist =
+ { lfun = lfun
+ ; poly
+ ; extra = TacStore.set ist.extra f_trace trace
+ } in
let tac = eval_tactic ist t in
let dummy = VFun (appl,extract_trace ist, Id.Map.empty, [], TacId []) in
catch_error_tac trace (tac <*> Ftactic.return (of_tacvalue dummy))
@@ -1872,7 +1885,7 @@ and interp_atomic ist tac : unit Proofview.tactic =
let default_ist () =
let extra = TacStore.set TacStore.empty f_debug (get_debug ()) in
- { lfun = Id.Map.empty; extra = extra }
+ { lfun = Id.Map.empty; poly = false; extra = extra }
let eval_tactic t =
Proofview.tclUNIT () >>= fun () -> (* delay for [default_ist] *)
@@ -1912,11 +1925,12 @@ end
let interp_tac_gen lfun avoid_ids debug t =
+ Proofview.tclProofInfo [@ocaml.warning "-3"] >>= fun (_name, poly) ->
Proofview.Goal.enter begin fun gl ->
let env = Proofview.Goal.env gl in
let extra = TacStore.set TacStore.empty f_debug debug in
let extra = TacStore.set extra f_avoid_ids avoid_ids in
- let ist = { lfun = lfun; extra = extra } in
+ let ist = { lfun; poly; extra } in
let ltacvars = Id.Map.domain lfun in
interp_tactic ist
(intern_pure_tactic { (Genintern.empty_glob_sign env) with ltacvars } t)
@@ -2057,20 +2071,15 @@ let interp_redexp env sigma r =
(* Backwarding recursive needs of tactic glob/interp/eval functions *)
let _ =
- let eval lfun env sigma ty tac =
+ let eval lfun poly env sigma ty tac =
let extra = TacStore.set TacStore.empty f_debug (get_debug ()) in
- let ist = { lfun = lfun; extra; } in
+ let ist = { lfun; poly; extra; } in
let tac = interp_tactic ist tac in
- (* XXX: This depends on the global state which is bad; the hooking
- mechanism should be modified. *)
- let name, poly =
- try
- let (_, poly, _) = Proof_global.get_current_persistence () in
- let name = Proof_global.get_current_proof_name () in
- name, poly
- with | Proof_global.NoCurrentProof ->
- Id.of_string "ltac_gen", false
- in
+ (* EJGA: We sould also pass the proof name if desired, for now
+ poly seems like enough to get reasonable behavior in practice
+ *)
+ let name, poly = Id.of_string "ltac_gen", poly in
+ let name, poly = Id.of_string "ltac_gen", poly in
let (c, sigma) = Pfedit.refine_by_tactic ~name ~poly env sigma ty tac in
(EConstr.of_constr c, sigma)
in
diff --git a/plugins/ltac/tacinterp.mli b/plugins/ltac/tacinterp.mli
index d9c80bb835..22a092fa8b 100644
--- a/plugins/ltac/tacinterp.mli
+++ b/plugins/ltac/tacinterp.mli
@@ -39,9 +39,10 @@ module TacStore : Store.S with
and type 'a field = 'a Geninterp.TacStore.field
(** Signature for interpretation: val\_interp and interpretation functions *)
-type interp_sign = Geninterp.interp_sign = {
- lfun : value Id.Map.t;
- extra : TacStore.t }
+type interp_sign = Geninterp.interp_sign =
+ { lfun : value Id.Map.t
+ ; poly : bool
+ ; extra : TacStore.t }
open Genintern
diff --git a/plugins/ltac/tactic_debug.ml b/plugins/ltac/tactic_debug.ml
index 52a83a038f..04f3116664 100644
--- a/plugins/ltac/tactic_debug.ml
+++ b/plugins/ltac/tactic_debug.ml
@@ -370,7 +370,10 @@ let explain_ltac_call_trace last trace loc =
strbrk " (with " ++
prlist_with_sep pr_comma
(fun (id,c) ->
- let sigma, env = Pfedit.get_current_context () in
+ (* XXX: This hooks into the ExplainErr extension API
+ so it is tricky to provide the right env for now. *)
+ let env = Global.env () in
+ let sigma = Evd.from_env env in
Id.print id ++ str ":=" ++ Printer.pr_lconstr_under_binders_env env sigma c)
(List.rev (Id.Map.bindings vars)) ++ str ")"
else mt())
diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml
index 7db47e13a5..6c04fe9a8a 100644
--- a/plugins/micromega/coq_micromega.ml
+++ b/plugins/micromega/coq_micromega.ml
@@ -877,9 +877,9 @@ struct
* This is the big generic function for expression parsers.
*)
- let parse_expr cenv sigma parse_constant parse_exp ops_spec env term =
+ let parse_expr env sigma parse_constant parse_exp ops_spec term_env term =
if debug
- then Feedback.msg_debug (Pp.str "parse_expr: " ++ Printer.pr_leconstr_env cenv sigma term);
+ then Feedback.msg_debug (Pp.str "parse_expr: " ++ Printer.pr_leconstr_env env sigma term);
(*
let constant_or_variable env term =
@@ -928,7 +928,7 @@ struct
| _ -> parse_variable env term
)
| _ -> parse_variable env term in
- parse_expr env term
+ parse_expr term_env term
let zop_spec =
[
@@ -1007,7 +1007,7 @@ struct
res
- let parse_zexpr env sigma = parse_expr env sigma
+ let parse_zexpr env sigma = parse_expr env sigma
(zconstant sigma)
(fun expr x ->
let exp = (parse_z sigma x) in
@@ -1038,16 +1038,17 @@ struct
Mc.PEpow(expr,exp))
rop_spec
- let parse_arith parse_op parse_expr env cstr gl =
+ let parse_arith parse_op parse_expr term_env cstr gl =
let sigma = gl.sigma in
+ let env = gl.env in
if debug
- then Feedback.msg_debug (Pp.str "parse_arith: " ++ Printer.pr_leconstr_env gl.env sigma cstr ++ fnl ());
+ then Feedback.msg_debug (Pp.str "parse_arith: " ++ Printer.pr_leconstr_env env sigma cstr ++ fnl ());
match EConstr.kind sigma cstr with
| App(op,args) ->
let (op,lhs,rhs) = parse_op gl (op,args) in
- let (e1,env) = parse_expr gl.env sigma env lhs in
- let (e2,env) = parse_expr gl.env sigma env rhs in
- ({Mc.flhs = e1; Mc.fop = op;Mc.frhs = e2},env)
+ let (e1,term_env) = parse_expr env sigma term_env lhs in
+ let (e2,term_env) = parse_expr env sigma term_env rhs in
+ ({Mc.flhs = e1; Mc.fop = op;Mc.frhs = e2},term_env)
| _ -> failwith "error : parse_arith(2)"
let parse_zarith = parse_arith parse_zop parse_zexpr
diff --git a/plugins/setoid_ring/g_newring.mlg b/plugins/setoid_ring/g_newring.mlg
index 3ce6478700..6be556b2ae 100644
--- a/plugins/setoid_ring/g_newring.mlg
+++ b/plugins/setoid_ring/g_newring.mlg
@@ -86,15 +86,20 @@ END
VERNAC COMMAND EXTEND AddSetoidRing CLASSIFIED AS SIDEFF
| [ "Add" "Ring" ident(id) ":" constr(t) ring_mods_opt(l) ] ->
{ let l = match l with None -> [] | Some l -> l in add_theory id t l }
- | [ "Print" "Rings" ] => { Vernacextend.classify_as_query } -> {
+ | ![proof] [ "Print" "Rings" ] => { Vernacextend.classify_as_query } -> {
+ fun ~pstate ->
Feedback.msg_notice (strbrk "The following ring structures have been declared:");
Spmap.iter (fun fn fi ->
- let sigma, env = Pfedit.get_current_context () in
+ (* We should use the global env here as this shouldn't contain proof
+ data, however preserving behavior as requested in review. *)
+ let sigma, env = Option.cata Pfedit.get_current_context
+ (let e = Global.env () in Evd.from_env e, e) pstate in
Feedback.msg_notice (hov 2
(Ppconstr.pr_id (Libnames.basename fn)++spc()++
str"with carrier "++ pr_constr_env env sigma fi.ring_carrier++spc()++
str"and equivalence relation "++ pr_constr_env env sigma fi.ring_req))
- ) !from_name }
+ ) !from_name;
+ pstate }
END
TACTIC EXTEND ring_lookup
@@ -130,15 +135,20 @@ END
VERNAC COMMAND EXTEND AddSetoidField CLASSIFIED AS SIDEFF
| [ "Add" "Field" ident(id) ":" constr(t) field_mods_opt(l) ] ->
{ let l = match l with None -> [] | Some l -> l in add_field_theory id t l }
-| [ "Print" "Fields" ] => {Vernacextend.classify_as_query} -> {
+| ![proof] [ "Print" "Fields" ] => {Vernacextend.classify_as_query} -> {
+ fun ~pstate ->
Feedback.msg_notice (strbrk "The following field structures have been declared:");
Spmap.iter (fun fn fi ->
- let sigma, env = Pfedit.get_current_context () in
+ (* We should use the global env here as this shouldn't
+ contain proof data. *)
+ let sigma, env = Option.cata Pfedit.get_current_context
+ (let e = Global.env () in Evd.from_env e, e) pstate in
Feedback.msg_notice (hov 2
(Ppconstr.pr_id (Libnames.basename fn)++spc()++
str"with carrier "++ pr_constr_env env sigma fi.field_carrier++spc()++
str"and equivalence relation "++ pr_constr_env env sigma fi.field_req))
- ) !field_from_name }
+ ) !field_from_name;
+ pstate }
END
TACTIC EXTEND field_lookup
diff --git a/plugins/ssr/ssrcommon.ml b/plugins/ssr/ssrcommon.ml
index 6956120a6a..2a84469af0 100644
--- a/plugins/ssr/ssrcommon.ml
+++ b/plugins/ssr/ssrcommon.ml
@@ -246,6 +246,7 @@ let interp_refine ist gl rc =
fail_evar = false;
expand_evars = true;
program_mode = false;
+ polymorphic = false;
}
in
let sigma, c = Pretyping.understand_ltac flags (pf_env gl) (project gl) vars kind rc in
@@ -1175,7 +1176,7 @@ let genstac (gens, clr) =
tclTHENLIST (old_cleartac clr :: List.rev_map gentac gens)
let gen_tmp_ids
- ?(ist=Geninterp.({ lfun = Id.Map.empty; extra = Tacinterp.TacStore.empty })) gl
+ ?(ist=Geninterp.({ lfun = Id.Map.empty; poly = false; extra = Tacinterp.TacStore.empty })) gl
=
let gl, ctx = pull_ctx gl in
push_ctxs ctx
@@ -1232,7 +1233,7 @@ let abs_wgen keep_let f gen (gl,args,c) =
let evar_closed t p =
if occur_existential sigma t then
CErrors.user_err ?loc:(loc_of_cpattern p) ~hdr:"ssreflect"
- (pr_constr_pat env sigma (EConstr.Unsafe.to_constr t) ++
+ (pr_econstr_pat env sigma t ++
str" contains holes and matches no subterm of the goal") in
match gen with
| _, Some ((x, mode), None) when mode = "@" || (mode = " " && keep_let) ->
diff --git a/plugins/ssr/ssrelim.ml b/plugins/ssr/ssrelim.ml
index 94f7d24242..350bb9019e 100644
--- a/plugins/ssr/ssrelim.ml
+++ b/plugins/ssr/ssrelim.ml
@@ -239,8 +239,10 @@ let ssrelim ?(is_case=false) deps what ?elim eqid elim_intro_tac =
let elimty = Reductionops.whd_all env (project gl) elimty in
seed, cty, elim, elimty, elim_args, n_elim_args, elim_is_dep, is_rec, pred, gl
in
- ppdebug(lazy Pp.(str"elim= "++ pr_constr_pat env (project gl) (EConstr.Unsafe.to_constr elim)));
- ppdebug(lazy Pp.(str"elimty= "++ pr_constr_pat env (project gl) (EConstr.Unsafe.to_constr elimty)));
+ let () =
+ let sigma = project gl in
+ ppdebug(lazy Pp.(str"elim= "++ pr_econstr_pat env sigma elim));
+ ppdebug(lazy Pp.(str"elimty= "++ pr_econstr_pat env sigma elimty)) in
let inf_deps_r = match EConstr.kind_of_type (project gl) elimty with
| AtomicType (_, args) -> List.rev (Array.to_list args)
| _ -> assert false in
@@ -304,7 +306,7 @@ let ssrelim ?(is_case=false) deps what ?elim eqid elim_intro_tac =
* looking at the ones provided by the user and the inferred ones looking at
* the type of the elimination principle *)
let pp_pat (_,p,_,occ) = Pp.(pr_occ occ ++ pp_pattern env p) in
- let pp_inf_pat gl (_,_,t,_) = pr_constr_pat env (project gl) (EConstr.Unsafe.to_constr (fire_subst gl t)) in
+ let pp_inf_pat gl (_,_,t,_) = pr_econstr_pat env (project gl) (fire_subst gl t) in
let patterns, clr, gl =
let rec loop patterns clr i = function
| [],[] -> patterns, clr, gl
@@ -318,7 +320,7 @@ let ssrelim ?(is_case=false) deps what ?elim eqid elim_intro_tac =
loop (patterns @ [i, p, inf_t, occ])
(clr_t @ clr) (i+1) (deps, inf_deps)
| [], c :: inf_deps ->
- ppdebug(lazy Pp.(str"adding inf pattern " ++ pr_constr_pat env (project gl) (EConstr.Unsafe.to_constr c)));
+ ppdebug(lazy Pp.(str"adding inf pattern " ++ pr_econstr_pat env (project gl) c));
loop (patterns @ [i, mkTpat gl c, c, allocc])
clr (i+1) ([], inf_deps)
| _::_, [] -> errorstrm Pp.(str "Too many dependent abstractions") in
@@ -341,7 +343,7 @@ let ssrelim ?(is_case=false) deps what ?elim eqid elim_intro_tac =
let elim_pred, gen_eq_tac, clr, gl =
let error gl t inf_t = errorstrm Pp.(str"The given pattern matches the term"++
spc()++pp_term gl t++spc()++str"while the inferred pattern"++
- spc()++pr_constr_pat env (project gl) (EConstr.Unsafe.to_constr (fire_subst gl inf_t))++spc()++ str"doesn't") in
+ spc()++pr_econstr_pat env (project gl) (fire_subst gl inf_t)++spc()++ str"doesn't") in
let match_or_postpone (cl, gl, post) (h, p, inf_t, occ) =
let p = unif_redex gl p inf_t in
if is_undef_pat p then
@@ -426,7 +428,7 @@ let ssrelim ?(is_case=false) deps what ?elim eqid elim_intro_tac =
if not (Evar.Set.is_empty inter) then begin
let i = Evar.Set.choose inter in
let pat = List.find (fun t -> Evar.Set.mem i (evars_of_term t)) patterns in
- errorstrm Pp.(str"Pattern"++spc()++pr_constr_pat env (project gl) (EConstr.Unsafe.to_constr pat)++spc()++
+ errorstrm Pp.(str"Pattern"++spc()++pr_econstr_pat env (project gl) pat++spc()++
str"was not completely instantiated and one of its variables"++spc()++
str"occurs in the type of another non-instantiated pattern variable");
end
diff --git a/plugins/ssr/ssrequality.ml b/plugins/ssr/ssrequality.ml
index 902098c8ce..5abbc214de 100644
--- a/plugins/ssr/ssrequality.ml
+++ b/plugins/ssr/ssrequality.ml
@@ -205,7 +205,7 @@ let rec get_evalref env sigma c = match EConstr.kind sigma c with
| App (c', _) -> get_evalref env sigma c'
| Cast (c', _, _) -> get_evalref env sigma c'
| Proj(c,_) -> EvalConstRef(Projection.constant c)
- | _ -> errorstrm Pp.(str "The term " ++ pr_constr_pat env sigma (EConstr.Unsafe.to_constr c) ++ str " is not unfoldable")
+ | _ -> errorstrm Pp.(str "The term " ++ pr_econstr_pat (Global.env ()) sigma c ++ str " is not unfoldable")
(* Strip a pattern generated by a prenex implicit to its constant. *)
let strip_unfold_term _ ((sigma, t) as p) kt = match EConstr.kind sigma t with
@@ -244,7 +244,7 @@ let unfoldintac occ rdx t (kt,_) gl =
try find_T env c h ~k:(fun env c _ _ -> EConstr.Unsafe.to_constr (body env t (EConstr.of_constr c)))
with NoMatch when easy -> c
| NoMatch | NoProgress -> errorstrm Pp.(str"No occurrence of "
- ++ pr_constr_pat env sigma (EConstr.Unsafe.to_constr t) ++ spc() ++ str "in " ++ Printer.pr_constr_env env sigma c)),
+ ++ pr_econstr_pat env sigma0 t ++ spc() ++ str "in " ++ Printer.pr_constr_env env sigma c)),
(fun () -> try end_T () with
| NoMatch when easy -> fake_pmatcher_end ()
| NoMatch -> anomaly "unfoldintac")
@@ -270,12 +270,12 @@ let unfoldintac occ rdx t (kt,_) gl =
else
try EConstr.Unsafe.to_constr @@ body env t (fs (unify_HO env sigma (EConstr.of_constr c) t) t)
with _ -> errorstrm Pp.(str "The term " ++
- pr_constr_env env sigma c ++spc()++ str "does not unify with " ++ pr_constr_pat env sigma (EConstr.Unsafe.to_constr t))),
+ pr_constr_env env sigma c ++spc()++ str "does not unify with " ++ pr_econstr_pat env sigma t)),
fake_pmatcher_end in
let concl =
let concl0 = EConstr.Unsafe.to_constr concl0 in
try beta env0 (EConstr.of_constr (eval_pattern env0 sigma0 concl0 rdx occ unfold))
- with Option.IsNone -> errorstrm Pp.(str"Failed to unfold " ++ pr_constr_pat env0 sigma (EConstr.Unsafe.to_constr t)) in
+ with Option.IsNone -> errorstrm Pp.(str"Failed to unfold " ++ pr_econstr_pat env0 sigma t) in
let _ = conclude () in
Proofview.V82.of_tactic (convert_concl concl) gl
;;
@@ -415,7 +415,7 @@ let rwcltac cl rdx dir sr gl =
let dc, r2 = EConstr.decompose_lam_n_assum (project gl) n r' in
let r3, _, r3t =
try EConstr.destCast (project gl) r2 with _ ->
- errorstrm Pp.(str "no cast from " ++ pr_constr_pat (pf_env gl) (project gl) (EConstr.Unsafe.to_constr (snd sr))
+ errorstrm Pp.(str "no cast from " ++ pr_econstr_pat (pf_env gl) (project gl) (snd sr)
++ str " to " ++ pr_econstr_env (pf_env gl) (project gl) r2) in
let cl' = EConstr.mkNamedProd (make_annot rule_id Sorts.Relevant) (EConstr.it_mkProd_or_LetIn r3t dc) (EConstr.Vars.lift 1 cl) in
let cl'' = EConstr.mkNamedProd (make_annot pattern_id Sorts.Relevant) rdxt cl' in
@@ -433,9 +433,8 @@ let rwcltac cl rdx dir sr gl =
if occur_existential (project gl) (Tacmach.pf_concl gl)
then errorstrm Pp.(str "Rewriting impacts evars" ++ error)
else errorstrm Pp.(str "Dependent type error in rewrite of "
- ++ pr_constr_env (pf_env gl) (project gl)
- (Term.mkNamedLambda (make_annot pattern_id Sorts.Relevant)
- (EConstr.Unsafe.to_constr rdxt) (EConstr.Unsafe.to_constr cl))
+ ++ pr_econstr_env (pf_env gl) (project gl)
+ (EConstr.mkNamedLambda (make_annot pattern_id Sorts.Relevant) rdxt cl)
++ error)
in
tclTHEN cvtac' rwtac gl
@@ -480,7 +479,7 @@ let rwprocess_rule dir rule gl =
let t =
if red = 1 then Tacred.hnf_constr env sigma t0
else Reductionops.whd_betaiotazeta sigma t0 in
- ppdebug(lazy Pp.(str"rewrule="++pr_constr_pat env sigma (EConstr.Unsafe.to_constr t)));
+ ppdebug(lazy Pp.(str"rewrule="++pr_econstr_pat env sigma t));
match EConstr.kind sigma t with
| Prod (_, xt, at) ->
let sigma = Evd.create_evar_defs sigma in
@@ -539,8 +538,8 @@ let rwprocess_rule dir rule gl =
sigma, (d, r', lhs, rhs) :: rs
| _ ->
if red = 0 then loop d sigma r t rs 1
- else errorstrm Pp.(str "not a rewritable relation: " ++ pr_constr_pat env sigma (EConstr.Unsafe.to_constr t)
- ++ spc() ++ str "in rule " ++ pr_constr_pat env sigma (EConstr.Unsafe.to_constr (snd rule)))
+ else errorstrm Pp.(str "not a rewritable relation: " ++ pr_econstr_pat env sigma t
+ ++ spc() ++ str "in rule " ++ pr_econstr_pat env sigma (snd rule))
in
let sigma, r = rule in
let t = Retyping.get_type_of env sigma r in
@@ -554,9 +553,9 @@ let rwrxtac occ rdx_pat dir rule gl =
let find_rule rdx =
let rec rwtac = function
| [] ->
- errorstrm Pp.(str "pattern " ++ pr_constr_pat env (project gl) (EConstr.Unsafe.to_constr rdx) ++
+ errorstrm Pp.(str "pattern " ++ pr_econstr_pat env (project gl) rdx ++
str " does not match " ++ pr_dir_side dir ++
- str " of " ++ pr_constr_pat env (project gl) (EConstr.Unsafe.to_constr (snd rule)))
+ str " of " ++ pr_econstr_pat env (project gl) (snd rule))
| (d, r, lhs, rhs) :: rs ->
try
let ise = unify_HO env (Evd.create_evar_defs r_sigma) lhs rdx in
diff --git a/plugins/ssr/ssrfwd.ml b/plugins/ssr/ssrfwd.ml
index be9586fdd7..3cadc92bcc 100644
--- a/plugins/ssr/ssrfwd.ml
+++ b/plugins/ssr/ssrfwd.ml
@@ -50,7 +50,7 @@ let ssrsettac id ((_, (pat, pty)), (_, occ)) gl =
let c = EConstr.of_constr c in
let cl = EConstr.of_constr cl in
if Termops.occur_existential sigma c then errorstrm(str"The pattern"++spc()++
- pr_constr_pat env sigma (EConstr.Unsafe.to_constr c)++spc()++str"did not match and has holes."++spc()++
+ pr_econstr_pat env sigma c++spc()++str"did not match and has holes."++spc()++
str"Did you mean pose?") else
let c, (gl, cty) = match EConstr.kind sigma c with
| Cast(t, DEFAULTcast, ty) -> t, (gl, ty)
diff --git a/plugins/ssr/ssrvernac.mlg b/plugins/ssr/ssrvernac.mlg
index d3f89147fa..0a0d9b12fa 100644
--- a/plugins/ssr/ssrvernac.mlg
+++ b/plugins/ssr/ssrvernac.mlg
@@ -566,17 +566,21 @@ let print_view_hints env sigma kind l =
}
VERNAC COMMAND EXTEND PrintView CLASSIFIED AS QUERY
-| [ "Print" "Hint" "View" ssrviewpos(i) ] ->
+| ![proof] [ "Print" "Hint" "View" ssrviewpos(i) ] ->
{
- let sigma, env = Pfedit.get_current_context () in
- match i with
+ fun ~pstate ->
+ (* XXX this is incorrect *)
+ let sigma, env = Option.cata Pfedit.get_current_context
+ (let e = Global.env () in Evd.from_env e, e) pstate in
+ (match i with
| Some k ->
print_view_hints env sigma k (Ssrview.AdaptorDb.get k)
| None ->
List.iter (fun k -> print_view_hints env sigma k (Ssrview.AdaptorDb.get k))
[ Ssrview.AdaptorDb.Forward;
Ssrview.AdaptorDb.Backward;
- Ssrview.AdaptorDb.Equivalence ]
+ Ssrview.AdaptorDb.Equivalence ]);
+ pstate
}
END
diff --git a/plugins/ssrmatching/ssrmatching.ml b/plugins/ssrmatching/ssrmatching.ml
index 5eb106cc26..1deb935d5c 100644
--- a/plugins/ssrmatching/ssrmatching.ml
+++ b/plugins/ssrmatching/ssrmatching.ml
@@ -373,6 +373,12 @@ let pr_constr_pat env sigma c0 =
if isEvar c then hole_var else map wipe_evar c in
pr_constr_env env sigma (wipe_evar c0)
+let ehole_var = EConstr.mkVar (Id.of_string "_")
+let pr_econstr_pat env sigma c0 =
+ let rec wipe_evar c = let open EConstr in
+ if isEvar sigma c then ehole_var else map sigma wipe_evar c in
+ pr_econstr_env env sigma (wipe_evar c0)
+
(* Turn (new) evars into metas *)
let evars_for_FO ~hack env sigma0 (ise0:evar_map) c0 =
let ise = ref ise0 in
@@ -694,8 +700,7 @@ let source env = match upats_origin, upats with
(if fixed_upat ise p then str"term " else str"partial term ") ++
pr_constr_pat env ise (p2t p) ++ spc()
| Some (dir,rule), [p] -> str"The " ++ pr_dir_side dir ++ str" of " ++
- pr_constr_pat env ise rule ++ fnl() ++ ws 4 ++
- pr_constr_pat env ise (p2t p) ++ fnl()
+ pr_constr_pat env ise rule ++ fnl() ++ ws 4 ++ pr_constr_pat env ise (p2t p) ++ fnl()
| Some (dir,rule), _ -> str"The " ++ pr_dir_side dir ++ str" of " ++
pr_constr_pat env ise rule ++ spc()
| _, [] | None, _::_::_ ->
@@ -732,13 +737,13 @@ let rec uniquize = function
env, 0, uniquize (instances ())
| NoMatch when (not raise_NoMatch) ->
if !failed_because_of_TC then
- errorstrm (source env++strbrk"matches but type classes inference fails")
+ errorstrm (source env ++ strbrk"matches but type classes inference fails")
else
errorstrm (source env ++ str "does not match any subterm of the goal")
| NoProgress when (not raise_NoMatch) ->
let dir = match upats_origin with Some (d,_) -> d | _ ->
CErrors.anomaly (str"mk_tpattern_matcher with no upats_origin.") in
- errorstrm (str"all matches of "++source env++
+ errorstrm (str"all matches of "++ source env ++
str"are equal to the " ++ pr_dir_side (inv_dir dir))
| NoProgress -> raise NoMatch);
let sigma, _, ({up_f = pf; up_a = pa} as u) =
@@ -823,7 +828,7 @@ let pr_pattern_aux pr_constr = function
| E_As_X_In_T (e,x,t) ->
pr_constr e ++ str " as " ++ pr_constr x ++ str " in " ++ pr_constr t
let pp_pattern env (sigma, p) =
- pr_pattern_aux (fun t -> pr_constr_pat env sigma (EConstr.Unsafe.to_constr (pi3 (nf_open_term sigma sigma (EConstr.of_constr t))))) p
+ pr_pattern_aux (fun t -> pr_econstr_pat env sigma (pi3 (nf_open_term sigma sigma (EConstr.of_constr t)))) p
let pr_cpattern = pr_term
let wit_rpatternty = add_genarg "rpatternty" (fun env sigma -> pr_pattern)
@@ -1253,10 +1258,8 @@ let fill_occ_term env cl occ sigma0 (sigma, t) =
if sigma' != sigma0 then raise NoMatch
else cl, (Evd.merge_universe_context sigma' uc, t')
with _ ->
- errorstrm (str "partial term " ++
- pr_constr_pat env sigma
- (EConstr.to_constr ~abort_on_undefined_evars:false sigma t) ++
- str " does not match any subterm of the goal")
+ errorstrm (str "partial term " ++ pr_econstr_pat env sigma t
+ ++ str " does not match any subterm of the goal")
let pf_fill_occ_term gl occ t =
let sigma0 = project gl and env = pf_env gl and concl = pf_concl gl in
@@ -1264,7 +1267,7 @@ let pf_fill_occ_term gl occ t =
cl, t
let cpattern_of_id id =
- ' ', (DAst.make @@ GRef (VarRef id, None), None), Some Geninterp.({ lfun = Id.Map.empty; extra = Tacinterp.TacStore.empty })
+ ' ', (DAst.make @@ GRef (VarRef id, None), None), Some Geninterp.({ lfun = Id.Map.empty; poly = false; extra = Tacinterp.TacStore.empty })
let is_wildcard ((_, (l, r), _) : cpattern) : bool = match DAst.get l, r with
| _, Some { CAst.v = CHole _ } | GHole _, None -> true
diff --git a/plugins/ssrmatching/ssrmatching.mli b/plugins/ssrmatching/ssrmatching.mli
index 1143bcc813..25975c84e8 100644
--- a/plugins/ssrmatching/ssrmatching.mli
+++ b/plugins/ssrmatching/ssrmatching.mli
@@ -223,6 +223,7 @@ val id_of_pattern : pattern -> Names.Id.t option
val is_wildcard : cpattern -> bool
val cpattern_of_id : Names.Id.t -> cpattern
val pr_constr_pat : env -> evar_map -> constr -> Pp.t
+val pr_econstr_pat : env -> evar_map -> econstr -> Pp.t
val pf_merge_uc : UState.t -> goal Evd.sigma -> goal Evd.sigma
val pf_unsafe_merge_uc : UState.t -> goal Evd.sigma -> goal Evd.sigma
diff --git a/plugins/syntax/g_numeral.mlg b/plugins/syntax/g_numeral.mlg
index 73a2b99434..baa4ae0306 100644
--- a/plugins/syntax/g_numeral.mlg
+++ b/plugins/syntax/g_numeral.mlg
@@ -35,8 +35,23 @@ ARGUMENT EXTEND numnotoption
END
VERNAC COMMAND EXTEND NumeralNotation CLASSIFIED AS SIDEFF
- | #[ locality = Attributes.locality; ] [ "Numeral" "Notation" reference(ty) reference(f) reference(g) ":"
+ | #[ locality = Attributes.locality; ] ![proof][ "Numeral" "Notation" reference(ty) reference(f) reference(g) ":"
ident(sc) numnotoption(o) ] ->
- { let (sigma, env) = Pfedit.get_current_context () in
- vernac_numeral_notation env sigma (Locality.make_module_locality locality) ty f g (Id.to_string sc) o }
+
+ { (* It is a bug to use the proof context here, but at the request of
+ * the reviewers we keep this broken behavior for now. The Global env
+ * should be used instead, and the `env, sigma` parameteter to the
+ * numeral notation command removed.
+ *)
+ fun ~pstate ->
+ let sigma, env = match pstate with
+ | None ->
+ let env = Global.env () in
+ let sigma = Evd.from_env env in
+ sigma, env
+ | Some pstate ->
+ Pfedit.get_current_context pstate
+ in
+ vernac_numeral_notation env sigma (Locality.make_module_locality locality) ty f g (Id.to_string sc) o;
+ pstate }
END
diff --git a/plugins/syntax/g_string.mlg b/plugins/syntax/g_string.mlg
index 171e0e213d..cc8c13a84b 100644
--- a/plugins/syntax/g_string.mlg
+++ b/plugins/syntax/g_string.mlg
@@ -19,8 +19,22 @@ open Stdarg
}
VERNAC COMMAND EXTEND StringNotation CLASSIFIED AS SIDEFF
- | #[ locality = Attributes.locality; ] [ "String" "Notation" reference(ty) reference(f) reference(g) ":"
+ | #[ locality = Attributes.locality; ] ![proof] [ "String" "Notation" reference(ty) reference(f) reference(g) ":"
ident(sc) ] ->
- { let (sigma, env) = Pfedit.get_current_context () in
- vernac_string_notation env sigma (Locality.make_module_locality locality) ty f g (Id.to_string sc) }
+ { (* It is a bug to use the proof context here, but at the request of
+ * the reviewers we keep this broken behavior for now. The Global env
+ * should be used instead, and the `env, sigma` parameteter to the
+ * numeral notation command removed.
+ *)
+ fun ~pstate ->
+ let sigma, env = match pstate with
+ | None ->
+ let env = Global.env () in
+ let sigma = Evd.from_env env in
+ sigma, env
+ | Some pstate ->
+ Pfedit.get_current_context pstate
+ in
+ vernac_string_notation env sigma (Locality.make_module_locality locality) ty f g (Id.to_string sc);
+ pstate }
END
diff --git a/pretyping/geninterp.ml b/pretyping/geninterp.ml
index 1f8b926365..32152ad0e4 100644
--- a/pretyping/geninterp.ml
+++ b/pretyping/geninterp.ml
@@ -82,9 +82,10 @@ let register_val0 wit tag =
(** Interpretation functions *)
-type interp_sign = {
- lfun : Val.t Id.Map.t;
- extra : TacStore.t }
+type interp_sign =
+ { lfun : Val.t Id.Map.t
+ ; poly : bool
+ ; extra : TacStore.t }
type ('glb, 'top) interp_fun = interp_sign -> 'glb -> 'top Ftactic.t
diff --git a/pretyping/geninterp.mli b/pretyping/geninterp.mli
index 606a6ebead..49d874289d 100644
--- a/pretyping/geninterp.mli
+++ b/pretyping/geninterp.mli
@@ -62,9 +62,10 @@ val register_val0 : ('raw, 'glb, 'top) genarg_type -> 'top Val.tag option -> uni
module TacStore : Store.S
-type interp_sign = {
- lfun : Val.t Id.Map.t;
- extra : TacStore.t }
+type interp_sign =
+ { lfun : Val.t Id.Map.t
+ ; poly : bool
+ ; extra : TacStore.t }
type ('glb, 'top) interp_fun = interp_sign -> 'glb -> 'top Ftactic.t
diff --git a/pretyping/globEnv.ml b/pretyping/globEnv.ml
index cd82b1993b..e76eb2a7de 100644
--- a/pretyping/globEnv.ml
+++ b/pretyping/globEnv.ml
@@ -183,7 +183,7 @@ let interp_ltac_id env id = ltac_interp_id env.lvar id
module ConstrInterpObj =
struct
type ('r, 'g, 't) obj =
- unbound_ltac_var_map -> env -> Evd.evar_map -> types -> 'g -> constr * Evd.evar_map
+ unbound_ltac_var_map -> bool -> env -> Evd.evar_map -> types -> 'g -> constr * Evd.evar_map
let name = "constr_interp"
let default _ = None
end
@@ -192,8 +192,8 @@ module ConstrInterp = Genarg.Register(ConstrInterpObj)
let register_constr_interp0 = ConstrInterp.register0
-let interp_glob_genarg env sigma ty arg =
+let interp_glob_genarg env poly sigma ty arg =
let open Genarg in
let GenArg (Glbwit tag, arg) = arg in
let interp = ConstrInterp.obj tag in
- interp env.lvar.ltac_genargs env.renamed_env sigma ty arg
+ interp env.lvar.ltac_genargs poly env.renamed_env sigma ty arg
diff --git a/pretyping/globEnv.mli b/pretyping/globEnv.mli
index 65ae495135..cdd36bbba6 100644
--- a/pretyping/globEnv.mli
+++ b/pretyping/globEnv.mli
@@ -19,7 +19,7 @@ open Evarutil
val register_constr_interp0 :
('r, 'g, 't) Genarg.genarg_type ->
- (unbound_ltac_var_map -> env -> evar_map -> types -> 'g -> constr * evar_map) -> unit
+ (unbound_ltac_var_map -> bool -> env -> evar_map -> types -> 'g -> constr * evar_map) -> unit
(** {6 Pretyping name management} *)
@@ -85,5 +85,5 @@ val interp_ltac_id : t -> Id.t -> Id.t
(** Interpreting a generic argument, typically a "ltac:(...)", taking
into account the possible renaming *)
-val interp_glob_genarg : t -> evar_map -> constr ->
+val interp_glob_genarg : t -> bool -> evar_map -> constr ->
Genarg.glob_generic_argument -> constr * evar_map
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index 3c8d31d671..bec939b911 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -198,6 +198,7 @@ type inference_flags = {
fail_evar : bool;
expand_evars : bool;
program_mode : bool;
+ polymorphic : bool;
}
(* Compute the set of still-undefined initial evars up to restriction
@@ -474,10 +475,10 @@ let mark_obligation_evar sigma k evc =
(* in environment [env], with existential variables [sigma] and *)
(* the type constraint tycon *)
-let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : GlobEnv.t) (sigma : evar_map) t =
+let rec pretype ~program_mode ~poly k0 resolve_tc (tycon : type_constraint) (env : GlobEnv.t) (sigma : evar_map) t =
let inh_conv_coerce_to_tycon ?loc = inh_conv_coerce_to_tycon ?loc ~program_mode resolve_tc in
- let pretype_type = pretype_type ~program_mode k0 resolve_tc in
- let pretype = pretype ~program_mode k0 resolve_tc in
+ let pretype_type = pretype_type ~program_mode ~poly k0 resolve_tc in
+ let pretype = pretype ~program_mode ~poly k0 resolve_tc in
let open Context.Rel.Declaration in
let loc = t.CAst.loc in
match DAst.get t with
@@ -497,7 +498,7 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
try Evd.evar_key id sigma
with Not_found -> error_evar_not_found ?loc !!env sigma id in
let hyps = evar_filtered_context (Evd.find sigma evk) in
- let sigma, args = pretype_instance ~program_mode k0 resolve_tc env sigma loc hyps evk inst in
+ let sigma, args = pretype_instance ~program_mode ~poly k0 resolve_tc env sigma loc hyps evk inst in
let c = mkEvar (evk, args) in
let j = Retyping.get_judgment_of !!env sigma c in
inh_conv_coerce_to_tycon ?loc env sigma j tycon
@@ -530,7 +531,7 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
match tycon with
| Some ty -> sigma, ty
| None -> new_type_evar env sigma loc in
- let c, sigma = GlobEnv.interp_glob_genarg env sigma ty arg in
+ let c, sigma = GlobEnv.interp_glob_genarg env poly sigma ty arg in
sigma, { uj_val = c; uj_type = ty }
| GRec (fixkind,names,bl,lar,vdef) ->
@@ -983,7 +984,7 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
in
inh_conv_coerce_to_tycon ?loc env sigma resj tycon
-and pretype_instance ~program_mode k0 resolve_tc env sigma loc hyps evk update =
+and pretype_instance ~program_mode ~poly k0 resolve_tc env sigma loc hyps evk update =
let f decl (subst,update,sigma) =
let id = NamedDecl.get_id decl in
let b = Option.map (replace_vars subst) (NamedDecl.get_value decl) in
@@ -1015,7 +1016,7 @@ and pretype_instance ~program_mode k0 resolve_tc env sigma loc hyps evk update =
let sigma, c, update =
try
let c = List.assoc id update in
- let sigma, c = pretype ~program_mode k0 resolve_tc (mk_tycon t) env sigma c in
+ let sigma, c = pretype ~program_mode ~poly k0 resolve_tc (mk_tycon t) env sigma c in
check_body sigma id (Some c.uj_val);
sigma, c.uj_val, List.remove_assoc id update
with Not_found ->
@@ -1040,7 +1041,7 @@ and pretype_instance ~program_mode k0 resolve_tc env sigma loc hyps evk update =
sigma, Array.map_of_list snd subst
(* [pretype_type valcon env sigma c] coerces [c] into a type *)
-and pretype_type ~program_mode k0 resolve_tc valcon (env : GlobEnv.t) sigma c = match DAst.get c with
+and pretype_type ~program_mode ~poly k0 resolve_tc valcon (env : GlobEnv.t) sigma c = match DAst.get c with
| GHole (knd, naming, None) ->
let loc = loc_of_glob_constr c in
(match valcon with
@@ -1067,7 +1068,7 @@ and pretype_type ~program_mode k0 resolve_tc valcon (env : GlobEnv.t) sigma c =
let sigma = if program_mode then mark_obligation_evar sigma knd utj_val else sigma in
sigma, { utj_val; utj_type = s})
| _ ->
- let sigma, j = pretype ~program_mode k0 resolve_tc empty_tycon env sigma c in
+ let sigma, j = pretype ~program_mode ~poly k0 resolve_tc empty_tycon env sigma c in
let loc = loc_of_glob_constr c in
let sigma, tj = Coercion.inh_coerce_to_sort ?loc !!env sigma j in
match valcon with
@@ -1082,6 +1083,7 @@ and pretype_type ~program_mode k0 resolve_tc valcon (env : GlobEnv.t) sigma c =
let ise_pretype_gen flags env sigma lvar kind c =
let program_mode = flags.program_mode in
+ let poly = flags.polymorphic in
let hypnaming =
if program_mode then ProgramNaming else KeepUserNameAndRenameExistingButSectionNames
in
@@ -1089,13 +1091,13 @@ let ise_pretype_gen flags env sigma lvar kind c =
let k0 = Context.Rel.length (rel_context !!env) in
let sigma', c', c'_ty = match kind with
| WithoutTypeConstraint ->
- let sigma, j = pretype ~program_mode k0 flags.use_typeclasses empty_tycon env sigma c in
+ let sigma, j = pretype ~program_mode ~poly k0 flags.use_typeclasses empty_tycon env sigma c in
sigma, j.uj_val, j.uj_type
| OfType exptyp ->
- let sigma, j = pretype ~program_mode k0 flags.use_typeclasses (mk_tycon exptyp) env sigma c in
+ let sigma, j = pretype ~program_mode ~poly k0 flags.use_typeclasses (mk_tycon exptyp) env sigma c in
sigma, j.uj_val, j.uj_type
| IsType ->
- let sigma, tj = pretype_type ~program_mode k0 flags.use_typeclasses empty_valcon env sigma c in
+ let sigma, tj = pretype_type ~program_mode ~poly k0 flags.use_typeclasses empty_valcon env sigma c in
sigma, tj.utj_val, mkSort tj.utj_type
in
process_inference_flags flags !!env sigma (sigma',c',c'_ty)
@@ -1106,6 +1108,7 @@ let default_inference_flags fail = {
fail_evar = fail;
expand_evars = true;
program_mode = false;
+ polymorphic = false;
}
let no_classes_no_fail_inference_flags = {
@@ -1114,6 +1117,7 @@ let no_classes_no_fail_inference_flags = {
fail_evar = false;
expand_evars = true;
program_mode = false;
+ polymorphic = false;
}
let all_and_fail_flags = default_inference_flags true
diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli
index 3c875e69d2..1037cf6cc5 100644
--- a/pretyping/pretyping.mli
+++ b/pretyping/pretyping.mli
@@ -38,6 +38,7 @@ type inference_flags = {
fail_evar : bool;
expand_evars : bool;
program_mode : bool;
+ polymorphic : bool;
}
val default_inference_flags : bool -> inference_flags
diff --git a/printing/proof_diffs.ml b/printing/proof_diffs.ml
index d620e14a94..ab4501fe75 100644
--- a/printing/proof_diffs.ml
+++ b/printing/proof_diffs.ml
@@ -547,13 +547,16 @@ let match_goals ot nt =
| None -> ());
!nevar_to_oevar
+let get_proof_context (p : Proof.t) =
+ let Proof.{goals; sigma} = Proof.data p in
+ sigma, Refiner.pf_env { Evd.it = List.(hd goals); sigma }
-let to_constr p =
+let to_constr pf =
let open CAst in
- let pprf = Proof.partial_proof p in
+ let pprf = Proof.partial_proof pf in
(* pprf generally has only one element, but it may have more in the derive plugin *)
let t = List.hd pprf in
- let sigma, env = Pfedit.get_current_context ~p () in
+ let sigma, env = get_proof_context pf in
let x = Constrextern.extern_constr false env sigma t in (* todo: right options?? *)
x.v
diff --git a/proofs/evar_refiner.ml b/proofs/evar_refiner.ml
index 0f97a942ed..1a34105ab6 100644
--- a/proofs/evar_refiner.ml
+++ b/proofs/evar_refiner.ml
@@ -55,6 +55,7 @@ let w_refine (evk,evi) (ltac_var,rawc) sigma =
Pretyping.fail_evar = false;
Pretyping.expand_evars = true;
Pretyping.program_mode = false;
+ Pretyping.polymorphic = false;
} in
try Pretyping.understand_ltac flags
env sigma ltac_var (Pretyping.OfType evi.evar_concl) rawc
diff --git a/proofs/pfedit.ml b/proofs/pfedit.ml
index 9509c36ec0..472db790f2 100644
--- a/proofs/pfedit.ml
+++ b/proofs/pfedit.ml
@@ -37,41 +37,35 @@ let get_nth_V82_goal p i =
try { it = List.nth goals (i-1) ; sigma }
with Failure _ -> raise NoSuchGoal
-let get_goal_context_gen p i =
- let { it=goal ; sigma=sigma; } = get_nth_V82_goal p i in
+let get_goal_context_gen pf i =
+ let { it=goal ; sigma=sigma; } = get_nth_V82_goal pf i in
(sigma, Refiner.pf_env { it=goal ; sigma=sigma; })
-let get_goal_context i =
- try get_goal_context_gen (Proof_global.give_me_the_proof ()) i
- with Proof_global.NoCurrentProof -> CErrors.user_err Pp.(str "No focused proof.")
- | NoSuchGoal -> CErrors.user_err Pp.(str "No such goal.")
-
-let get_current_goal_context () =
- try get_goal_context_gen (Proof_global.give_me_the_proof ()) 1
- with Proof_global.NoCurrentProof -> CErrors.user_err Pp.(str "No focused proof.")
- | NoSuchGoal ->
- (* spiwack: returning empty evar_map, since if there is no goal, under focus,
- there is no accessible evar either *)
- let env = Global.env () in
- (Evd.from_env env, env)
+let get_goal_context pf i =
+ let p = Proof_global.give_me_the_proof pf in
+ get_goal_context_gen p i
-let get_current_context ?p () =
- let current_proof_by_default = function
- | Some p -> p
- | None -> Proof_global.give_me_the_proof ()
- in
- try get_goal_context_gen (current_proof_by_default p) 1
- with Proof_global.NoCurrentProof ->
+let get_current_goal_context pf =
+ let p = Proof_global.give_me_the_proof pf in
+ try get_goal_context_gen p 1
+ with
+ | NoSuchGoal ->
+ (* spiwack: returning empty evar_map, since if there is no goal,
+ under focus, there is no accessible evar either. EJGA: this
+ seems strange, as we have pf *)
let env = Global.env () in
- (Evd.from_env env, env)
- | NoSuchGoal ->
- (* No more focused goals ? *)
- let p = (current_proof_by_default p) in
- let evd = Proof.in_proof p (fun x -> x) in
- (evd, Global.env ())
+ Evd.from_env env, env
+
+let get_current_context pf =
+ let p = Proof_global.give_me_the_proof pf in
+ try get_goal_context_gen p 1
+ with
+ | NoSuchGoal ->
+ (* No more focused goals *)
+ let evd = Proof.in_proof p (fun x -> x) in
+ evd, Global.env ()
let solve ?with_end_tac gi info_lvl tac pr =
- try
let tac = match with_end_tac with
| None -> tac
| Some etac -> Proofview.tclTHEN tac etac in
@@ -112,15 +106,12 @@ let solve ?with_end_tac gi info_lvl tac pr =
| Some i -> Feedback.msg_info (hov 0 (Proofview.Trace.pr_info env sigma ~lvl:i info))
in
(p,status)
- with
- Proof_global.NoCurrentProof -> CErrors.user_err Pp.(str "No focused proof")
let by tac = Proof_global.with_current_proof (fun _ -> solve (Goal_select.SelectNth 1) None tac)
-let instantiate_nth_evar_com n com =
+let instantiate_nth_evar_com n com =
Proof_global.simple_with_current_proof (fun _ p ->
- Proof.V82.instantiate_evar Global.(env ())n com p)
-
+ Proof.V82.instantiate_evar Global.(env ()) n com p)
(**********************************************************************)
(* Shortcut to build a term using tactics *)
@@ -133,21 +124,19 @@ let build_constant_by_tactic id ctx sign ?(goal_kind = Global, false, Proof Theo
let evd = Evd.from_ctx ctx in
let terminator = Proof_global.make_terminator (fun _ -> ()) in
let goals = [ (Global.env_of_context sign , typ) ] in
- Proof_global.start_proof evd id goal_kind goals terminator;
+ let pf = Proof_global.start_proof ~ontop:None evd id goal_kind goals terminator in
try
- let status = by tac in
+ let pf, status = by tac pf in
let open Proof_global in
- let { entries; universes } = fst @@ close_proof ~opaque:Transparent ~keep_body_ucst_separate:false (fun x -> x) in
+ let { entries; universes } = fst @@ close_proof ~opaque:Transparent ~keep_body_ucst_separate:false (fun x -> x) pf in
match entries with
| [entry] ->
- discard_current ();
let univs = UState.demote_seff_univs entry universes in
entry, status, univs
| _ ->
CErrors.anomaly Pp.(str "[build_constant_by_tactic] close_proof returned more than one proof term")
with reraise ->
let reraise = CErrors.push reraise in
- Proof_global.discard_current ();
iraise reraise
let build_by_tactic ?(side_eff=true) env sigma ?(poly=false) typ tac =
diff --git a/proofs/pfedit.mli b/proofs/pfedit.mli
index 29ab00876a..2fe4bc6385 100644
--- a/proofs/pfedit.mli
+++ b/proofs/pfedit.mli
@@ -16,29 +16,29 @@ open Environ
open Decl_kinds
(** {6 ... } *)
+
+exception NoSuchGoal
+
(** [get_goal_context n] returns the context of the [n]th subgoal of
the current focused proof or raises a [UserError] if there is no
focused proof or if there is no more subgoals *)
-val get_goal_context : int -> Evd.evar_map * env
+val get_goal_context : Proof_global.t -> int -> Evd.evar_map * env
(** [get_current_goal_context ()] works as [get_goal_context 1] *)
-
-val get_current_goal_context : unit -> Evd.evar_map * env
+val get_current_goal_context : Proof_global.t -> Evd.evar_map * env
(** [get_current_context ()] returns the context of the
current focused goal. If there is no focused goal but there
is a proof in progress, it returns the corresponding evar_map.
If there is no pending proof then it returns the current global
environment and empty evar_map. *)
-
-val get_current_context : ?p:Proof.t -> unit -> Evd.evar_map * env
+val get_current_context : Proof_global.t -> Evd.evar_map * env
(** {6 ... } *)
(** [solve (SelectNth n) tac] applies tactic [tac] to the [n]th
- subgoal of the current focused proof or raises a [UserError] if no
- proof is focused or if there is no [n]th subgoal. [solve SelectAll
+ subgoal of the current focused proof. [solve SelectAll
tac] applies [tac] to all subgoals. *)
val solve : ?with_end_tac:unit Proofview.tactic ->
@@ -46,11 +46,10 @@ val solve : ?with_end_tac:unit Proofview.tactic ->
Proof.t -> Proof.t * bool
(** [by tac] applies tactic [tac] to the 1st subgoal of the current
- focused proof or raises a UserError if there is no focused proof or
- if there is no more subgoals.
+ focused proof.
Returns [false] if an unsafe tactic has been used. *)
-val by : unit Proofview.tactic -> bool
+val by : unit Proofview.tactic -> Proof_global.t -> Proof_global.t * bool
(** Option telling if unification heuristics should be used. *)
val use_unification_heuristics : unit -> bool
@@ -60,7 +59,7 @@ val use_unification_heuristics : unit -> bool
UserError if no proof is focused or if there is no such [n]th
existential variable *)
-val instantiate_nth_evar_com : int -> Constrexpr.constr_expr -> unit
+val instantiate_nth_evar_com : int -> Constrexpr.constr_expr -> Proof_global.t -> Proof_global.t
(** [build_by_tactic typ tac] returns a term of type [typ] by calling
[tac]. The return boolean, if [false] indicates the use of an unsafe
diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml
index 6174b75a96..86d3d9601e 100644
--- a/proofs/proof_global.ml
+++ b/proofs/proof_global.ml
@@ -17,7 +17,6 @@
(***********************************************************************)
open Util
-open Pp
open Names
open Context
@@ -55,108 +54,66 @@ type pstate = {
strength : Decl_kinds.goal_kind;
}
-type t = pstate list
+(* The head of [t] is the actual current proof, the other ones are
+ to be resumed when the current proof is closed or aborted. *)
+type t = pstate * pstate list
+
+let pstate_map f (pf, pfl) = (f pf, List.map f pfl)
let make_terminator f = f
let apply_terminator f = f
-(* The head of [!pstates] is the actual current proof, the other ones are
- to be resumed when the current proof is closed or aborted. *)
-let pstates = ref ([] : pstate list)
-
(* combinators for the current_proof lists *)
-let push a l = l := a::!l
-
-exception NoCurrentProof
-let () = CErrors.register_handler begin function
- | NoCurrentProof -> CErrors.user_err Pp.(str "No focused proof (No proof-editing in progress).")
- | _ -> raise CErrors.Unhandled
-end
+let push ~ontop a =
+ match ontop with
+ | None -> a , []
+ | Some (l,ls) -> a, (l :: ls)
(*** Proof Global manipulation ***)
-let get_all_proof_names () =
- List.map Proof.(function pf -> (data pf.proof).name) !pstates
-
-let cur_pstate () =
- match !pstates with
- | np::_ -> np
- | [] -> raise NoCurrentProof
-
-let give_me_the_proof () = (cur_pstate ()).proof
-let give_me_the_proof_opt () = try Some (give_me_the_proof ()) with | NoCurrentProof -> None
-let get_current_proof_name () = (Proof.data (cur_pstate ()).proof).Proof.name
-let get_current_persistence () = (cur_pstate ()).strength
-
-let with_current_proof f =
- match !pstates with
- | [] -> raise NoCurrentProof
- | p :: rest ->
- let et =
- match p.endline_tactic with
- | None -> Proofview.tclUNIT ()
- | Some tac ->
- let open Geninterp in
- let ist = { lfun = Id.Map.empty; extra = TacStore.empty } in
- let Genarg.GenArg (Genarg.Glbwit tag, tac) = tac in
- let tac = Geninterp.interp tag ist tac in
- Ftactic.run tac (fun _ -> Proofview.tclUNIT ())
- in
- let (newpr,ret) = f et p.proof in
- let p = { p with proof = newpr } in
- pstates := p :: rest;
- ret
-
-let simple_with_current_proof f = with_current_proof (fun t p -> f t p , ())
-
-let compact_the_proof () = simple_with_current_proof (fun _ -> Proof.compact)
+let get_all_proof_names (pf : t) =
+ let (pn, pns) = pstate_map Proof.(function pf -> (data pf.proof).name) pf in
+ pn :: pns
+
+let give_me_the_proof (ps,_) = ps.proof
+let get_current_proof_name (ps,_) = (Proof.data ps.proof).Proof.name
+let get_current_persistence (ps,_) = ps.strength
+
+let with_current_proof f (ps, psl) =
+ let et =
+ match ps.endline_tactic with
+ | None -> Proofview.tclUNIT ()
+ | Some tac ->
+ let open Geninterp in
+ let ist = { lfun = Id.Map.empty; poly = pi2 ps.strength; extra = TacStore.empty } in
+ let Genarg.GenArg (Genarg.Glbwit tag, tac) = tac in
+ let tac = Geninterp.interp tag ist tac in
+ Ftactic.run tac (fun _ -> Proofview.tclUNIT ())
+ in
+ let (newpr,ret) = f et ps.proof in
+ let ps = { ps with proof = newpr } in
+ (ps, psl), ret
+
+let simple_with_current_proof f pf =
+ let p, () = with_current_proof (fun t p -> f t p , ()) pf in p
+
+let compact_the_proof pf = simple_with_current_proof (fun _ -> Proof.compact) pf
(* Sets the tactic to be used when a tactic line is closed with [...] *)
-let set_endline_tactic tac =
- match !pstates with
- | [] -> raise NoCurrentProof
- | p :: rest -> pstates := { p with endline_tactic = Some tac } :: rest
-
-(* spiwack: it might be considered to move error messages away.
- Or else to remove special exceptions from Proof_global.
- Arguments for the former: there is no reason Proof_global is only
- accessed directly through vernacular commands. Error message should be
- pushed to external layers, and so we should be able to have a finer
- control on error message on complex actions. *)
-let msg_proofs () =
- match get_all_proof_names () with
- | [] -> (spc () ++ str"(No proof-editing in progress).")
- | l -> (str"." ++ fnl () ++ str"Proofs currently edited:" ++ spc () ++
- (pr_sequence Id.print l) ++ str".")
-
-let there_is_a_proof () = not (List.is_empty !pstates)
-let there_are_pending_proofs () = there_is_a_proof ()
-let check_no_pending_proof () =
- if not (there_are_pending_proofs ()) then
- ()
- else begin
- CErrors.user_err
- (str"Proof editing in progress" ++ msg_proofs () ++ fnl() ++
- str"Use \"Abort All\" first or complete proof(s).")
- end
+let set_endline_tactic tac (ps, psl) =
+ { ps with endline_tactic = Some tac }, psl
let pf_name_eq id ps =
let Proof.{ name } = Proof.data ps.proof in
Id.equal name id
-let discard_gen id =
- pstates := List.filter (fun pf -> not (pf_name_eq id pf)) !pstates
-
-let discard {CAst.loc;v=id} =
- let n = List.length !pstates in
- discard_gen id;
- if Int.equal (List.length !pstates) n then
- CErrors.user_err ?loc
- ~hdr:"Pfedit.delete_proof" (str"No such proof" ++ msg_proofs ())
+let discard {CAst.loc;v=id} (ps, psl) =
+ match List.filter (fun pf -> not (pf_name_eq id pf)) (ps :: psl) with
+ | [] -> None
+ | ps :: psl -> Some (ps, psl)
-let discard_current () =
- if List.is_empty !pstates then raise NoCurrentProof else pstates := List.tl !pstates
-let discard_all () = pstates := []
+let discard_current (ps, psl) =
+ if List.is_empty psl then None else Some List.(hd psl, tl psl)
(** [start_proof sigma id pl str goals terminator] starts a proof of name
[id] with goals [goals] (a list of pairs of environment and
@@ -166,30 +123,30 @@ let discard_all () = pstates := []
end of the proof to close the proof. The proof is started in the
evar map [sigma] (which can typically contain universe
constraints), and with universe bindings pl. *)
-let start_proof sigma name ?(pl=UState.default_univ_decl) kind goals terminator =
+let start_proof ~ontop sigma name ?(pl=UState.default_univ_decl) kind goals terminator =
let initial_state = {
terminator = CEphemeron.create terminator;
proof = Proof.start ~name ~poly:(pi2 kind) sigma goals;
endline_tactic = None;
section_vars = None;
- strength = kind;
- universe_decl = pl } in
- push initial_state pstates
+ universe_decl = pl;
+ strength = kind } in
+ push ~ontop initial_state
-let start_dependent_proof name ?(pl=UState.default_univ_decl) kind goals terminator =
+let start_dependent_proof ~ontop name ?(pl=UState.default_univ_decl) kind goals terminator =
let initial_state = {
terminator = CEphemeron.create terminator;
proof = Proof.dependent_start ~name ~poly:(pi2 kind) goals;
endline_tactic = None;
section_vars = None;
- strength = kind;
- universe_decl = pl } in
- push initial_state pstates
+ universe_decl = pl;
+ strength = kind } in
+ push ~ontop initial_state
-let get_used_variables () = (cur_pstate ()).section_vars
-let get_universe_decl () = (cur_pstate ()).universe_decl
+let get_used_variables (pf,_) = pf.section_vars
+let get_universe_decl (pf,_) = pf.universe_decl
-let set_used_variables l =
+let set_used_variables (ps,psl) l =
let open Context.Named.Declaration in
let env = Global.env () in
let ids = List.fold_right Id.Set.add l Id.Set.empty in
@@ -210,20 +167,17 @@ let set_used_variables l =
else (ctx, all_safe) in
let ctx, _ =
Environ.fold_named_context aux env ~init:(ctx,ctx_set) in
- match !pstates with
- | [] -> raise NoCurrentProof
- | p :: rest ->
- if not (Option.is_empty p.section_vars) then
- CErrors.user_err Pp.(str "Used section variables can be declared only once");
- pstates := { p with section_vars = Some ctx} :: rest;
- ctx, []
-
-let get_open_goals () =
- let Proof.{ goals; stack; shelf } = Proof.data (cur_pstate ()).proof in
+ if not (Option.is_empty ps.section_vars) then
+ CErrors.user_err Pp.(str "Used section variables can be declared only once");
+ (* EJGA: This is always empty thus we should modify the type *)
+ (ctx, []), ({ ps with section_vars = Some ctx}, psl)
+
+let get_open_goals (ps, _) =
+ let Proof.{ goals; stack; shelf } = Proof.data ps.proof in
List.length goals +
- List.fold_left (+) 0
+ List.fold_left (+) 0
(List.map (fun (l1,l2) -> List.length l1 + List.length l2) stack) +
- List.length shelf
+ List.length shelf
type closed_proof_output = (Constr.t * Safe_typing.private_constants) list * UState.t
@@ -240,8 +194,8 @@ let private_poly_univs =
fun () -> !b
let close_proof ~opaque ~keep_body_ucst_separate ?feedback_id ~now
- (fpl : closed_proof_output Future.computation) =
- let { section_vars; proof; terminator; universe_decl; strength } = cur_pstate () in
+ (fpl : closed_proof_output Future.computation) ps =
+ let { section_vars; proof; terminator; universe_decl; strength } = ps in
let Proof.{ name; poly; entry; initial_euctx } = Proof.data proof in
let opaque = match opaque with Opaque -> true | Transparent -> false in
let constrain_variables ctx =
@@ -339,8 +293,8 @@ let close_proof ~opaque ~keep_body_ucst_separate ?feedback_id ~now
universes },
fun pr_ending -> CEphemeron.get terminator pr_ending
-let return_proof ?(allow_partial=false) () =
- let { proof } = cur_pstate () in
+let return_proof ?(allow_partial=false) (ps,_) =
+ let { proof } = ps in
if allow_partial then begin
let proofs = Proof.partial_proof proof in
let Proof.{sigma=evd} = Proof.data proof in
@@ -368,43 +322,44 @@ let return_proof ?(allow_partial=false) () =
List.map (fun (c, _) -> (proof_opt c, eff)) initial_goals in
proofs, Evd.evar_universe_context evd
-let close_future_proof ~opaque ~feedback_id proof =
- close_proof ~opaque ~keep_body_ucst_separate:true ~feedback_id ~now:false proof
-let close_proof ~opaque ~keep_body_ucst_separate fix_exn =
+let close_future_proof ~opaque ~feedback_id (ps, psl) proof =
+ close_proof ~opaque ~keep_body_ucst_separate:true ~feedback_id ~now:false proof ps
+
+let close_proof ~opaque ~keep_body_ucst_separate fix_exn (ps, psl) =
close_proof ~opaque ~keep_body_ucst_separate ~now:true
- (Future.from_val ~fix_exn (return_proof ()))
+ (Future.from_val ~fix_exn (return_proof (ps,psl))) ps
(** Gets the current terminator without checking that the proof has
been completed. Useful for the likes of [Admitted]. *)
-let get_terminator () = CEphemeron.get ( cur_pstate() ).terminator
-let set_terminator hook =
- match !pstates with
- | [] -> raise NoCurrentProof
- | p :: ps -> pstates := { p with terminator = CEphemeron.create hook } :: ps
-
-let freeze ~marshallable =
- if marshallable then CErrors.anomaly (Pp.str"full marshalling of proof state not supported.")
- else !pstates
-let unfreeze s = pstates := s
-let proof_of_state = function { proof }::_ -> proof | _ -> raise NoCurrentProof
+let get_terminator (ps, _) = CEphemeron.get ps.terminator
+let set_terminator hook (ps, psl) =
+ { ps with terminator = CEphemeron.create hook }, psl
+
let copy_terminators ~src ~tgt =
- assert(List.length src = List.length tgt);
- List.map2 (fun op p -> { p with terminator = op.terminator }) src tgt
+ let (ps, psl), (ts,tsl) = src, tgt in
+ assert(List.length psl = List.length tsl);
+ {ts with terminator = ps.terminator}, List.map2 (fun op p -> { p with terminator = op.terminator }) psl tsl
-let update_global_env pf_info =
+let update_global_env (pf : t) =
+ let res, () =
with_current_proof (fun _ p ->
Proof.in_proof p (fun sigma ->
let tac = Proofview.Unsafe.tclEVARS (Evd.update_sigma_env sigma (Global.env ())) in
let (p,(status,info)) = Proof.run_tactic (Global.env ()) tac p in
- (p, ())))
-
-(* XXX: Bullet hook, should be really moved elsewhere *)
-let () =
- let hook n =
- try
- let prf = give_me_the_proof () in
- (Proof_bullet.suggest prf)
- with NoCurrentProof -> mt ()
- in
- Proofview.set_nosuchgoals_hook hook
+ (p, ()))) pf
+ in res
+
+(* XXX: This hook is used to provide a better error w.r.t. bullets,
+ however the proof engine [surprise!] knows nothing about bullets so
+ here we have a layering violation. The right fix is to modify the
+ entry point to handle this and reraise the exception with the
+ needed information. *)
+(* let _ =
+ * let hook n =
+ * try
+ * let prf = give_me_the_proof pf in
+ * (Proof_bullet.suggest prf)
+ * with NoCurrentProof -> mt ()
+ * in
+ * Proofview.set_nosuchgoals_hook hook *)
diff --git a/proofs/proof_global.mli b/proofs/proof_global.mli
index 38e234eaee..e2e457483b 100644
--- a/proofs/proof_global.mli
+++ b/proofs/proof_global.mli
@@ -13,23 +13,15 @@
environment. *)
type t
-val there_are_pending_proofs : unit -> bool
-val check_no_pending_proof : unit -> unit
+val get_current_proof_name : t -> Names.Id.t
+val get_current_persistence : t -> Decl_kinds.goal_kind
+val get_all_proof_names : t -> Names.Id.t list
-val get_current_proof_name : unit -> Names.Id.t
-val get_current_persistence : unit -> Decl_kinds.goal_kind
-val get_all_proof_names : unit -> Names.Id.t list
+val discard : Names.lident -> t -> t option
+val discard_current : t -> t option
-val discard : Names.lident -> unit
-val discard_current : unit -> unit
-val discard_all : unit -> unit
-
-val give_me_the_proof_opt : unit -> Proof.t option
-exception NoCurrentProof
-val give_me_the_proof : unit -> Proof.t
-(** @raise NoCurrentProof when outside proof mode. *)
-
-val compact_the_proof : unit -> unit
+val give_me_the_proof : t -> Proof.t
+val compact_the_proof : t -> t
(** When a proof is closed, it is reified into a [proof_object], where
[id] is the name of the proof, [entries] the list of the proof terms
@@ -60,7 +52,7 @@ type closed_proof = proof_object * proof_terminator
val make_terminator : (proof_ending -> unit) -> proof_terminator
val apply_terminator : proof_terminator -> proof_ending -> unit
-(** [start_proof id str pl goals terminator] starts a proof of name
+(** [start_proof ~ontop id str pl goals terminator] starts a proof of name
[id] with goals [goals] (a list of pairs of environment and
conclusion); [str] describes what kind of theorem/definition this
is; [terminator] is used at the end of the proof to close the proof
@@ -68,25 +60,25 @@ val apply_terminator : proof_terminator -> proof_ending -> unit
morphism). The proof is started in the evar map [sigma] (which can
typically contain universe constraints), and with universe bindings
pl. *)
-val start_proof :
+val start_proof : ontop:t option ->
Evd.evar_map -> Names.Id.t -> ?pl:UState.universe_decl ->
Decl_kinds.goal_kind -> (Environ.env * EConstr.types) list ->
- proof_terminator -> unit
+ proof_terminator -> t
(** Like [start_proof] except that there may be dependencies between
initial goals. *)
-val start_dependent_proof :
+val start_dependent_proof : ontop:t option ->
Names.Id.t -> ?pl:UState.universe_decl -> Decl_kinds.goal_kind ->
- Proofview.telescope -> proof_terminator -> unit
+ Proofview.telescope -> proof_terminator -> t
(** Update the proofs global environment after a side-effecting command
(e.g. a sublemma definition) has been run inside it. Assumes
there_are_pending_proofs. *)
-val update_global_env : unit -> unit
+val update_global_env : t -> t
(* Takes a function to add to the exceptions data relative to the
state in which the proof was built *)
-val close_proof : opaque:opacity_flag -> keep_body_ucst_separate:bool -> Future.fix_exn -> closed_proof
+val close_proof : opaque:opacity_flag -> keep_body_ucst_separate:bool -> Future.fix_exn -> t -> closed_proof
(* Intermediate step necessary to delegate the future.
* Both access the current proof state. The former is supposed to be
@@ -96,39 +88,36 @@ type closed_proof_output = (Constr.t * Safe_typing.private_constants) list * USt
(* If allow_partial is set (default no) then an incomplete proof
* is allowed (no error), and a warn is given if the proof is complete. *)
-val return_proof : ?allow_partial:bool -> unit -> closed_proof_output
-val close_future_proof : opaque:opacity_flag -> feedback_id:Stateid.t ->
+val return_proof : ?allow_partial:bool -> t -> closed_proof_output
+val close_future_proof : opaque:opacity_flag -> feedback_id:Stateid.t -> t ->
closed_proof_output Future.computation -> closed_proof
(** Gets the current terminator without checking that the proof has
been completed. Useful for the likes of [Admitted]. *)
-val get_terminator : unit -> proof_terminator
-val set_terminator : proof_terminator -> unit
-
-val get_open_goals : unit -> int
+val get_terminator : t -> proof_terminator
+val set_terminator : proof_terminator -> t -> t
+val get_open_goals : t -> int
(** Runs a tactic on the current proof. Raises [NoCurrentProof] is there is
no current proof.
The return boolean is set to [false] if an unsafe tactic has been used. *)
val with_current_proof :
- (unit Proofview.tactic -> Proof.t -> Proof.t * 'a) -> 'a
+ (unit Proofview.tactic -> Proof.t -> Proof.t * 'a) -> t -> t * 'a
val simple_with_current_proof :
- (unit Proofview.tactic -> Proof.t -> Proof.t) -> unit
+ (unit Proofview.tactic -> Proof.t -> Proof.t) -> t -> t
(** Sets the tactic to be used when a tactic line is closed with [...] *)
-val set_endline_tactic : Genarg.glob_generic_argument -> unit
+val set_endline_tactic : Genarg.glob_generic_argument -> t -> t
(** Sets the section variables assumed by the proof, returns its closure
* (w.r.t. type dependencies and let-ins covered by it) + a list of
* ids to be cleared *)
-val set_used_variables :
- Names.Id.t list -> Constr.named_context * Names.lident list
-val get_used_variables : unit -> Constr.named_context option
+val set_used_variables : t ->
+ Names.Id.t list -> (Constr.named_context * Names.lident list) * t
+
+val get_used_variables : t -> Constr.named_context option
(** Get the universe declaration associated to the current proof. *)
-val get_universe_decl : unit -> UState.universe_decl
+val get_universe_decl : t -> UState.universe_decl
-val freeze : marshallable:bool -> t
-val unfreeze : t -> unit
-val proof_of_state : t -> Proof.t
val copy_terminators : src:t -> tgt:t -> t
diff --git a/stm/proofBlockDelimiter.ml b/stm/proofBlockDelimiter.ml
index 230a3207a8..d13763cdec 100644
--- a/stm/proofBlockDelimiter.ml
+++ b/stm/proofBlockDelimiter.ml
@@ -49,12 +49,13 @@ let is_focused_goal_simple ~doc id =
match state_of_id ~doc id with
| `Expired | `Error _ | `Valid None -> `Not
| `Valid (Some { Vernacstate.proof }) ->
- let proof = Proof_global.proof_of_state proof in
- let Proof.{ goals=focused; stack=r1; shelf=r2; given_up=r3; sigma } = Proof.data proof in
- let rest = List.(flatten (map (fun (x,y) -> x @ y) r1)) @ r2 @ r3 in
- if List.for_all (fun x -> simple_goal sigma x rest) focused
- then `Simple focused
- else `Not
+ Option.cata (fun proof ->
+ let proof = Proof_global.give_me_the_proof proof in
+ let Proof.{ goals=focused; stack=r1; shelf=r2; given_up=r3; sigma } = Proof.data proof in
+ let rest = List.(flatten (map (fun (x,y) -> x @ y) r1)) @ r2 @ r3 in
+ if List.for_all (fun x -> simple_goal sigma x rest) focused
+ then `Simple focused
+ else `Not) `Not proof
type 'a until = [ `Stop | `Found of static_block_declaration | `Cont of 'a ]
diff --git a/stm/stm.ml b/stm/stm.ml
index 0c5d0c7b5d..cc0de0e9df 100644
--- a/stm/stm.ml
+++ b/stm/stm.ml
@@ -139,8 +139,8 @@ let may_pierce_opaque = function
| _ -> false
let update_global_env () =
- if Proof_global.there_are_pending_proofs () then
- Proof_global.update_global_env ()
+ if Vernacstate.Proof_global.there_are_pending_proofs () then
+ Vernacstate.Proof_global.update_global_env ()
module Vcs_ = Vcs.Make(Stateid.Self)
type future_proof = Proof_global.closed_proof_output Future.computation
@@ -872,7 +872,7 @@ end = struct (* {{{ *)
let invalidate_cur_state () = cur_id := Stateid.dummy
type proof_part =
- Proof_global.t *
+ Proof_global.t option *
int * (* Evarutil.meta_counter_summary_tag *)
int * (* Evd.evar_counter_summary_tag *)
Obligations.program_info Names.Id.Map.t (* Obligations.program_tcc_summary_tag *)
@@ -948,8 +948,8 @@ end = struct (* {{{ *)
let prev = (VCS.visit id).next in
if is_cached_and_valid prev
then { s with proof =
- Proof_global.copy_terminators
- ~src:(get_cached prev).proof ~tgt:s.proof }
+ Vernacstate.Proof_global.copy_terminators
+ ~src:((get_cached prev).proof) ~tgt:s.proof }
else s
with VCS.Expired -> s in
VCS.set_state id (FullState s)
@@ -957,7 +957,7 @@ end = struct (* {{{ *)
if is_cached_and_valid ontop then
let s = get_cached ontop in
let s = { s with proof =
- Proof_global.copy_terminators ~src:s.proof ~tgt:pstate } in
+ Vernacstate.Proof_global.copy_terminators ~src:s.proof ~tgt:pstate } in
let s = { s with system =
States.replace_summary s.system
begin
@@ -1009,8 +1009,8 @@ end = struct (* {{{ *)
if feedback_processed then
Hooks.(call state_computed ~doc id ~in_cache:false);
VCS.reached id;
- if Proof_global.there_are_pending_proofs () then
- VCS.goals id (Proof_global.get_open_goals ())
+ if Vernacstate.Proof_global.there_are_pending_proofs () then
+ VCS.goals id (Vernacstate.Proof_global.get_open_goals ())
with e ->
let (e, info) = CErrors.push e in
let good_id = !cur_id in
@@ -1130,9 +1130,9 @@ let show_script ?proof () =
try
let prf =
try match proof with
- | None -> Some (Proof_global.get_current_proof_name ())
+ | None -> Some (Vernacstate.Proof_global.get_current_proof_name ())
| Some (p,_) -> Some (p.Proof_global.id)
- with Proof_global.NoCurrentProof -> None
+ with Vernacstate.Proof_global.NoCurrentProof -> None
in
let cmds = get_script prf in
let _,_,_,indented_cmds =
@@ -1255,9 +1255,8 @@ end = struct (* {{{ *)
if Int.equal n 0 then `Stop id else `Cont (n-value)
let get_proof ~doc id =
- let open Vernacstate in
match state_of_id ~doc id with
- | `Valid (Some vstate) -> Some (Proof_global.proof_of_state vstate.proof)
+ | `Valid (Some vstate) -> Option.map Proof_global.give_me_the_proof vstate.Vernacstate.proof
| _ -> None
let undo_vernac_classifier v ~doc =
@@ -1296,7 +1295,7 @@ end = struct (* {{{ *)
| Some vcs, _ -> vcs in
let cb, _ =
try Vcs_aux.find_proof_at_depth vcs (Vcs_aux.proof_nesting vcs)
- with Failure _ -> raise Proof_global.NoCurrentProof in
+ with Failure _ -> raise Vernacstate.Proof_global.NoCurrentProof in
let n = fold_until (fun n (_,vcs,_,_,_) ->
if List.mem cb (Vcs_.branches vcs) then `Cont (n+1) else `Stop n)
0 id in
@@ -1334,7 +1333,7 @@ end = struct (* {{{ *)
| None -> true
done;
!rv
- with Not_found | Proof_global.NoCurrentProof -> None
+ with Not_found | Vernacstate.Proof_global.NoCurrentProof -> None
end (* }}} *)
@@ -1595,7 +1594,7 @@ end = struct (* {{{ *)
let wall_clock2 = Unix.gettimeofday () in
Aux_file.record_in_aux_at ?loc "proof_build_time"
(Printf.sprintf "%.3f" (wall_clock2 -. wall_clock1));
- let p = Proof_global.return_proof ~allow_partial:drop_pt () in
+ let p = Vernacstate.Proof_global.return_proof ~allow_partial:drop_pt () in
if drop_pt then feedback ~id Complete;
p)
@@ -1622,7 +1621,7 @@ end = struct (* {{{ *)
to set the state manually here *)
State.unfreeze st;
let pobject, _ =
- Proof_global.close_future_proof ~opaque ~feedback_id:stop (Future.from_val ~fix_exn p) in
+ Vernacstate.Proof_global.close_future_proof ~opaque ~feedback_id:stop (Future.from_val ~fix_exn p) in
let terminator = (* The one sent by master is an InvalidKey *)
Lemmas.(standard_proof_terminator []) in
@@ -1759,15 +1758,15 @@ end = struct (* {{{ *)
try
Reach.known_state ~doc:dummy_doc (* XXX should be document *) ~cache:false stop;
if drop then
- let _proof = Proof_global.return_proof ~allow_partial:true () in
+ let _proof = Vernacstate.Proof_global.return_proof ~allow_partial:true () in
`OK_ADMITTED
else begin
(* The original terminator, a hook, has not been saved in the .vio*)
- Proof_global.set_terminator (Lemmas.standard_proof_terminator []);
+ Vernacstate.Proof_global.set_terminator (Lemmas.standard_proof_terminator []);
let opaque = Proof_global.Opaque in
let proof =
- Proof_global.close_proof ~opaque ~keep_body_ucst_separate:true (fun x -> x) in
+ Vernacstate.Proof_global.close_proof ~opaque ~keep_body_ucst_separate:true (fun x -> x) in
(* We jump at the beginning since the kernel handles side effects by also
* looking at the ones that happen to be present in the current env *)
Reach.known_state ~doc:dummy_doc (* XXX should be document *) ~cache:false start;
@@ -2017,7 +2016,7 @@ end = struct (* {{{ *)
try
Reach.known_state ~doc:dummy_doc (* XXX should be vcs *) ~cache:false id;
State.purify (fun () ->
- let Proof.{sigma=sigma0} = Proof.data (Proof_global.give_me_the_proof ()) in
+ let Proof.{sigma=sigma0} = Proof.data (Vernacstate.Proof_global.give_me_the_proof ()) in
let g = Evd.find sigma0 r_goal in
let is_ground c = Evarutil.is_ground_term sigma0 c in
if not (
@@ -2029,7 +2028,7 @@ end = struct (* {{{ *)
"goals only"))
else begin
let (i, ast) = r_ast in
- Proof_global.simple_with_current_proof (fun _ p -> Proof.focus focus_cond () i p);
+ Vernacstate.Proof_global.simple_with_current_proof (fun _ p -> Proof.focus focus_cond () i p);
(* STATE SPEC:
* - start : id
* - return: id
@@ -2038,7 +2037,7 @@ end = struct (* {{{ *)
*)
let st = Vernacstate.freeze_interp_state ~marshallable:false in
ignore(stm_vernac_interp r_state_fb st ast);
- let Proof.{sigma} = Proof.data (Proof_global.give_me_the_proof ()) in
+ let Proof.{sigma} = Proof.data (Vernacstate.Proof_global.give_me_the_proof ()) in
match Evd.(evar_body (find sigma r_goal)) with
| Evd.Evar_empty -> RespNoProgress
| Evd.Evar_defined t ->
@@ -2065,8 +2064,14 @@ end = struct (* {{{ *)
module TaskQueue = AsyncTaskQueue.MakeQueue(TacTask) ()
+ let stm_fail ~st fail f =
+ if fail then
+ Vernacentries.with_fail ~st f
+ else
+ f ()
+
let vernac_interp ~solve ~abstract ~cancel_switch nworkers safe_id id
- { indentation; verbose; loc; expr = e; strlen }
+ { indentation; verbose; loc; expr = e; strlen } : unit
=
let e, time, batch, fail =
let rec find ~time ~batch ~fail = function
@@ -2076,10 +2081,10 @@ end = struct (* {{{ *)
| e -> e, time, batch, fail in
find ~time:false ~batch:false ~fail:false e in
let st = Vernacstate.freeze_interp_state ~marshallable:false in
- Vernacentries.with_fail st fail (fun () ->
+ stm_fail ~st fail (fun () ->
(if time then System.with_time ~batch ~header:(Pp.mt ()) else (fun x -> x)) (fun () ->
ignore(TaskQueue.with_n_workers nworkers (fun queue ->
- Proof_global.with_current_proof (fun _ p ->
+ Vernacstate.Proof_global.with_current_proof (fun _ p ->
let Proof.{goals} = Proof.data p in
let open TacTask in
let res = CList.map_i (fun i g ->
@@ -2112,7 +2117,7 @@ end = struct (* {{{ *)
let open Notations in
match Future.join f with
| Some (pt, uc) ->
- let sigma, env = Pfedit.get_current_context () in
+ let sigma, env = Vernacstate.Proof_global.get_current_context () in
stm_pperr_endline (fun () -> hov 0 (
str"g=" ++ int (Evar.repr gid) ++ spc () ++
str"t=" ++ (Printer.pr_constr_env env sigma pt) ++ spc () ++
@@ -2392,10 +2397,10 @@ let known_state ~doc ?(redefine_qed=false) ~cache id =
(Proofview.Goal.goal gl) goals_to_admit then
Proofview.give_up else Proofview.tclUNIT ()
end in
- match VCS.get_state base_state with
+ match (VCS.get_info base_state).state with
| FullState { Vernacstate.proof } ->
- Proof_global.unfreeze proof;
- Proof_global.with_current_proof (fun _ p ->
+ Option.iter Vernacstate.Proof_global.unfreeze proof;
+ Vernacstate.Proof_global.with_current_proof (fun _ p ->
feedback ~id:id Feedback.AddedAxiom;
fst (Pfedit.solve Goal_select.SelectAll None tac p), ());
(* STATE SPEC:
@@ -2565,7 +2570,7 @@ let known_state ~doc ?(redefine_qed=false) ~cache id =
| VtKeepDefined -> Proof_global.Transparent
in
let proof =
- Proof_global.close_future_proof ~opaque ~feedback_id:id fp in
+ Vernacstate.Proof_global.close_future_proof ~opaque ~feedback_id:id fp in
if not delegate then ignore(Future.compute fp);
reach view.next;
let st = Vernacstate.freeze_interp_state ~marshallable:false in
@@ -2573,13 +2578,13 @@ let known_state ~doc ?(redefine_qed=false) ~cache id =
feedback ~id:id Incomplete
| { VCS.kind = `Master }, _ -> assert false
end;
- Proof_global.discard_all ()
+ Vernacstate.Proof_global.discard_all ()
), not redefine_qed, true
| `Sync (name, `Immediate) -> (fun () ->
reach eop;
let st = Vernacstate.freeze_interp_state ~marshallable:false in
ignore(stm_vernac_interp id st x);
- Proof_global.discard_all ()
+ Vernacstate.Proof_global.discard_all ()
), true, true
| `Sync (name, reason) -> (fun () ->
log_processing_sync id name reason;
@@ -2598,7 +2603,7 @@ let known_state ~doc ?(redefine_qed=false) ~cache id =
| VtKeepOpaque -> Opaque | VtKeepDefined -> Transparent
| VtKeepAxiom -> assert false
in
- Some(Proof_global.close_proof ~opaque
+ Some(Vernacstate.Proof_global.close_proof ~opaque
~keep_body_ucst_separate:false
(State.exn_on id ~valid:eop)) in
if keep <> VtKeep VtKeepAxiom then
@@ -2609,7 +2614,7 @@ let known_state ~doc ?(redefine_qed=false) ~cache id =
let wall_clock3 = Unix.gettimeofday () in
Aux_file.record_in_aux_at ?loc:x.loc "proof_check_time"
(Printf.sprintf "%.3f" (wall_clock3 -. wall_clock2));
- Proof_global.discard_all ()
+ Vernacstate.Proof_global.discard_all ()
), true, true
| `MaybeASync (start, nodes, name, delegate) -> (fun () ->
reach ~cache:true start;
@@ -2870,7 +2875,7 @@ let merge_proof_branch ~valid ?id qast keep brname =
VCS.checkout VCS.Branch.master;
`Unfocus qed_id
| { VCS.kind = `Master } ->
- Exninfo.iraise (State.exn_on ~valid Stateid.dummy (Proof_global.NoCurrentProof, Exninfo.null))
+ Exninfo.iraise (State.exn_on ~valid Stateid.dummy (Vernacstate.Proof_global.NoCurrentProof, Exninfo.null))
(* When tty is true, this code also does some of the job of the user interface:
jump back to a state that is valid *)
@@ -3062,7 +3067,7 @@ let process_transaction ~doc ?(newtip=Stateid.fresh ())
let st = Vernacstate.freeze_interp_state ~marshallable:false in
ignore(stm_vernac_interp id st x);
(* Vernac x may or may not start a proof *)
- if not in_proof && Proof_global.there_are_pending_proofs () then
+ if not in_proof && Vernacstate.Proof_global.there_are_pending_proofs () then
begin
let bname = VCS.mk_branch_name x in
let opacity_of_produced_term = function
diff --git a/tactics/auto.ml b/tactics/auto.ml
index 2619620eb8..4e0ec1f7e4 100644
--- a/tactics/auto.ml
+++ b/tactics/auto.ml
@@ -143,7 +143,8 @@ let conclPattern concl pat tac =
Proofview.Goal.enter begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
- constr_bindings env sigma >>= fun constr_bindings ->
+ constr_bindings env sigma >>= fun constr_bindings ->
+ Proofview.tclProofInfo [@ocaml.warning "-3"] >>= fun (_name, poly) ->
let open Genarg in
let open Geninterp in
let inj c = match val_tag (topwit Stdarg.wit_constr) with
@@ -152,7 +153,9 @@ let conclPattern concl pat tac =
in
let fold id c accu = Id.Map.add id (inj c) accu in
let lfun = Id.Map.fold fold constr_bindings Id.Map.empty in
- let ist = { lfun; extra = TacStore.empty } in
+ let ist = { lfun
+ ; poly
+ ; extra = TacStore.empty } in
match tac with
| GenArg (Glbwit wit, tac) ->
Ftactic.run (Geninterp.interp wit ist tac) (fun _ -> Proofview.tclUNIT ())
diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml
index d9c0a26f91..51708670f5 100644
--- a/tactics/autorewrite.ml
+++ b/tactics/autorewrite.ml
@@ -99,11 +99,15 @@ let one_base general_rewrite_maybe_in tac_main bas =
Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma)
(general_rewrite_maybe_in dir c' tc)
end in
- let lrul = List.map (fun h ->
+ let open Proofview.Notations in
+ Proofview.tclProofInfo [@ocaml.warning "-3"] >>= fun (_name, poly) ->
+ let lrul = List.map (fun h ->
let tac = match h.rew_tac with
| None -> Proofview.tclUNIT ()
| Some (Genarg.GenArg (Genarg.Glbwit wit, tac)) ->
- let ist = { Geninterp.lfun = Id.Map.empty; extra = Geninterp.TacStore.empty } in
+ let ist = { Geninterp.lfun = Id.Map.empty
+ ; poly
+ ; extra = Geninterp.TacStore.empty } in
Ftactic.run (Geninterp.interp wit ist tac) (fun _ -> Proofview.tclUNIT ())
in
(h.rew_ctx,h.rew_lemma,h.rew_l2r,tac)) lrul in
diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml
index a3620f4081..44102afd74 100644
--- a/tactics/class_tactics.ml
+++ b/tactics/class_tactics.ml
@@ -933,11 +933,12 @@ module Search = struct
try
(* Instance may try to call this before a proof is set up!
Thus, give_me_the_proof will fail. Beware! *)
- let name, poly = try
- let Proof.{ name; poly } = Proof.data Proof_global.(give_me_the_proof ()) in
- name, poly
- with | Proof_global.NoCurrentProof ->
- Id.of_string "instance", false
+ let name, poly =
+ (* try
+ * let Proof.{ name; poly } = Proof.data Proof_global.(give_me_the_proof ()) in
+ * name, poly
+ * with | Proof_global.NoCurrentProof -> *)
+ Id.of_string "instance", false
in
let (), pv', (unsafe, shelved, gaveup), _ =
Proofview.apply ~name ~poly env tac pv
diff --git a/tactics/hints.ml b/tactics/hints.ml
index 85d75f1010..3a7e67cb3f 100644
--- a/tactics/hints.ml
+++ b/tactics/hints.ml
@@ -1514,9 +1514,9 @@ let pr_hint_term env sigma cl =
(str "No hint applicable for current goal")
(* print all hints that apply to the concl of the current goal *)
-let pr_applicable_hint () =
+let pr_applicable_hint pf =
let env = Global.env () in
- let pts = Proof_global.give_me_the_proof () in
+ let pts = Proof_global.give_me_the_proof pf in
let Proof.{goals;sigma} = Proof.data pts in
match goals with
| [] -> CErrors.user_err Pp.(str "No focused goal.")
diff --git a/tactics/hints.mli b/tactics/hints.mli
index dd2c63d351..e84e423faa 100644
--- a/tactics/hints.mli
+++ b/tactics/hints.mli
@@ -294,7 +294,7 @@ val wrap_hint_warning_fun : env -> evar_map ->
(** Printing hints *)
val pr_searchtable : env -> evar_map -> Pp.t
-val pr_applicable_hint : unit -> Pp.t
+val pr_applicable_hint : Proof_global.t -> Pp.t
val pr_hint_ref : env -> evar_map -> GlobRef.t -> Pp.t
val pr_hint_db_by_name : env -> evar_map -> hint_db_name -> Pp.t
val pr_hint_db_env : env -> evar_map -> Hint_db.t -> Pp.t
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index b8308dc49b..206f35c8ba 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -1161,6 +1161,7 @@ let tactic_infer_flags with_evar = {
Pretyping.fail_evar = not with_evar;
Pretyping.expand_evars = true;
Pretyping.program_mode = false;
+ Pretyping.polymorphic = false;
}
type evars_flag = bool (* true = pose evars false = fail on evars *)
diff --git a/toplevel/ccompile.ml b/toplevel/ccompile.ml
index 3fe6ad0718..416ea88c1b 100644
--- a/toplevel/ccompile.ml
+++ b/toplevel/ccompile.ml
@@ -85,7 +85,7 @@ let ensure_exists f =
let compile opts copts ~echo ~f_in ~f_out =
let open Vernac.State in
let check_pending_proofs () =
- let pfs = Proof_global.get_all_proof_names () in
+ let pfs = Vernacstate.Proof_global.get_all_proof_names () in
if not (CList.is_empty pfs) then
fatal_error (str "There are pending proofs: "
++ (pfs
diff --git a/toplevel/coqc.ml b/toplevel/coqc.ml
index d4107177a7..fd4c515209 100644
--- a/toplevel/coqc.ml
+++ b/toplevel/coqc.ml
@@ -46,8 +46,9 @@ let coqc_main () =
outputstate copts;
flush_all();
+
if opts.Coqargs.output_context then begin
- let sigma, env = Pfedit.get_current_context () in
+ let sigma, env = let e = Global.env () in Evd.from_env e, e in
Feedback.msg_notice Pp.(Flags.(with_option raw_print (Prettyp.print_full_pure_context env) sigma) ++ fnl ())
end;
CProfile.print_profile ()
diff --git a/toplevel/coqloop.ml b/toplevel/coqloop.ml
index 1094fc86b4..b3de8dd85f 100644
--- a/toplevel/coqloop.ml
+++ b/toplevel/coqloop.ml
@@ -191,8 +191,8 @@ end
from cycling. *)
let make_prompt () =
try
- (Names.Id.to_string (Proof_global.get_current_proof_name ())) ^ " < "
- with Proof_global.NoCurrentProof ->
+ (Names.Id.to_string (Vernacstate.Proof_global.get_current_proof_name ())) ^ " < "
+ with Vernacstate.Proof_global.NoCurrentProof ->
"Coq < "
(* the coq prompt added to the default one when in emacs mode
@@ -353,7 +353,7 @@ let print_anyway c =
let top_goal_print ~doc c oldp newp =
try
let proof_changed = not (Option.equal cproof oldp newp) in
- let print_goals = proof_changed && Proof_global.there_are_pending_proofs () ||
+ let print_goals = proof_changed && Vernacstate.Proof_global.there_are_pending_proofs () ||
print_anyway c in
if not !Flags.quiet && print_goals then begin
let dproof = Stm.get_prev_proof ~doc (Stm.get_current_state ~doc) in
diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml
index ef1dc6993b..ce584f1109 100644
--- a/toplevel/vernac.ml
+++ b/toplevel/vernac.ml
@@ -70,7 +70,7 @@ let interp_vernac ~check ~interactive ~state ({CAst.loc;_} as com) =
(* Force the command *)
let ndoc = if check then Stm.observe ~doc nsid else doc in
- let new_proof = Proof_global.give_me_the_proof_opt () in
+ let new_proof = Vernacstate.Proof_global.give_me_the_proof_opt () in
{ state with doc = ndoc; sid = nsid; proof = new_proof; }
with reraise ->
(* XXX: In non-interactive mode edit_at seems to do very weird
diff --git a/vernac/classes.ml b/vernac/classes.ml
index 61b8cc3dcb..6a67a1b5d0 100644
--- a/vernac/classes.ml
+++ b/vernac/classes.ml
@@ -144,7 +144,7 @@ let do_declare_instance env sigma ~global ~poly k u ctx ctx' pri decl imps subst
Declare.declare_univ_binders (ConstRef cst) (Evd.universe_binders sigma);
instance_hook k pri global imps (ConstRef cst)
-let declare_instance_open env sigma ?hook ~tac ~program_mode ~global ~poly k id pri imps decl ids term termtype =
+let declare_instance_open ~pstate env sigma ?hook ~tac ~program_mode ~global ~poly k id pri imps decl ids term termtype =
let kind = Decl_kinds.Global, poly, Decl_kinds.DefinitionBody Decl_kinds.Instance in
if program_mode then
let hook _ _ vis gr =
@@ -163,33 +163,44 @@ let declare_instance_open env sigma ?hook ~tac ~program_mode ~global ~poly k id
in
let hook = Lemmas.mk_hook hook in
let ctx = Evd.evar_universe_context sigma in
- ignore (Obligations.add_definition id ?term:constr
- ~univdecl:decl typ ctx ~kind:(Global,poly,Instance) ~hook obls)
+ let _progress = Obligations.add_definition id ?term:constr
+ ~univdecl:decl typ ctx ~kind:(Global,poly,Instance) ~hook obls in
+ pstate
else
- Flags.silently (fun () ->
+ Some Flags.(silently (fun () ->
(* spiwack: it is hard to reorder the actions to do
the pretyping after the proof has opened. As a
consequence, we use the low-level primitives to code
the refinement manually.*)
let gls = List.rev (Evd.future_goals sigma) in
let sigma = Evd.reset_future_goals sigma in
- Lemmas.start_proof id ~pl:decl kind sigma (EConstr.of_constr termtype)
+ let pstate = Lemmas.start_proof ~ontop:pstate id ~pl:decl kind sigma (EConstr.of_constr termtype)
~hook:(Lemmas.mk_hook
- (fun _ _ _ -> instance_hook k pri global imps ?hook));
+ (fun _ _ _ -> instance_hook k pri global imps ?hook)) in
(* spiwack: I don't know what to do with the status here. *)
- if not (Option.is_empty term) then
- let init_refine =
- Tacticals.New.tclTHENLIST [
- Refine.refine ~typecheck:false (fun sigma -> (sigma,EConstr.of_constr (Option.get term)));
- Proofview.Unsafe.tclNEWGOALS (CList.map Proofview.with_empty_state gls);
- Tactics.New.reduce_after_refine;
- ]
- in
- ignore (Pfedit.by init_refine)
- else ignore (Pfedit.by (Tactics.auto_intros_tac ids));
- (match tac with Some tac -> ignore (Pfedit.by tac) | None -> ())) ()
+ let pstate =
+ if not (Option.is_empty term) then
+ let init_refine =
+ Tacticals.New.tclTHENLIST [
+ Refine.refine ~typecheck:false (fun sigma -> (sigma,EConstr.of_constr (Option.get term)));
+ Proofview.Unsafe.tclNEWGOALS (CList.map Proofview.with_empty_state gls);
+ Tactics.New.reduce_after_refine;
+ ]
+ in
+ let pstate, _ = Pfedit.by init_refine pstate in
+ pstate
+ else
+ let pstate, _ = Pfedit.by (Tactics.auto_intros_tac ids) pstate in
+ pstate
+ in
+ match tac with
+ | Some tac ->
+ let pstate, _ = Pfedit.by tac pstate in
+ pstate
+ | None ->
+ pstate) ())
-let do_instance env env' sigma ?hook ~refine ~tac ~global ~poly ~program_mode cty k u ctx ctx' pri decl imps subst id props =
+let do_instance ~pstate env env' sigma ?hook ~refine ~tac ~global ~poly ~program_mode cty k u ctx ctx' pri decl imps subst id props =
let props =
match props with
| Some (true, { CAst.v = CRecord fs }) ->
@@ -269,12 +280,14 @@ let do_instance env env' sigma ?hook ~refine ~tac ~global ~poly ~program_mode ct
Pretyping.check_evars env (Evd.from_env env) sigma termtype;
let termtype = to_constr sigma termtype in
let term = Option.map (to_constr ~abort_on_undefined_evars:false sigma) term in
- if not (Evd.has_undefined sigma) && not (Option.is_empty props) then
- declare_instance_constant k pri global imps ?hook id decl poly sigma (Option.get term) termtype
- else if program_mode || refine || Option.is_empty props then
- declare_instance_open env sigma ?hook ~tac ~program_mode ~global ~poly k id pri imps decl (List.map RelDecl.get_name ctx) term termtype
- else CErrors.user_err Pp.(str "Unsolved obligations remaining.");
- id
+ let pstate =
+ if not (Evd.has_undefined sigma) && not (Option.is_empty props) then
+ (declare_instance_constant k pri global imps ?hook id decl poly sigma (Option.get term) termtype;
+ None)
+ else if program_mode || refine || Option.is_empty props then
+ declare_instance_open ~pstate env sigma ?hook ~tac ~program_mode ~global ~poly k id pri imps decl (List.map RelDecl.get_name ctx) term termtype
+ else CErrors.user_err Pp.(str "Unsolved obligations remaining.") in
+ id, pstate
let interp_instance_context ~program_mode env ctx ?(generalize=false) pl bk cl =
let sigma, decl = Constrexpr_ops.interp_univ_decl_opt env pl in
@@ -318,7 +331,7 @@ let interp_instance_context ~program_mode env ctx ?(generalize=false) pl bk cl =
sigma, cl, u, c', ctx', ctx, imps, args, decl
-let new_instance ?(global=false) ?(refine= !refine_instance) ~program_mode
+let new_instance ~pstate ?(global=false) ?(refine= !refine_instance) ~program_mode
poly ctx (instid, bk, cl) props
?(generalize=true) ?(tac:unit Proofview.tactic option) ?hook pri =
let env = Global.env() in
@@ -334,7 +347,7 @@ let new_instance ?(global=false) ?(refine= !refine_instance) ~program_mode
Namegen.next_global_ident_away i (Termops.vars_of_env env)
in
let env' = push_rel_context ctx env in
- do_instance env env' sigma ?hook ~refine ~tac ~global ~poly ~program_mode
+ do_instance ~pstate env env' sigma ?hook ~refine ~tac ~global ~poly ~program_mode
cty k u ctx ctx' pri decl imps subst id props
let declare_new_instance ?(global=false) ~program_mode poly ctx (instid, bk, cl) pri =
@@ -358,7 +371,7 @@ let named_of_rel_context l =
l ([], [])
in ctx
-let context poly l =
+let context ~pstate poly l =
let env = Global.env() in
let sigma = Evd.from_env env in
let sigma, (_, ((env', fullctx), impls)) = interp_context_evars ~program_mode:false env sigma l in
@@ -426,12 +439,12 @@ let context poly l =
let decl = (Discharge, poly, Definitional) in
let nstatus = match b with
| None ->
- pi3 (ComAssumption.declare_assumption false decl (t, univs) UnivNames.empty_binders [] impl
+ pi3 (ComAssumption.declare_assumption ~pstate false decl (t, univs) UnivNames.empty_binders [] impl
Declaremods.NoInline (CAst.make id))
| Some b ->
let decl = (Discharge, poly, Definition) in
let entry = Declare.definition_entry ~univs ~types:t b in
- let _ = DeclareDef.declare_definition id decl entry UnivNames.empty_binders [] in
+ let _gr = DeclareDef.declare_definition ~ontop:pstate id decl entry UnivNames.empty_binders [] in
Lib.sections_are_opened () || Lib.is_modtype_strict ()
in
status && nstatus
diff --git a/vernac/classes.mli b/vernac/classes.mli
index 7e0ec42625..73e4b024ef 100644
--- a/vernac/classes.mli
+++ b/vernac/classes.mli
@@ -40,6 +40,7 @@ val declare_instance_constant :
unit
val new_instance :
+ pstate:Proof_global.t option ->
?global:bool (** Not global by default. *) ->
?refine:bool (** Allow refinement *) ->
program_mode:bool ->
@@ -51,7 +52,8 @@ val new_instance :
?tac:unit Proofview.tactic ->
?hook:(GlobRef.t -> unit) ->
Hints.hint_info_expr ->
- Id.t
+ (* May open a proof *)
+ Id.t * Proof_global.t option
val declare_new_instance :
?global:bool (** Not global by default. *) ->
@@ -74,4 +76,8 @@ val id_of_class : typeclass -> Id.t
(** returns [false] if, for lack of section, it declares an assumption
(unless in a module type). *)
-val context : Decl_kinds.polymorphic -> local_binder_expr list -> bool
+val context
+ : pstate:Proof_global.t option
+ -> Decl_kinds.polymorphic
+ -> local_binder_expr list
+ -> bool
diff --git a/vernac/comAssumption.ml b/vernac/comAssumption.ml
index 37a33daf8f..d7bd64067b 100644
--- a/vernac/comAssumption.ml
+++ b/vernac/comAssumption.ml
@@ -42,7 +42,7 @@ let should_axiom_into_instance = function
true
| Global | Local -> !axiom_into_instance
-let declare_assumption is_coe (local,p,kind) (c,ctx) pl imps impl nl {CAst.v=ident} =
+let declare_assumption ~pstate is_coe (local,p,kind) (c,ctx) pl imps impl nl {CAst.v=ident} =
match local with
| Discharge when Lib.sections_are_opened () ->
let ctx = match ctx with
@@ -53,7 +53,7 @@ match local with
let _ = declare_variable ident decl in
let () = assumption_message ident in
let () =
- if not !Flags.quiet && Proof_global.there_are_pending_proofs () then
+ if not !Flags.quiet && Option.has_some pstate then
Feedback.msg_info Pp.(str"Variable" ++ spc () ++ Id.print ident ++
strbrk " is not visible from current goals")
in
@@ -96,11 +96,11 @@ let next_uctx =
| Polymorphic_entry _ as uctx -> uctx
| Monomorphic_entry _ -> empty_uctx
-let declare_assumptions idl is_coe k (c,uctx) pl imps nl =
+let declare_assumptions ~pstate idl is_coe k (c,uctx) pl imps nl =
let refs, status, _ =
List.fold_left (fun (refs,status,uctx) id ->
let ref',u',status' =
- declare_assumption is_coe k (c,uctx) pl imps false nl id in
+ declare_assumption ~pstate is_coe k (c,uctx) pl imps false nl id in
(ref',u')::refs, status' && status, next_uctx uctx)
([],true,uctx) idl
in
@@ -132,7 +132,7 @@ let process_assumptions_udecls kind l =
in
udecl, List.map (fun (coe, (idl, c)) -> coe, (List.map fst idl, c)) l
-let do_assumptions ~program_mode kind nl l =
+let do_assumptions ~pstate ~program_mode kind nl l =
let open Context.Named.Declaration in
let env = Global.env () in
let udecl, l = process_assumptions_udecls kind l in
@@ -173,7 +173,7 @@ let do_assumptions ~program_mode kind nl l =
let ubinders = Evd.universe_binders sigma in
pi2 (List.fold_left (fun (subst,status,uctx) ((is_coe,idl),t,imps) ->
let t = replace_vars subst t in
- let refs, status' = declare_assumptions idl is_coe kind (t,uctx) ubinders imps nl in
+ let refs, status' = declare_assumptions ~pstate idl is_coe kind (t,uctx) ubinders imps nl in
let subst' = List.map2
(fun {CAst.v=id} (c,u) -> (id, Constr.mkRef (c,u)))
idl refs
diff --git a/vernac/comAssumption.mli b/vernac/comAssumption.mli
index 2b794b001a..32914cc11b 100644
--- a/vernac/comAssumption.mli
+++ b/vernac/comAssumption.mli
@@ -17,8 +17,13 @@ open Decl_kinds
(** {6 Parameters/Assumptions} *)
-val do_assumptions : program_mode:bool -> locality * polymorphic * assumption_object_kind ->
- Declaremods.inline -> (ident_decl list * constr_expr) with_coercion list -> bool
+val do_assumptions
+ : pstate:Proof_global.t option
+ -> program_mode:bool
+ -> locality * polymorphic * assumption_object_kind
+ -> Declaremods.inline
+ -> (ident_decl list * constr_expr) with_coercion list
+ -> bool
(************************************************************************)
(** Internal API *)
@@ -28,10 +33,16 @@ val do_assumptions : program_mode:bool -> locality * polymorphic * assumption_ob
(** returns [false] if the assumption is neither local to a section,
nor in a module type and meant to be instantiated. *)
-val declare_assumption : coercion_flag -> assumption_kind ->
- types in_universes_entry ->
- UnivNames.universe_binders -> Impargs.manual_implicits ->
- bool (** implicit *) -> Declaremods.inline -> variable CAst.t ->
- GlobRef.t * Univ.Instance.t * bool
+val declare_assumption
+ : pstate:Proof_global.t option
+ -> coercion_flag
+ -> assumption_kind
+ -> types in_universes_entry
+ -> UnivNames.universe_binders
+ -> Impargs.manual_implicits
+ -> bool (** implicit *)
+ -> Declaremods.inline
+ -> variable CAst.t
+ -> GlobRef.t * Univ.Instance.t * bool
val do_primitive : lident -> CPrimitives.op_or_type -> constr_expr option -> unit
diff --git a/vernac/comDefinition.ml b/vernac/comDefinition.ml
index 28773a3965..feaf47df18 100644
--- a/vernac/comDefinition.ml
+++ b/vernac/comDefinition.ml
@@ -90,7 +90,7 @@ let check_definition ~program_mode (ce, evd, _, imps) =
check_evars_are_solved ~program_mode env evd;
ce
-let do_definition ~program_mode ?hook ident k univdecl bl red_option c ctypopt =
+let do_definition ~ontop ~program_mode ?hook ident k univdecl bl red_option c ctypopt =
let (ce, evd, univdecl, imps as def) =
interp_definition ~program_mode univdecl bl (pi2 k) red_option c ctypopt
in
@@ -114,4 +114,4 @@ let do_definition ~program_mode ?hook ident k univdecl bl red_option c ctypopt =
let ce = check_definition ~program_mode def in
let uctx = Evd.evar_universe_context evd in
let hook_data = Option.map (fun hook -> hook, uctx, []) hook in
- ignore(DeclareDef.declare_definition ident k ?hook_data ce (Evd.universe_binders evd) imps )
+ ignore(DeclareDef.declare_definition ~ontop ident k ?hook_data ce (Evd.universe_binders evd) imps)
diff --git a/vernac/comDefinition.mli b/vernac/comDefinition.mli
index 9cb6190fcc..12853d83e0 100644
--- a/vernac/comDefinition.mli
+++ b/vernac/comDefinition.mli
@@ -16,11 +16,18 @@ open Constrexpr
(** {6 Definitions/Let} *)
-val do_definition : program_mode:bool ->
- ?hook:Lemmas.declaration_hook ->
- Id.t -> definition_kind -> universe_decl_expr option ->
- local_binder_expr list -> red_expr option -> constr_expr ->
- constr_expr option -> unit
+val do_definition
+ : ontop:Proof_global.t option
+ -> program_mode:bool
+ -> ?hook:Lemmas.declaration_hook
+ -> Id.t
+ -> definition_kind
+ -> universe_decl_expr option
+ -> local_binder_expr list
+ -> red_expr option
+ -> constr_expr
+ -> constr_expr option
+ -> unit
(************************************************************************)
(** Internal API *)
diff --git a/vernac/comFixpoint.ml b/vernac/comFixpoint.ml
index 2f00b41b7c..2aadbd224f 100644
--- a/vernac/comFixpoint.ml
+++ b/vernac/comFixpoint.ml
@@ -255,7 +255,8 @@ let interp_fixpoint ~cofix l ntns =
let uctx,fix = ground_fixpoint env evd fix in
(fix,pl,uctx,info)
-let declare_fixpoint local poly ((fixnames,fixrs,fixdefs,fixtypes),pl,ctx,fiximps) indexes ntns =
+let declare_fixpoint ~ontop local poly ((fixnames,fixrs,fixdefs,fixtypes),pl,ctx,fiximps) indexes ntns =
+ let pstate =
if List.exists Option.is_empty fixdefs then
(* Some bodies to define by proof *)
let thms =
@@ -265,8 +266,9 @@ let declare_fixpoint local poly ((fixnames,fixrs,fixdefs,fixtypes),pl,ctx,fiximp
Some (List.map (Option.cata (EConstr.of_constr %> Tactics.exact_no_check) Tacticals.New.tclIDTAC)
fixdefs) in
let evd = Evd.from_ctx ctx in
- Lemmas.start_proof_with_initialization (local,poly,DefinitionBody Fixpoint)
- evd pl (Some(false,indexes,init_tac)) thms None
+ Some
+ (Lemmas.start_proof_with_initialization ~ontop (local,poly,DefinitionBody Fixpoint)
+ evd pl (Some(false,indexes,init_tac)) thms None)
else begin
(* We shortcut the proof process *)
let fixdefs = List.map Option.get fixdefs in
@@ -282,15 +284,18 @@ let declare_fixpoint local poly ((fixnames,fixrs,fixdefs,fixtypes),pl,ctx,fiximp
let ctx = Evd.check_univ_decl ~poly evd pl in
let pl = Evd.universe_binders evd in
let fixdecls = List.map Safe_typing.mk_pure_proof fixdecls in
- ignore (List.map4 (DeclareDef.declare_fix (local, poly, Fixpoint) pl ctx)
+ ignore (List.map4 (DeclareDef.declare_fix ~ontop (local, poly, Fixpoint) pl ctx)
fixnames fixdecls fixtypes fiximps);
(* Declare the recursive definitions *)
fixpoint_message (Some indexes) fixnames;
- end;
+ None
+ end in
(* Declare notations *)
- List.iter (Metasyntax.add_notation_interpretation (Global.env())) ntns
+ List.iter (Metasyntax.add_notation_interpretation (Global.env())) ntns;
+ pstate
-let declare_cofixpoint local poly ((fixnames,fixrs,fixdefs,fixtypes),pl,ctx,fiximps) ntns =
+let declare_cofixpoint ~ontop local poly ((fixnames,fixrs,fixdefs,fixtypes),pl,ctx,fiximps) ntns =
+ let pstate =
if List.exists Option.is_empty fixdefs then
(* Some bodies to define by proof *)
let thms =
@@ -300,8 +305,8 @@ let declare_cofixpoint local poly ((fixnames,fixrs,fixdefs,fixtypes),pl,ctx,fixi
Some (List.map (Option.cata (EConstr.of_constr %> Tactics.exact_no_check) Tacticals.New.tclIDTAC)
fixdefs) in
let evd = Evd.from_ctx ctx in
- Lemmas.start_proof_with_initialization (Global,poly, DefinitionBody CoFixpoint)
- evd pl (Some(true,[],init_tac)) thms None
+ Some (Lemmas.start_proof_with_initialization ~ontop (Global,poly, DefinitionBody CoFixpoint)
+ evd pl (Some(true,[],init_tac)) thms None)
else begin
(* We shortcut the proof process *)
let fixdefs = List.map Option.get fixdefs in
@@ -314,13 +319,15 @@ let declare_cofixpoint local poly ((fixnames,fixrs,fixdefs,fixtypes),pl,ctx,fixi
let evd = Evd.restrict_universe_context evd vars in
let ctx = Evd.check_univ_decl ~poly evd pl in
let pl = Evd.universe_binders evd in
- ignore (List.map4 (DeclareDef.declare_fix (local, poly, CoFixpoint) pl ctx)
+ ignore (List.map4 (DeclareDef.declare_fix ~ontop (local, poly, CoFixpoint) pl ctx)
fixnames fixdecls fixtypes fiximps);
(* Declare the recursive definitions *)
- cofixpoint_message fixnames
- end;
+ cofixpoint_message fixnames;
+ None
+ end in
(* Declare notations *)
- List.iter (Metasyntax.add_notation_interpretation (Global.env())) ntns
+ List.iter (Metasyntax.add_notation_interpretation (Global.env())) ntns;
+ pstate
let extract_decreasing_argument limit = function
| (na,CStructRec) -> na
@@ -348,16 +355,18 @@ let check_safe () =
let flags = Environ.typing_flags (Global.env ()) in
flags.check_universes && flags.check_guarded
-let do_fixpoint local poly l =
+let do_fixpoint ~ontop local poly l =
let fixl, ntns = extract_fixpoint_components true l in
let (_, _, _, info as fix) = interp_fixpoint ~cofix:false fixl ntns in
let possible_indexes =
List.map compute_possible_guardness_evidences info in
- declare_fixpoint local poly fix possible_indexes ntns;
- if not (check_safe ()) then Feedback.feedback Feedback.AddedAxiom else ()
+ let pstate = declare_fixpoint ~ontop local poly fix possible_indexes ntns in
+ if not (check_safe ()) then Feedback.feedback Feedback.AddedAxiom else ();
+ pstate
-let do_cofixpoint local poly l =
+let do_cofixpoint ~ontop local poly l =
let fixl,ntns = extract_cofixpoint_components l in
let cofix = interp_fixpoint ~cofix:true fixl ntns in
- declare_cofixpoint local poly cofix ntns;
- if not (check_safe ()) then Feedback.feedback Feedback.AddedAxiom else ()
+ let pstate = declare_cofixpoint ~ontop local poly cofix ntns in
+ if not (check_safe ()) then Feedback.feedback Feedback.AddedAxiom else ();
+ pstate
diff --git a/vernac/comFixpoint.mli b/vernac/comFixpoint.mli
index 9bcb53697b..15ff5f4498 100644
--- a/vernac/comFixpoint.mli
+++ b/vernac/comFixpoint.mli
@@ -19,12 +19,14 @@ open Vernacexpr
(** Entry points for the vernacular commands Fixpoint and CoFixpoint *)
val do_fixpoint :
+ ontop:Proof_global.t option ->
(* When [false], assume guarded. *)
- locality -> polymorphic -> (fixpoint_expr * decl_notation list) list -> unit
+ locality -> polymorphic -> (fixpoint_expr * decl_notation list) list -> Proof_global.t option
val do_cofixpoint :
+ ontop:Proof_global.t option ->
(* When [false], assume guarded. *)
- locality -> polymorphic -> (cofixpoint_expr * decl_notation list) list -> unit
+ locality -> polymorphic -> (cofixpoint_expr * decl_notation list) list -> Proof_global.t option
(************************************************************************)
(** Internal API *)
@@ -81,15 +83,20 @@ val interp_fixpoint :
(** [Not used so far] *)
val declare_fixpoint :
+ ontop:Proof_global.t option ->
locality -> polymorphic ->
recursive_preentry * UState.universe_decl * UState.t *
(Constr.rel_context * Impargs.manual_implicits * int option) list ->
- Proof_global.lemma_possible_guards -> decl_notation list -> unit
+ Proof_global.lemma_possible_guards -> decl_notation list ->
+ Proof_global.t option
-val declare_cofixpoint : locality -> polymorphic ->
+val declare_cofixpoint :
+ ontop:Proof_global.t option ->
+ locality -> polymorphic ->
recursive_preentry * UState.universe_decl * UState.t *
(Constr.rel_context * Impargs.manual_implicits * int option) list ->
- decl_notation list -> unit
+ decl_notation list ->
+ Proof_global.t option
(** Very private function, do not use *)
val compute_possible_guardness_evidences :
diff --git a/vernac/declareDef.ml b/vernac/declareDef.ml
index 7dcd098183..052832244b 100644
--- a/vernac/declareDef.ml
+++ b/vernac/declareDef.ml
@@ -33,12 +33,12 @@ let get_locality id ~kind = function
| Local -> true
| Global -> false
-let declare_definition ident (local, p, k) ?hook_data ce pl imps =
+let declare_definition ~ontop ident (local, p, k) ?hook_data ce pl imps =
let fix_exn = Future.fix_exn_of ce.const_entry_body in
let gr = match local with
| Discharge when Lib.sections_are_opened () ->
let _ = declare_variable ident (Lib.cwd(), SectionLocalDef ce, IsDefinition k) in
- let () = if Proof_global.there_are_pending_proofs () then warn_definition_not_visible ident in
+ let () = if Option.has_some ontop then warn_definition_not_visible ident in
VarRef ident
| Discharge | Local | Global ->
let local = get_locality ident ~kind:"definition" local in
@@ -57,9 +57,9 @@ let declare_definition ident (local, p, k) ?hook_data ce pl imps =
end;
gr
-let declare_fix ?(opaque = false) ?hook_data (_,poly,_ as kind) pl univs f ((def,_),eff) t imps =
+let declare_fix ~ontop ?(opaque = false) ?hook_data (_,poly,_ as kind) pl univs f ((def,_),eff) t imps =
let ce = definition_entry ~opaque ~types:t ~univs ~eff def in
- declare_definition f kind ?hook_data ce pl imps
+ declare_definition ~ontop f kind ?hook_data ce pl imps
let check_definition_evars ~allow_evars sigma =
let env = Global.env () in
diff --git a/vernac/declareDef.mli b/vernac/declareDef.mli
index 3f95ec7107..8e4f4bf7fb 100644
--- a/vernac/declareDef.mli
+++ b/vernac/declareDef.mli
@@ -14,7 +14,8 @@ open Decl_kinds
val get_locality : Id.t -> kind:string -> Decl_kinds.locality -> bool
val declare_definition
- : Id.t
+ : ontop:Proof_global.t option
+ -> Id.t
-> definition_kind
-> ?hook_data:(Lemmas.declaration_hook * UState.t * (Id.t * Constr.t) list)
-> Safe_typing.private_constants Entries.definition_entry
@@ -23,7 +24,8 @@ val declare_definition
-> GlobRef.t
val declare_fix
- : ?opaque:bool
+ : ontop:Proof_global.t option
+ -> ?opaque:bool
-> ?hook_data:(Lemmas.declaration_hook * UState.t * (Id.t * Constr.t) list)
-> definition_kind
-> UnivNames.universe_binders
diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml
index 0d0732cbb4..1c7cc5e636 100644
--- a/vernac/lemmas.ml
+++ b/vernac/lemmas.ml
@@ -213,8 +213,11 @@ let save ?export_seff id const uctx do_guard (locality,poly,kind) hook universes
let default_thm_id = Id.of_string "Unnamed_thm"
-let fresh_name_for_anonymous_theorem () =
- let avoid = Id.Set.of_list (Proof_global.get_all_proof_names ()) in
+let fresh_name_for_anonymous_theorem ~pstate =
+ let avoid = match pstate with
+ | None -> Id.Set.empty
+ | Some pstate -> Id.Set.of_list (Proof_global.get_all_proof_names pstate)
+ in
next_global_ident_away default_thm_id avoid
let check_name_freshness locality {CAst.loc;v=id} : unit =
@@ -224,7 +227,7 @@ let check_name_freshness locality {CAst.loc;v=id} : unit =
then
user_err ?loc (Id.print id ++ str " already exists.")
-let save_remaining_recthms (locality,p,kind) norm univs body opaq i (id,(t_i,(_,imps))) =
+let save_remaining_recthms env sigma (locality,p,kind) norm univs body opaq i (id,(t_i,(_,imps))) =
let t_i = norm t_i in
let k = IsAssumption Conjectural in
match body with
@@ -260,7 +263,6 @@ let save_remaining_recthms (locality,p,kind) norm univs body opaq i (id,(t_i,(_,
| Lambda(na,ty,t) -> mkLambda(na,ty,body_i t)
| App (t, args) -> mkApp (body_i t, args)
| _ ->
- let sigma, env = Pfedit.get_current_context () in
anomaly Pp.(str "Not a proof by induction: " ++ Printer.pr_constr_env env sigma body ++ str ".") in
let body_i = body_i body in
match locality with
@@ -333,7 +335,7 @@ let initialize_named_context_for_proof () =
let d = if variable_opacity id then NamedDecl.drop_body d else d in
Environ.push_named_context_val d signv) sign Environ.empty_named_context_val
-let start_proof id ?pl kind sigma ?terminator ?sign ?(compute_guard=[]) ?(hook : declaration_hook option) c =
+let start_proof ~ontop id ?pl kind sigma ?terminator ?sign ?(compute_guard=[]) ?hook c =
let terminator = match terminator with
| None -> standard_proof_terminator ?hook compute_guard
| Some terminator -> terminator ?hook compute_guard
@@ -344,7 +346,7 @@ let start_proof id ?pl kind sigma ?terminator ?sign ?(compute_guard=[]) ?(hook :
| None -> initialize_named_context_for_proof ()
in
let goals = [ Global.env_of_context sign , c ] in
- Proof_global.start_proof sigma id ?pl kind goals terminator
+ Proof_global.start_proof ~ontop sigma id ?pl kind goals terminator
let rec_tac_initializer finite guard thms snl =
if finite then
@@ -360,7 +362,7 @@ let rec_tac_initializer finite guard thms snl =
| (id,n,_)::l -> Tactics.mutual_fix id n l 0
| _ -> assert false
-let start_proof_with_initialization ?hook kind sigma decl recguard thms snl =
+let start_proof_with_initialization ~ontop ?hook kind sigma decl recguard thms snl =
let intro_tac (_, (_, (ids, _))) = Tactics.auto_intros_tac ids in
let init_tac,guard = match recguard with
| Some (finite,guard,init_tac) ->
@@ -386,18 +388,20 @@ let start_proof_with_initialization ?hook kind sigma decl recguard thms snl =
let norm c = EConstr.to_constr (Evd.from_ctx ctx) c in
let body = Option.map EConstr.of_constr body in
let uctx = UState.check_univ_decl ~poly:(pi2 kind) ctx decl in
- List.map_i (save_remaining_recthms kind norm uctx body opaq) 1 other_thms in
+ let env = Global.env () in
+ List.map_i (save_remaining_recthms env sigma kind norm uctx body opaq) 1 other_thms in
let thms_data = (strength,ref,imps)::other_thms_data in
List.iter (fun (strength,ref,imps) ->
maybe_declare_manual_implicits false ref imps;
call_hook ?hook ctx [] strength ref) thms_data in
- start_proof id ~pl:decl kind sigma t ~hook ~compute_guard:guard;
- ignore (Proof_global.with_current_proof (fun _ p ->
+ let pstate = start_proof ~ontop id ~pl:decl kind sigma t ~hook ~compute_guard:guard in
+ let pstate, _ = Proof_global.with_current_proof (fun _ p ->
match init_tac with
| None -> p,(true,[])
- | Some tac -> Proof.run_tactic Global.(env ()) tac p))
+ | Some tac -> Proof.run_tactic Global.(env ()) tac p) pstate in
+ pstate
-let start_proof_com ~program_mode ?inference_hook ?hook kind thms =
+let start_proof_com ~program_mode ~ontop ?inference_hook ?hook kind thms =
let env0 = Global.env () in
let decl = fst (List.hd thms) in
let evd, decl = Constrexpr_ops.interp_univ_decl_opt env0 (snd decl) in
@@ -429,7 +433,7 @@ let start_proof_com ~program_mode ?inference_hook ?hook kind thms =
else (* We fix the variables to ensure they won't be lowered to Set *)
Evd.fix_undefined_variables evd
in
- start_proof_with_initialization ?hook kind evd decl recguard thms snl
+ start_proof_with_initialization ~ontop ?hook kind evd decl recguard thms snl
(* Saving a proof *)
@@ -444,58 +448,65 @@ let () =
optread = (fun () -> !keep_admitted_vars);
optwrite = (fun b -> keep_admitted_vars := b) }
-let save_proof ?proof = function
- | Vernacexpr.Admitted ->
- let pe =
- let open Proof_global in
- match proof with
- | Some ({ id; entries; persistence = k; universes }, _) ->
- if List.length entries <> 1 then
- user_err Pp.(str "Admitted does not support multiple statements");
- let { const_entry_secctx; const_entry_type } = List.hd entries in
- if const_entry_type = None then
- user_err Pp.(str "Admitted requires an explicit statement");
- let typ = Option.get const_entry_type in
- let ctx = UState.univ_entry ~poly:(pi2 k) universes in
- let sec_vars = if !keep_admitted_vars then const_entry_secctx else None in
- Admitted(id, k, (sec_vars, (typ, ctx), None), universes)
- | None ->
- let pftree = Proof_global.give_me_the_proof () in
- let gk = Proof_global.get_current_persistence () in
- let Proof.{ name; poly; entry } = Proof.data pftree in
- let typ = match Proofview.initial_goals entry with
- | [typ] -> snd typ
- | _ ->
- CErrors.anomaly
- ~label:"Lemmas.save_proof" (Pp.str "more than one statement.")
- in
- let typ = EConstr.Unsafe.to_constr typ in
- let universes = Proof.((data pftree).initial_euctx) in
- (* This will warn if the proof is complete *)
- let pproofs, _univs =
- Proof_global.return_proof ~allow_partial:true () in
- let sec_vars =
- if not !keep_admitted_vars then None
- else match Proof_global.get_used_variables(), pproofs with
- | Some _ as x, _ -> x
- | None, (pproof, _) :: _ ->
- let env = Global.env () in
- let ids_typ = Environ.global_vars_set env typ in
- let ids_def = Environ.global_vars_set env pproof in
- Some (Environ.keep_hyps env (Id.Set.union ids_typ ids_def))
- | _ -> None in
- let decl = Proof_global.get_universe_decl () in
- let ctx = UState.check_univ_decl ~poly universes decl in
- Admitted(name,gk,(sec_vars, (typ, ctx), None), universes)
- in
- Proof_global.apply_terminator (Proof_global.get_terminator ()) pe
- | Vernacexpr.Proved (opaque,idopt) ->
- let (proof_obj,terminator) =
- match proof with
- | None ->
- Proof_global.close_proof ~opaque ~keep_body_ucst_separate:false (fun x -> x)
- | Some proof -> proof
+let save_proof_admitted ?proof ~pstate =
+ let pe =
+ let open Proof_global in
+ match proof with
+ | Some ({ id; entries; persistence = k; universes }, _) ->
+ if List.length entries <> 1 then
+ user_err Pp.(str "Admitted does not support multiple statements");
+ let { const_entry_secctx; const_entry_type } = List.hd entries in
+ if const_entry_type = None then
+ user_err Pp.(str "Admitted requires an explicit statement");
+ let typ = Option.get const_entry_type in
+ let ctx = UState.univ_entry ~poly:(pi2 k) universes in
+ let sec_vars = if !keep_admitted_vars then const_entry_secctx else None in
+ Admitted(id, k, (sec_vars, (typ, ctx), None), universes)
+ | None ->
+ let pftree = Proof_global.give_me_the_proof pstate in
+ let gk = Proof_global.get_current_persistence pstate in
+ let Proof.{ name; poly; entry } = Proof.data pftree in
+ let typ = match Proofview.initial_goals entry with
+ | [typ] -> snd typ
+ | _ ->
+ CErrors.anomaly
+ ~label:"Lemmas.save_proof" (Pp.str "more than one statement.")
in
- (* if the proof is given explicitly, nothing has to be deleted *)
- if Option.is_empty proof then Proof_global.discard_current ();
- Proof_global.(apply_terminator terminator (Proved (opaque,idopt,proof_obj)))
+ let typ = EConstr.Unsafe.to_constr typ in
+ let universes = Proof.((data pftree).initial_euctx) in
+ (* This will warn if the proof is complete *)
+ let pproofs, _univs =
+ Proof_global.return_proof ~allow_partial:true pstate in
+ let sec_vars =
+ if not !keep_admitted_vars then None
+ else match Proof_global.get_used_variables pstate, pproofs with
+ | Some _ as x, _ -> x
+ | None, (pproof, _) :: _ ->
+ let env = Global.env () in
+ let ids_typ = Environ.global_vars_set env typ in
+ let ids_def = Environ.global_vars_set env pproof in
+ Some (Environ.keep_hyps env (Id.Set.union ids_typ ids_def))
+ | _ -> None in
+ let decl = Proof_global.get_universe_decl pstate in
+ let ctx = UState.check_univ_decl ~poly universes decl in
+ Admitted(name,gk,(sec_vars, (typ, ctx), None), universes)
+ in
+ Proof_global.apply_terminator (Proof_global.get_terminator pstate) pe
+
+let save_proof_proved ?proof ?pstate ~opaque ~idopt =
+ (* Invariant (uh) *)
+ if Option.is_empty pstate && Option.is_empty proof then
+ user_err (str "No focused proof (No proof-editing in progress).");
+ let (proof_obj,terminator) =
+ match proof with
+ | None ->
+ (* XXX: The close_proof and proof state API should be refactored
+ so it is possible to insert proofs properly into the state *)
+ let pstate = Option.get pstate in
+ Proof_global.close_proof ~opaque ~keep_body_ucst_separate:false (fun x -> x) pstate
+ | Some proof -> proof
+ in
+ (* if the proof is given explicitly, nothing has to be deleted *)
+ let pstate = if Option.is_empty proof then Proof_global.discard_current Option.(get pstate) else pstate in
+ Proof_global.(apply_terminator terminator (Proved (opaque,idopt,proof_obj)));
+ pstate
diff --git a/vernac/lemmas.mli b/vernac/lemmas.mli
index 72c666e903..1f70cfa1ad 100644
--- a/vernac/lemmas.mli
+++ b/vernac/lemmas.mli
@@ -37,30 +37,32 @@ val call_hook
-> ?fix_exn:Future.fix_exn
-> hook_type
-val start_proof : Id.t -> ?pl:UState.universe_decl -> goal_kind -> Evd.evar_map ->
+val start_proof : ontop:Proof_global.t option -> Id.t -> ?pl:UState.universe_decl -> goal_kind -> Evd.evar_map ->
?terminator:(?hook:declaration_hook -> Proof_global.lemma_possible_guards -> Proof_global.proof_terminator) ->
?sign:Environ.named_context_val ->
?compute_guard:Proof_global.lemma_possible_guards ->
- ?hook:declaration_hook -> EConstr.types -> unit
+ ?hook:declaration_hook -> EConstr.types -> Proof_global.t
-val start_proof_com :
- program_mode:bool -> ?inference_hook:Pretyping.inference_hook ->
- ?hook:declaration_hook -> goal_kind -> Vernacexpr.proof_expr list ->
- unit
+val start_proof_com
+ : program_mode:bool
+ -> ontop:Proof_global.t option
+ -> ?inference_hook:Pretyping.inference_hook
+ -> ?hook:declaration_hook -> goal_kind -> Vernacexpr.proof_expr list
+ -> Proof_global.t
-val start_proof_with_initialization :
+val start_proof_with_initialization : ontop:Proof_global.t option ->
?hook:declaration_hook ->
goal_kind -> Evd.evar_map -> UState.universe_decl ->
(bool * Proof_global.lemma_possible_guards * unit Proofview.tactic list option) option ->
(Id.t (* name of thm *) *
- (EConstr.types (* type of thm *) * (Name.t list (* names to pre-introduce *) * Impargs.manual_explicitation list))) list ->
- int list option -> unit
+ (EConstr.types (* type of thm *) * (Name.t list (* names to pre-introduce *) * Impargs.manual_explicitation list))) list
+ -> int list option -> Proof_global.t
val standard_proof_terminator :
?hook:declaration_hook -> Proof_global.lemma_possible_guards ->
Proof_global.proof_terminator
-val fresh_name_for_anonymous_theorem : unit -> Id.t
+val fresh_name_for_anonymous_theorem : pstate:Proof_global.t option -> Id.t
(* Prepare global named context for proof session: remove proofs of
opaque section definitions and remove vm-compiled code *)
@@ -69,4 +71,14 @@ val initialize_named_context_for_proof : unit -> Environ.named_context_val
(** {6 ... } *)
-val save_proof : ?proof:Proof_global.closed_proof -> Vernacexpr.proof_end -> unit
+val save_proof_admitted
+ : ?proof:Proof_global.closed_proof
+ -> pstate:Proof_global.t
+ -> unit
+
+val save_proof_proved
+ : ?proof:Proof_global.closed_proof
+ -> ?pstate:Proof_global.t
+ -> opaque:Proof_global.opacity_flag
+ -> idopt:Names.lident option
+ -> Proof_global.t option
diff --git a/vernac/obligations.ml b/vernac/obligations.ml
index 9aca48f529..07194578c1 100644
--- a/vernac/obligations.ml
+++ b/vernac/obligations.ml
@@ -456,7 +456,7 @@ let obligation_substitution expand prg =
let ints = intset_to (pred (Array.length obls)) in
obl_substitution expand obls ints
-let declare_definition prg =
+let declare_definition ~ontop prg =
let varsubst = obligation_substitution true prg in
let body, typ = subst_prog varsubst prg in
let nf = UnivSubst.nf_evars_and_universes_opt_subst (fun x -> None)
@@ -475,7 +475,7 @@ let declare_definition prg =
let () = progmap_remove prg in
let ubinders = UState.universe_binders uctx in
let hook_data = Option.map (fun hook -> hook, uctx, obls) prg.prg_hook in
- DeclareDef.declare_definition prg.prg_name
+ DeclareDef.declare_definition ~ontop prg.prg_name
prg.prg_kind ce ubinders prg.prg_implicits ?hook_data
let rec lam_index n t acc =
@@ -554,16 +554,14 @@ let declare_mutual_definition l =
(* Declare the recursive definitions *)
let univs = UState.univ_entry ~poly first.prg_ctx in
let fix_exn = Hook.get get_fix_exn () in
- let kns = List.map4
- (DeclareDef.declare_fix ~opaque (local, poly, kind) UnivNames.empty_binders univs)
- fixnames fixdecls fixtypes fiximps
- in
- (* Declare notations *)
- List.iter (Metasyntax.add_notation_interpretation (Global.env())) first.prg_notations;
- Declare.recursive_message (fixkind != IsCoFixpoint) indexes fixnames;
- let gr = List.hd kns in
- Lemmas.call_hook ?hook:first.prg_hook ~fix_exn first.prg_ctx obls local gr;
- List.iter progmap_remove l; gr
+ let kns = List.map4 (DeclareDef.declare_fix ~ontop:None ~opaque (local, poly, kind) UnivNames.empty_binders univs)
+ fixnames fixdecls fixtypes fiximps in
+ (* Declare notations *)
+ List.iter (Metasyntax.add_notation_interpretation (Global.env())) first.prg_notations;
+ Declare.recursive_message (fixkind != IsCoFixpoint) indexes fixnames;
+ let gr = List.hd kns in
+ Lemmas.call_hook ?hook:first.prg_hook ~fix_exn first.prg_ctx obls local gr;
+ List.iter progmap_remove l; gr
let decompose_lam_prod c ty =
let open Context.Rel.Declaration in
@@ -763,7 +761,7 @@ let update_obls prg obls rem =
else (
match prg'.prg_deps with
| [] ->
- let kn = declare_definition prg' in
+ let kn = declare_definition ~ontop:None prg' in
progmap_remove prg';
Defined kn
| l ->
@@ -948,7 +946,7 @@ let obligation_hook prg obl num auto ctx' _ _ gr =
ignore (auto (Some prg.prg_name) None deps)
end
-let rec solve_obligation prg num tac =
+let rec solve_obligation ~ontop prg num tac =
let user_num = succ num in
let obls, rem = prg.prg_obligations in
let obl = obls.(num) in
@@ -967,20 +965,21 @@ let rec solve_obligation prg num tac =
let auto n tac oblset = auto_solve_obligations n ~oblset tac in
let terminator ?hook guard =
Proof_global.make_terminator
- (obligation_terminator ?hook prg.prg_name num guard auto) in
+ (obligation_terminator prg.prg_name num guard ?hook auto) in
let hook = Lemmas.mk_hook (obligation_hook prg obl num auto) in
- let () = Lemmas.start_proof ~sign:prg.prg_sign obl.obl_name kind evd (EConstr.of_constr obl.obl_type) ~terminator ~hook in
- let _ = Pfedit.by !default_tactic in
- Option.iter (fun tac -> Proof_global.set_endline_tactic tac) tac
+ let pstate = Lemmas.start_proof ~ontop ~sign:prg.prg_sign obl.obl_name kind evd (EConstr.of_constr obl.obl_type) ~terminator ~hook in
+ let pstate = fst @@ Pfedit.by !default_tactic pstate in
+ let pstate = Option.cata (fun tac -> Proof_global.set_endline_tactic tac pstate) pstate tac in
+ pstate
-and obligation (user_num, name, typ) tac =
+and obligation ~ontop (user_num, name, typ) tac =
let num = pred user_num in
let prg = get_prog_err name in
let obls, rem = prg.prg_obligations in
if num >= 0 && num < Array.length obls then
let obl = obls.(num) in
match obl.obl_body with
- None -> solve_obligation prg num tac
+ | None -> solve_obligation ~ontop prg num tac
| Some r -> error "Obligation already solved"
else error (sprintf "Unknown obligation number %i" (succ num))
@@ -1113,7 +1112,7 @@ let add_definition n ?term t ctx ?(univdecl=UState.default_univ_decl)
let obls,_ = prg.prg_obligations in
if Int.equal (Array.length obls) 0 then (
Flags.if_verbose Feedback.msg_info (info ++ str ".");
- let cst = declare_definition prg in
+ let cst = declare_definition ~ontop:None prg in
Defined cst)
else (
let len = Array.length obls in
@@ -1180,7 +1179,7 @@ let admit_obligations n =
let prg = get_prog_err n in
admit_prog prg
-let next_obligation n tac =
+let next_obligation ~ontop n tac =
let prg = match n with
| None -> get_any_prog_err ()
| Some _ -> get_prog_err n
@@ -1191,7 +1190,7 @@ let next_obligation n tac =
| Some i -> i
| None -> anomaly (Pp.str "Could not find a solvable obligation.")
in
- solve_obligation prg i tac
+ solve_obligation ~ontop prg i tac
let check_program_libraries () =
Coqlib.check_required_library Coqlib.datatypes_module_name;
diff --git a/vernac/obligations.mli b/vernac/obligations.mli
index c5720363b4..b1b7b1ec90 100644
--- a/vernac/obligations.mli
+++ b/vernac/obligations.mli
@@ -85,10 +85,17 @@ val add_mutual_definitions :
notations ->
fixpoint_kind -> unit
-val obligation : int * Names.Id.t option * Constrexpr.constr_expr option ->
- Genarg.glob_generic_argument option -> unit
-
-val next_obligation : Names.Id.t option -> Genarg.glob_generic_argument option -> unit
+val obligation
+ : ontop:Proof_global.t option
+ -> int * Names.Id.t option * Constrexpr.constr_expr option
+ -> Genarg.glob_generic_argument option
+ -> Proof_global.t
+
+val next_obligation
+ : ontop:Proof_global.t option
+ -> Names.Id.t option
+ -> Genarg.glob_generic_argument option
+ -> Proof_global.t
val solve_obligations : Names.Id.t option -> unit Proofview.tactic option -> progress
(* Number of remaining obligations to be solved for this program *)
diff --git a/vernac/search.ml b/vernac/search.ml
index 6610789626..e41378908f 100644
--- a/vernac/search.ml
+++ b/vernac/search.ml
@@ -59,11 +59,16 @@ let iter_constructors indsp u fn env nconstr =
let iter_named_context_name_type f =
List.iter (fun decl -> f (NamedDecl.get_id decl) (NamedDecl.get_type decl))
+let get_current_or_goal_context ?pstate glnum =
+ match pstate with
+ | None -> let env = Global.env () in Evd.(from_env env, env)
+ | Some p -> Pfedit.get_goal_context p glnum
+
(* General search over hypothesis of a goal *)
-let iter_hypothesis glnum (fn : GlobRef.t -> env -> constr -> unit) =
+let iter_hypothesis ?pstate glnum (fn : GlobRef.t -> env -> constr -> unit) =
let env = Global.env () in
let iter_hyp idh typ = fn (VarRef idh) env typ in
- let evmap,e = Pfedit.get_goal_context glnum in
+ let evmap,e = get_current_or_goal_context ?pstate glnum in
let pfctxt = named_context e in
iter_named_context_name_type iter_hyp pfctxt
@@ -99,10 +104,10 @@ let iter_declarations (fn : GlobRef.t -> env -> constr -> unit) =
try Declaremods.iter_all_segments iter_obj
with Not_found -> ()
-let generic_search glnumopt fn =
+let generic_search ?pstate glnumopt fn =
(match glnumopt with
| None -> ()
- | Some glnum -> iter_hypothesis glnum fn);
+ | Some glnum -> iter_hypothesis ?pstate glnum fn);
iter_declarations fn
(** This module defines a preference on constrs in the form of a
@@ -221,7 +226,7 @@ let search_about_filter query gr env typ = match query with
(** SearchPattern *)
-let search_pattern gopt pat mods pr_search =
+let search_pattern ?pstate gopt pat mods pr_search =
let blacklist_filter = blacklist_filter_aux () in
let filter ref env typ =
module_filter mods ref env typ &&
@@ -231,7 +236,7 @@ let search_pattern gopt pat mods pr_search =
let iter ref env typ =
if filter ref env typ then pr_search ref env typ
in
- generic_search gopt iter
+ generic_search ?pstate gopt iter
(** SearchRewrite *)
@@ -243,7 +248,7 @@ let rewrite_pat1 pat =
let rewrite_pat2 pat =
PApp (PRef (eq ()), [| PMeta None; PMeta None; pat |])
-let search_rewrite gopt pat mods pr_search =
+let search_rewrite ?pstate gopt pat mods pr_search =
let pat1 = rewrite_pat1 pat in
let pat2 = rewrite_pat2 pat in
let blacklist_filter = blacklist_filter_aux () in
@@ -256,11 +261,11 @@ let search_rewrite gopt pat mods pr_search =
let iter ref env typ =
if filter ref env typ then pr_search ref env typ
in
- generic_search gopt iter
+ generic_search ?pstate gopt iter
(** Search *)
-let search_by_head gopt pat mods pr_search =
+let search_by_head ?pstate gopt pat mods pr_search =
let blacklist_filter = blacklist_filter_aux () in
let filter ref env typ =
module_filter mods ref env typ &&
@@ -270,11 +275,11 @@ let search_by_head gopt pat mods pr_search =
let iter ref env typ =
if filter ref env typ then pr_search ref env typ
in
- generic_search gopt iter
+ generic_search ?pstate gopt iter
(** SearchAbout *)
-let search_about gopt items mods pr_search =
+let search_about ?pstate gopt items mods pr_search =
let blacklist_filter = blacklist_filter_aux () in
let filter ref env typ =
let eqb b1 b2 = if b1 then b2 else not b2 in
@@ -286,7 +291,7 @@ let search_about gopt items mods pr_search =
let iter ref env typ =
if filter ref env typ then pr_search ref env typ
in
- generic_search gopt iter
+ generic_search ?pstate gopt iter
type search_constraint =
| Name_Pattern of Str.regexp
@@ -301,7 +306,7 @@ type 'a coq_object = {
coq_object_object : 'a;
}
-let interface_search =
+let interface_search ?pstate =
let rec extract_flags name tpe subtpe mods blacklist = function
| [] -> (name, tpe, subtpe, mods, blacklist)
| (Name_Pattern regexp, b) :: l ->
@@ -371,7 +376,7 @@ let interface_search =
let iter ref env typ =
if filter_function ref env typ then print_function ref env typ
in
- let () = generic_search glnum iter in
+ let () = generic_search ?pstate glnum iter in
!ans
let blacklist_filter ref env typ =
diff --git a/vernac/search.mli b/vernac/search.mli
index ecbb02bc68..0f94ddc5b6 100644
--- a/vernac/search.mli
+++ b/vernac/search.mli
@@ -39,13 +39,13 @@ val search_about_filter : glob_search_about_item -> filter_function
goal and the global environment for things matching [pattern] and
satisfying module exclude/include clauses of [modinout]. *)
-val search_by_head : int option -> constr_pattern -> DirPath.t list * bool
+val search_by_head : ?pstate:Proof_global.t -> int option -> constr_pattern -> DirPath.t list * bool
-> display_function -> unit
-val search_rewrite : int option -> constr_pattern -> DirPath.t list * bool
+val search_rewrite : ?pstate:Proof_global.t -> int option -> constr_pattern -> DirPath.t list * bool
-> display_function -> unit
-val search_pattern : int option -> constr_pattern -> DirPath.t list * bool
+val search_pattern : ?pstate:Proof_global.t -> int option -> constr_pattern -> DirPath.t list * bool
-> display_function -> unit
-val search_about : int option -> (bool * glob_search_about_item) list
+val search_about : ?pstate:Proof_global.t -> int option -> (bool * glob_search_about_item) list
-> DirPath.t list * bool -> display_function -> unit
type search_constraint =
@@ -66,12 +66,12 @@ type 'a coq_object = {
coq_object_object : 'a;
}
-val interface_search : ?glnum:int -> (search_constraint * bool) list ->
+val interface_search : ?pstate:Proof_global.t -> ?glnum:int -> (search_constraint * bool) list ->
constr coq_object list
(** {6 Generic search function} *)
-val generic_search : int option -> display_function -> unit
+val generic_search : ?pstate:Proof_global.t -> int option -> display_function -> unit
(** This function iterates over all hypothesis of the goal numbered
[glnum] (if present) and all known declarations. *)
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
index 4250ddb02c..91965e56b1 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -44,6 +44,28 @@ let vernac_pperr_endline pp =
(* Misc *)
+let there_are_pending_proofs ~pstate =
+ not Option.(is_empty pstate)
+
+let check_no_pending_proof ~pstate =
+ if there_are_pending_proofs ~pstate then
+ user_err Pp.(str "Command not supported (Open proofs remain)")
+
+let vernac_require_open_proof ~pstate f =
+ match pstate with
+ | Some pstate -> f ~pstate
+ | None -> user_err Pp.(str "Command not supported (No proof-editing in progress)")
+
+let get_current_or_global_context ~pstate =
+ match pstate with
+ | None -> let env = Global.env () in Evd.(from_env env, env)
+ | Some p -> Pfedit.get_current_context p
+
+let get_goal_or_global_context ~pstate glnum =
+ match pstate with
+ | None -> let env = Global.env () in Evd.(from_env env, env)
+ | Some p -> Pfedit.get_goal_context p glnum
+
let cl_of_qualid = function
| FunClass -> Classops.CL_FUN
| SortClass -> Classops.CL_SORT
@@ -72,30 +94,37 @@ end
(*******************)
(* "Show" commands *)
-let show_proof () =
+let show_proof ~pstate =
(* spiwack: this would probably be cooler with a bit of polishing. *)
- let p = Proof_global.give_me_the_proof () in
- let sigma, env = Pfedit.get_current_context () in
- let pprf = Proof.partial_proof p in
- Pp.prlist_with_sep Pp.fnl (Printer.pr_econstr_env env sigma) pprf
+ try
+ let pstate = Option.get pstate in
+ let p = Proof_global.give_me_the_proof pstate in
+ let sigma, env = Pfedit.get_current_context pstate in
+ let pprf = Proof.partial_proof p in
+ Pp.prlist_with_sep Pp.fnl (Printer.pr_econstr_env env sigma) pprf
+ (* We print nothing if there are no goals left *)
+ with
+ | Pfedit.NoSuchGoal
+ | Option.IsNone ->
+ user_err (str "No goals to show.")
-let show_top_evars () =
+let show_top_evars ~pstate =
(* spiwack: new as of Feb. 2010: shows goal evars in addition to non-goal evars. *)
- let pfts = Proof_global.give_me_the_proof () in
+ let pfts = Proof_global.give_me_the_proof pstate in
let Proof.{goals;shelf;given_up;sigma} = Proof.data pfts in
pr_evars_int sigma ~shelf ~given_up 1 (Evd.undefined_map sigma)
-let show_universes () =
- let pfts = Proof_global.give_me_the_proof () in
+let show_universes ~pstate =
+ let pfts = Proof_global.give_me_the_proof pstate in
let Proof.{goals;sigma} = Proof.data pfts in
let ctx = Evd.universe_context_set (Evd.minimize_universes sigma) in
Termops.pr_evar_universe_context (Evd.evar_universe_context sigma) ++ fnl () ++
str "Normalized constraints: " ++ Univ.pr_universe_context_set (Termops.pr_evd_level sigma) ctx
(* Simulate the Intro(s) tactic *)
-let show_intro all =
+let show_intro ~pstate all =
let open EConstr in
- let pf = Proof_global.give_me_the_proof() in
+ let pf = Proof_global.give_me_the_proof pstate in
let Proof.{goals;sigma} = Proof.data pf in
if not (List.is_empty goals) then begin
let gl = {Evd.it=List.hd goals ; sigma = sigma; } in
@@ -224,7 +253,7 @@ let print_modtype qid =
with Not_found ->
user_err (str"Unknown Module Type or Module " ++ pr_qualid qid)
-let print_namespace ns =
+let print_namespace ~pstate ns =
let ns = List.rev (Names.DirPath.repr ns) in
(* [match_dirpath], [match_modulpath] are helpers for [matches]
which checks whether a constant is in the namespace [ns]. *)
@@ -272,10 +301,10 @@ let print_namespace ns =
let qn = (qualified_minus (List.length ns) mp)@[Names.Label.to_id lbl] in
print_list Id.print qn
in
- let print_constant k body =
+ let print_constant ~pstate k body =
(* FIXME: universes *)
let t = body.Declarations.const_type in
- let sigma, env = Pfedit.get_current_context () in
+ let sigma, env = get_current_or_global_context ~pstate in
print_kn k ++ str":" ++ spc() ++ Printer.pr_type_env env sigma t
in
let matches mp = match match_modulepath ns mp with
@@ -285,7 +314,7 @@ let print_namespace ns =
Environ.fold_constants (fun c body acc ->
let kn = Constant.user c in
if matches (KerName.modpath kn)
- then acc++fnl()++hov 2 (print_constant kn body)
+ then acc++fnl()++hov 2 (print_constant ~pstate kn body)
else acc)
(Global.env ()) (str"")
in
@@ -515,7 +544,7 @@ let () =
(***********)
(* Gallina *)
-let start_proof_and_print ~program_mode ?hook k l =
+let start_proof_and_print ~program_mode ~pstate ?hook k l =
let inference_hook =
if program_mode then
let hook env sigma ev =
@@ -537,7 +566,7 @@ let start_proof_and_print ~program_mode ?hook k l =
in Some hook
else None
in
- start_proof_com ~program_mode ?inference_hook ?hook k l
+ start_proof_com ~program_mode ~ontop:pstate ?inference_hook ?hook k l
let vernac_definition_hook p = function
| Coercion ->
@@ -548,7 +577,7 @@ let vernac_definition_hook p = function
Some (Class.add_subclass_hook p)
| _ -> None
-let vernac_definition ~atts discharge kind ({loc;v=id}, pl) def =
+let vernac_definition ~atts ~pstate discharge kind ({loc;v=id}, pl) def =
let open DefAttributes in
let local = enforce_locality_exp atts.locality discharge in
let hook = vernac_definition_hook atts.polymorphic kind in
@@ -563,41 +592,47 @@ let vernac_definition ~atts discharge kind ({loc;v=id}, pl) def =
let program_mode = atts.program in
let name =
match id with
- | Anonymous -> fresh_name_for_anonymous_theorem ()
+ | Anonymous -> fresh_name_for_anonymous_theorem ~pstate
| Name n -> n
in
(match def with
| ProveBody (bl,t) -> (* local binders, typ *)
- start_proof_and_print ~program_mode (local, atts.polymorphic, DefinitionBody kind)
- ?hook [(CAst.make ?loc name, pl), (bl, t)]
+ Some (start_proof_and_print ~program_mode ~pstate (local, atts.polymorphic, DefinitionBody kind)
+ ?hook [(CAst.make ?loc name, pl), (bl, t)])
| DefineBody (bl,red_option,c,typ_opt) ->
let red_option = match red_option with
- | None -> None
- | Some r ->
- let sigma, env = Pfedit.get_current_context () in
- Some (snd (Hook.get f_interp_redexp env sigma r)) in
- ComDefinition.do_definition ~program_mode name
- (local, atts.polymorphic, kind) pl bl red_option c typ_opt ?hook)
-
-let vernac_start_proof ~atts kind l =
+ | None -> None
+ | Some r ->
+ let sigma, env = get_current_or_global_context ~pstate in
+ Some (snd (Hook.get f_interp_redexp env sigma r)) in
+ ComDefinition.do_definition ~ontop:pstate ~program_mode name
+ (local, atts.polymorphic, kind) pl bl red_option c typ_opt ?hook;
+ pstate
+ )
+
+let vernac_start_proof ~atts ~pstate kind l =
let open DefAttributes in
let local = enforce_locality_exp atts.locality NoDischarge in
if Dumpglob.dump () then
List.iter (fun ((id, _), _) -> Dumpglob.dump_definition id false "prf") l;
- start_proof_and_print ~program_mode:atts.program (local, atts.polymorphic, Proof kind) l
+ Some (start_proof_and_print ~pstate ~program_mode:atts.program (local, atts.polymorphic, Proof kind) l)
-let vernac_end_proof ?proof = function
- | Admitted -> save_proof ?proof Admitted
- | Proved (_,_) as e -> save_proof ?proof e
+let vernac_end_proof ?pstate ?proof = function
+ | Admitted ->
+ vernac_require_open_proof ~pstate (save_proof_admitted ?proof);
+ pstate
+ | Proved (opaque,idopt) ->
+ save_proof_proved ?pstate ?proof ~opaque ~idopt
-let vernac_exact_proof c =
+let vernac_exact_proof ~pstate c =
(* spiwack: for simplicity I do not enforce that "Proof proof_term" is
called only at the begining of a proof. *)
- let status = Pfedit.by (Tactics.exact_proof c) in
- save_proof (Vernacexpr.(Proved(Proof_global.Opaque,None)));
- if not status then Feedback.feedback Feedback.AddedAxiom
+ let pstate, status = Pfedit.by (Tactics.exact_proof c) pstate in
+ let pstate = save_proof_proved ?proof:None ~pstate ~opaque:Proof_global.Opaque ~idopt:None in
+ if not status then Feedback.feedback Feedback.AddedAxiom;
+ pstate
-let vernac_assumption ~atts discharge kind l nl =
+let vernac_assumption ~atts ~pstate discharge kind l nl =
let open DefAttributes in
let local = enforce_locality_exp atts.locality discharge in
let global = local == Global in
@@ -607,7 +642,7 @@ let vernac_assumption ~atts discharge kind l nl =
List.iter (fun (lid, _) ->
if global then Dumpglob.dump_definition lid false "ax"
else Dumpglob.dump_definition lid true "var") idl) l;
- let status = ComAssumption.do_assumptions ~program_mode:atts.program kind nl l in
+ let status = ComAssumption.do_assumptions ~pstate ~program_mode:atts.program kind nl l in
if not status then Feedback.feedback Feedback.AddedAxiom
let is_polymorphic_inductive_cumulativity =
@@ -772,28 +807,28 @@ let vernac_inductive ~atts cum lo finite indl =
in vernac_record cum (Class true) atts.polymorphic finite [id, bl, c, None, [f]]
*)
-let vernac_fixpoint ~atts discharge l =
+let vernac_fixpoint ~atts ~pstate discharge l : Proof_global.t option =
let open DefAttributes in
let local = enforce_locality_exp atts.locality discharge in
if Dumpglob.dump () then
List.iter (fun (((lid,_), _, _, _, _), _) -> Dumpglob.dump_definition lid false "def") l;
(* XXX: Switch to the attribute system and match on ~atts *)
let do_fixpoint = if atts.program then
- ComProgramFixpoint.do_fixpoint
+ fun local sign l -> ComProgramFixpoint.do_fixpoint local sign l; None
else
- ComFixpoint.do_fixpoint
+ ComFixpoint.do_fixpoint ~ontop:pstate
in
do_fixpoint local atts.polymorphic l
-let vernac_cofixpoint ~atts discharge l =
+let vernac_cofixpoint ~atts ~pstate discharge l =
let open DefAttributes in
let local = enforce_locality_exp atts.locality discharge in
if Dumpglob.dump () then
List.iter (fun (((lid,_), _, _, _), _) -> Dumpglob.dump_definition lid false "def") l;
let do_cofixpoint = if atts.program then
- ComProgramFixpoint.do_cofixpoint
+ fun local sign l -> ComProgramFixpoint.do_cofixpoint local sign l; None
else
- ComFixpoint.do_cofixpoint
+ ComFixpoint.do_cofixpoint ~ontop:pstate
in
do_cofixpoint local atts.polymorphic l
@@ -851,14 +886,14 @@ let vernac_declare_module export {loc;v=id} binders_ast mty_ast =
Flags.if_verbose Feedback.msg_info (str "Module " ++ Id.print id ++ str " is declared");
Option.iter (fun export -> vernac_import export [qualid_of_ident id]) export
-let vernac_define_module export {loc;v=id} (binders_ast : module_binder list) mty_ast_o mexpr_ast_l =
+let vernac_define_module ~pstate export {loc;v=id} (binders_ast : module_binder list) mty_ast_o mexpr_ast_l =
(* We check the state of the system (in section, in module type)
and what module information is supplied *)
if Lib.sections_are_opened () then
user_err Pp.(str "Modules and Module Types are not allowed inside sections.");
match mexpr_ast_l with
| [] ->
- Proof_global.check_no_pending_proof ();
+ check_no_pending_proof ~pstate;
let binders_ast,argsexport =
List.fold_right
(fun (export,idl,ty) (args,argsexport) ->
@@ -898,13 +933,13 @@ let vernac_end_module export {loc;v=id} =
Flags.if_verbose Feedback.msg_info (str "Module " ++ Id.print id ++ str " is defined");
Option.iter (fun export -> vernac_import export [qualid_of_ident ?loc id]) export
-let vernac_declare_module_type {loc;v=id} binders_ast mty_sign mty_ast_l =
+let vernac_declare_module_type ~pstate {loc;v=id} binders_ast mty_sign mty_ast_l =
if Lib.sections_are_opened () then
user_err Pp.(str "Modules and Module Types are not allowed inside sections.");
match mty_ast_l with
| [] ->
- Proof_global.check_no_pending_proof ();
+ check_no_pending_proof ~pstate;
let binders_ast,argsexport =
List.fold_right
(fun (export,idl,ty) (args,argsexport) ->
@@ -951,8 +986,8 @@ let vernac_include l =
(* Sections *)
-let vernac_begin_section ({v=id} as lid) =
- Proof_global.check_no_pending_proof ();
+let vernac_begin_section ~pstate ({v=id} as lid) =
+ check_no_pending_proof ~pstate;
Dumpglob.dump_definition lid true "sec";
Lib.open_section id
@@ -965,8 +1000,8 @@ let vernac_name_sec_hyp {v=id} set = Proof_using.name_set id set
(* Dispatcher of the "End" command *)
-let vernac_end_segment ({v=id} as lid) =
- Proof_global.check_no_pending_proof ();
+let vernac_end_segment ~pstate ({v=id} as lid) =
+ check_no_pending_proof ~pstate;
match Lib.find_opening_node id with
| Lib.OpenedModule (false,export,_,_) -> vernac_end_module export lid
| Lib.OpenedModule (true,_,_,_) -> vernac_end_modtype lid
@@ -1031,7 +1066,7 @@ let vernac_instance ~atts sup inst props pri =
let global = not (make_section_locality atts.locality) in
Dumpglob.dump_constraint (fst (pi1 inst)) false "inst";
let program_mode = atts.program in
- ignore(Classes.new_instance ~program_mode ~global atts.polymorphic sup inst props pri)
+ Classes.new_instance ~program_mode ~global atts.polymorphic sup inst props pri
let vernac_declare_instance ~atts sup inst pri =
let open DefAttributes in
@@ -1039,8 +1074,8 @@ let vernac_declare_instance ~atts sup inst pri =
Dumpglob.dump_definition (fst (pi1 inst)) false "inst";
Classes.declare_new_instance ~program_mode:atts.program ~global atts.polymorphic sup inst pri
-let vernac_context ~poly l =
- if not (Classes.context poly l) then Feedback.feedback Feedback.AddedAxiom
+let vernac_context ~pstate ~poly l =
+ if not (Classes.context ~pstate poly l) then Feedback.feedback Feedback.AddedAxiom
let vernac_existing_instance ~section_local insts =
let glob = not section_local in
@@ -1061,21 +1096,19 @@ let focus_command_cond = Proof.no_cond command_focus
there are no more goals to solve. It cannot be a tactic since
all tactics fail if there are no further goals to prove. *)
-let vernac_solve_existential = Pfedit.instantiate_nth_evar_com
+let vernac_solve_existential ~pstate i e = Pfedit.instantiate_nth_evar_com i e pstate
-let vernac_set_end_tac tac =
+let vernac_set_end_tac ~pstate tac =
let env = Genintern.empty_glob_sign (Global.env ()) in
let _, tac = Genintern.generic_intern env tac in
- if not (Proof_global.there_are_pending_proofs ()) then
- user_err Pp.(str "Unknown command of the non proof-editing mode.");
- Proof_global.set_endline_tactic tac
- (* TO DO verifier s'il faut pas mettre exist s | TacId s ici*)
+ (* TO DO verifier s'il faut pas mettre exist s | TacId s ici*)
+ Proof_global.set_endline_tactic tac pstate
-let vernac_set_used_variables e =
+let vernac_set_used_variables ~(pstate : Proof_global.t) e : Proof_global.t =
let env = Global.env () in
let initial_goals pf = Proofview.initial_goals Proof.(data pf).Proof.entry in
let tys =
- List.map snd (initial_goals (Proof_global.give_me_the_proof ())) in
+ List.map snd (initial_goals (Proof_global.give_me_the_proof pstate)) in
let tys = List.map EConstr.Unsafe.to_constr tys in
let l = Proof_using.process_expr env e tys in
let vars = Environ.named_context env in
@@ -1084,10 +1117,10 @@ let vernac_set_used_variables e =
user_err ~hdr:"vernac_set_used_variables"
(str "Unknown variable: " ++ Id.print id))
l;
- ignore (Proof_global.set_used_variables l);
- Proof_global.with_current_proof begin fun _ p ->
+ let _, pstate = Proof_global.set_used_variables pstate l in
+ fst @@ Proof_global.with_current_proof begin fun _ p ->
(p, ())
- end
+ end pstate
(*****************************)
(* Auxiliary file management *)
@@ -1132,12 +1165,10 @@ let vernac_chdir = function
(* State management *)
let vernac_write_state file =
- Proof_global.discard_all ();
let file = CUnix.make_suffix file ".coq" in
States.extern_state file
let vernac_restore_state file =
- Proof_global.discard_all ();
let file = Loadpath.locate_file (CUnix.make_suffix file ".coq") in
States.intern_state file
@@ -1730,9 +1761,14 @@ let vernac_print_option key =
try print_option_value key
with Not_found -> error_undeclared_key key
-let get_current_context_of_args = function
- | Some n -> Pfedit.get_goal_context n
- | None -> Pfedit.get_current_context ()
+let get_current_context_of_args ~pstate =
+ match pstate with
+ | None -> fun _ ->
+ let env = Global.env () in Evd.(from_env env, env)
+ | Some pstate ->
+ function
+ | Some n -> Pfedit.get_goal_context pstate n
+ | None -> Pfedit.get_current_context pstate
let query_command_selector ?loc = function
| None -> None
@@ -1740,9 +1776,9 @@ let query_command_selector ?loc = function
| _ -> user_err ?loc ~hdr:"query_command_selector"
(str "Query commands only support the single numbered goal selector.")
-let vernac_check_may_eval ~atts redexp glopt rc =
+let vernac_check_may_eval ~pstate ~atts redexp glopt rc =
let glopt = query_command_selector glopt in
- let (sigma, env) = get_current_context_of_args glopt in
+ let sigma, env = get_current_context_of_args ~pstate glopt in
let sigma, c = interp_open_constr env sigma rc in
let sigma = Evarconv.solve_unif_constraints_with_heuristics env sigma in
Evarconv.check_problems_are_solved env sigma;
@@ -1796,27 +1832,33 @@ let vernac_global_check c =
pr_universe_ctx_set sigma uctx
-let get_nth_goal n =
- let pf = Proof_global.give_me_the_proof() in
+let get_nth_goal ~pstate n =
+ let pf = Proof_global.give_me_the_proof pstate in
let Proof.{goals;sigma} = Proof.data pf in
let gl = {Evd.it=List.nth goals (n-1) ; sigma = sigma; } in
gl
exception NoHyp
+
(* Printing "About" information of a hypothesis of the current goal.
We only print the type and a small statement to this comes from the
goal. Precondition: there must be at least one current goal. *)
-let print_about_hyp_globs ?loc ref_or_by_not udecl glopt =
+let print_about_hyp_globs ~pstate ?loc ref_or_by_not udecl glopt =
let open Context.Named.Declaration in
try
+ (* Fallback early to globals *)
+ let pstate = match pstate with
+ | None -> raise Not_found
+ | Some pstate -> pstate
+ in
(* FIXME error on non None udecl if we find the hyp. *)
let glnumopt = query_command_selector ?loc glopt in
let gl,id =
match glnumopt, ref_or_by_not.v with
| None,AN qid when qualid_is_ident qid -> (* goal number not given, catch any failure *)
- (try get_nth_goal 1, qualid_basename qid with _ -> raise NoHyp)
+ (try get_nth_goal ~pstate 1, qualid_basename qid with _ -> raise NoHyp)
| Some n,AN qid when qualid_is_ident qid -> (* goal number given, catch if wong *)
- (try get_nth_goal n, qualid_basename qid
+ (try get_nth_goal ~pstate n, qualid_basename qid
with
Failure _ -> user_err ?loc ~hdr:"print_about_hyp_globs"
(str "No such goal: " ++ int n ++ str "."))
@@ -1826,15 +1868,16 @@ let print_about_hyp_globs ?loc ref_or_by_not udecl glopt =
let natureofid = match decl with
| LocalAssum _ -> "Hypothesis"
| LocalDef (_,bdy,_) ->"Constant (let in)" in
- let sigma, env = Pfedit.get_current_context () in
+ let sigma, env = Pfedit.get_current_context pstate in
v 0 (Id.print id ++ str":" ++ pr_econstr_env env sigma (NamedDecl.get_type decl) ++ fnl() ++ fnl()
++ str natureofid ++ str " of the goal context.")
with (* fallback to globals *)
| NoHyp | Not_found ->
- let sigma, env = Pfedit.get_current_context () in
+ let sigma, env = get_current_or_global_context ~pstate in
print_about env sigma ref_or_by_not udecl
-let vernac_print ~atts env sigma =
+let vernac_print ~(pstate : Proof_global.t option) ~atts =
+ let sigma, env = get_current_or_global_context ~pstate in
function
| PrintTables -> print_tables ()
| PrintFullContext-> print_full_context_typ env sigma
@@ -1845,7 +1888,7 @@ let vernac_print ~atts env sigma =
| PrintModules -> print_modules ()
| PrintModule qid -> print_module qid
| PrintModuleType qid -> print_modtype qid
- | PrintNamespace ns -> print_namespace ns
+ | PrintNamespace ns -> print_namespace ~pstate ns
| PrintMLLoadPath -> Mltop.print_ml_path ()
| PrintMLModules -> Mltop.print_ml_modules ()
| PrintDebugGC -> Mltop.print_gc ()
@@ -1862,7 +1905,13 @@ let vernac_print ~atts env sigma =
| PrintCanonicalConversions -> Prettyp.print_canonical_projections env sigma
| PrintUniverses (sort, subgraph, dst) -> print_universes ~sort ~subgraph dst
| PrintHint r -> Hints.pr_hint_ref env sigma (smart_global r)
- | PrintHintGoal -> Hints.pr_applicable_hint ()
+ | PrintHintGoal ->
+ begin match pstate with
+ | Some pstate ->
+ Hints.pr_applicable_hint pstate
+ | None ->
+ str "No proof in progress"
+ end
| PrintHintDbName s -> Hints.pr_hint_db_by_name env sigma s
| PrintHintDb -> Hints.pr_searchtable env sigma
| PrintScopes ->
@@ -1872,7 +1921,7 @@ let vernac_print ~atts env sigma =
| PrintVisibility s ->
Notation.pr_visibility (Constrextern.without_symbols (pr_lglob_constr_env env)) s
| PrintAbout (ref_or_by_not,udecl,glnumopt) ->
- print_about_hyp_globs ref_or_by_not udecl glnumopt
+ print_about_hyp_globs ~pstate ref_or_by_not udecl glnumopt
| PrintImplicit qid ->
dump_global qid;
print_impargs qid
@@ -1937,16 +1986,16 @@ let () =
optread = (fun () -> !search_output_name_only);
optwrite = (:=) search_output_name_only }
-let vernac_search ~atts s gopt r =
+let vernac_search ~pstate ~atts s gopt r =
let gopt = query_command_selector gopt in
let r = interp_search_restriction r in
let env,gopt =
match gopt with | None ->
(* 1st goal by default if it exists, otherwise no goal at all *)
- (try snd (Pfedit.get_goal_context 1) , Some 1
+ (try snd (get_goal_or_global_context ~pstate 1) , Some 1
with _ -> Global.env (),None)
(* if goal selector is given and wrong, then let exceptions be raised. *)
- | Some g -> snd (Pfedit.get_goal_context g) , Some g
+ | Some g -> snd (get_goal_or_global_context ~pstate g) , Some g
in
let get_pattern c = snd (intern_constr_pattern env Evd.(from_env env) c) in
let pr_search ref env c =
@@ -1961,21 +2010,21 @@ let vernac_search ~atts s gopt r =
in
match s with
| SearchPattern c ->
- (Search.search_pattern gopt (get_pattern c) r |> Search.prioritize_search) pr_search
+ (Search.search_pattern ?pstate gopt (get_pattern c) r |> Search.prioritize_search) pr_search
| SearchRewrite c ->
- (Search.search_rewrite gopt (get_pattern c) r |> Search.prioritize_search) pr_search
+ (Search.search_rewrite ?pstate gopt (get_pattern c) r |> Search.prioritize_search) pr_search
| SearchHead c ->
- (Search.search_by_head gopt (get_pattern c) r |> Search.prioritize_search) pr_search
+ (Search.search_by_head ?pstate 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 Evd.(from_env env))) sl) r |>
+ (Search.search_about ?pstate 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
+let vernac_locate ~pstate = function
| LocateAny {v=AN qid} -> print_located_qualid qid
| LocateTerm {v=AN qid} -> print_located_term qid
| LocateAny {v=ByNotation (ntn, sc)} (* TODO : handle Ltac notations *)
| LocateTerm {v=ByNotation (ntn, sc)} ->
- let _, env = Pfedit.get_current_context () in
+ let _, env = get_current_or_global_context ~pstate in
Notation.locate_notation
(Constrextern.without_symbols (pr_lglob_constr_env env)) ntn sc
| LocateLibrary qid -> print_located_library qid
@@ -1983,9 +2032,9 @@ let vernac_locate = function
| LocateOther (s, qid) -> print_located_other s qid
| LocateFile f -> locate_file f
-let vernac_register qid r =
+let vernac_register ~pstate qid r =
let gr = Smartlocate.global_with_alias qid in
- if Proof_global.there_are_pending_proofs () then
+ if there_are_pending_proofs ~pstate then
user_err Pp.(str "Cannot register a primitive while in proof editing mode.");
match r with
| RegisterInline ->
@@ -2029,8 +2078,8 @@ let vernac_unfocus () =
(fun _ p -> Proof.unfocus command_focus p ())
(* Checks that a proof is fully unfocused. Raises an error if not. *)
-let vernac_unfocused () =
- let p = Proof_global.give_me_the_proof () in
+let vernac_unfocused ~pstate =
+ let p = Proof_global.give_me_the_proof pstate in
if Proof.unfocused p then
str"The proof is indeed fully unfocused."
else
@@ -2060,25 +2109,39 @@ let vernac_bullet (bullet : Proof_bullet.t) =
Proof_global.simple_with_current_proof (fun _ p ->
Proof_bullet.put p bullet)
-let vernac_show = function
- | ShowScript -> assert false (* Only the stm knows the script *)
- | ShowGoal goalref ->
- let proof = Proof_global.give_me_the_proof () in
- begin match goalref with
- | OpenSubgoals -> pr_open_subgoals ~proof
- | NthGoal n -> pr_nth_open_subgoal ~proof n
- | GoalId id -> pr_goal_by_id ~proof id
+let vernac_show ~pstate =
+ match pstate with
+ (* Show functions that don't require a proof state *)
+ | None ->
+ begin function
+ | ShowProof -> show_proof ~pstate
+ | ShowMatch id -> show_match id
+ | ShowScript -> assert false (* Only the stm knows the script *)
+ | _ ->
+ user_err (str "This command requires an open proof.")
end
- | ShowProof -> show_proof ()
- | ShowExistentials -> show_top_evars ()
- | ShowUniverses -> show_universes ()
- | ShowProofNames ->
- pr_sequence Id.print (Proof_global.get_all_proof_names())
- | ShowIntros all -> show_intro all
- | ShowMatch id -> show_match id
-
-let vernac_check_guard () =
- let pts = Proof_global.give_me_the_proof () in
+ (* Show functions that require a proof state *)
+ | Some pstate ->
+ begin function
+ | ShowGoal goalref ->
+ let proof = Proof_global.give_me_the_proof pstate in
+ begin match goalref with
+ | OpenSubgoals -> pr_open_subgoals ~proof
+ | NthGoal n -> pr_nth_open_subgoal ~proof n
+ | GoalId id -> pr_goal_by_id ~proof id
+ end
+ | ShowExistentials -> show_top_evars ~pstate
+ | ShowUniverses -> show_universes ~pstate
+ | ShowProofNames ->
+ pr_sequence Id.print (Proof_global.get_all_proof_names pstate)
+ | ShowIntros all -> show_intro ~pstate all
+ | ShowProof -> show_proof ~pstate:(Some pstate)
+ | ShowMatch id -> show_match id
+ | ShowScript -> assert false (* Only the stm knows the script *)
+ end
+
+let vernac_check_guard ~pstate =
+ let pts = Proof_global.give_me_the_proof pstate in
let pfterm = List.hd (Proof.partial_proof pts) in
let message =
try
@@ -2097,8 +2160,9 @@ exception End_of_input
the way the proof mode is set there makes the task non trivial
without a considerable amount of refactoring.
*)
-let vernac_load interp fname =
- if Proof_global.there_are_pending_proofs () then
+let vernac_load ~st interp fname =
+ let pstate = st.Vernacstate.proof in
+ if there_are_pending_proofs ~pstate then
CErrors.user_err Pp.(str "Load is not supported inside proofs.");
let parse_sentence proof_mode = Flags.with_option Flags.we_are_parsing
(fun po ->
@@ -2112,21 +2176,21 @@ let vernac_load interp fname =
let longfname = Loadpath.locate_file fname in
let in_chan = open_utf8_file_in longfname in
Pcoq.Parsable.make ~file:(Loc.InFile longfname) (Stream.of_channel in_chan) in
- begin
- try while true do
- let proof_mode =
- if Proof_global.there_are_pending_proofs () then
- Some (get_default_proof_mode ())
- else
- None
- in
- interp (parse_sentence proof_mode input).CAst.v;
- done
- with End_of_input -> ()
- end;
+ let rec load_loop ~pstate =
+ try
+ let proof_mode = Option.map (fun _ -> get_default_proof_mode ()) pstate in
+ let pstate = interp ~st:{ st with Vernacstate.proof = pstate }
+ (parse_sentence proof_mode input).CAst.v in
+ load_loop ~pstate
+ with
+ End_of_input ->
+ pstate
+ in
+ let pstate = load_loop ~pstate in
(* 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.")
+ if there_are_pending_proofs ~pstate then
+ CErrors.user_err Pp.(str "Files processed by Load cannot leave open proofs.");
+ pstate
let with_locality ~atts f =
let local = Attributes.(parse locality atts) in
@@ -2151,7 +2215,8 @@ let with_def_attributes ~atts f =
* is the outdated/deprecated "Local" attribute of some vernacular commands
* still parsed as the obsolete_locality grammar entry for retrocompatibility.
* loc is the Loc.t of the vernacular command being interpreted. *)
-let interp ?proof ~atts ~st c =
+let interp ?proof ~atts ~st c : Proof_global.t option =
+ let pstate = st.Vernacstate.proof in
vernac_pperr_endline (fun () -> str "interpreting: " ++ Ppvernac.pr_vernac_expr c);
match c with
@@ -2175,145 +2240,309 @@ let interp ?proof ~atts ~st c =
(* Syntax *)
| VernacSyntaxExtension (infix, sl) ->
- with_module_locality ~atts vernac_syntax_extension infix sl
- | VernacDeclareScope sc -> with_module_locality ~atts vernac_declare_scope sc
- | VernacDelimiters (sc,lr) -> with_module_locality ~atts vernac_delimiters sc lr
- | VernacBindScope (sc,rl) -> with_module_locality ~atts vernac_bind_scope sc rl
- | VernacOpenCloseScope (b, s) -> with_section_locality ~atts vernac_open_close_scope (b,s)
- | VernacInfix (mv,qid,sc) -> with_module_locality ~atts vernac_infix mv qid sc
- | VernacNotation (c,infpl,sc) -> with_module_locality ~atts vernac_notation c infpl sc
+ with_module_locality ~atts vernac_syntax_extension infix sl;
+ pstate
+ | VernacDeclareScope sc ->
+ with_module_locality ~atts vernac_declare_scope sc;
+ pstate
+ | VernacDelimiters (sc,lr) ->
+ with_module_locality ~atts vernac_delimiters sc lr;
+ pstate
+ | VernacBindScope (sc,rl) ->
+ with_module_locality ~atts vernac_bind_scope sc rl;
+ pstate
+ | VernacOpenCloseScope (b, s) ->
+ with_section_locality ~atts vernac_open_close_scope (b,s);
+ pstate
+ | VernacInfix (mv,qid,sc) ->
+ with_module_locality ~atts vernac_infix mv qid sc;
+ pstate
+ | VernacNotation (c,infpl,sc) ->
+ with_module_locality ~atts vernac_notation c infpl sc;
+ pstate
| VernacNotationAddFormat(n,k,v) ->
unsupported_attributes atts;
- Metasyntax.add_notation_extra_printing_rule n k v
+ Metasyntax.add_notation_extra_printing_rule n k v;
+ pstate
| VernacDeclareCustomEntry s ->
- with_module_locality ~atts vernac_custom_entry s
+ with_module_locality ~atts vernac_custom_entry s;
+ pstate
(* Gallina *)
| VernacDefinition ((discharge,kind),lid,d) ->
- with_def_attributes ~atts vernac_definition discharge kind lid d
- | VernacStartTheoremProof (k,l) -> with_def_attributes vernac_start_proof ~atts k l
- | VernacEndProof e -> unsupported_attributes atts; vernac_end_proof ?proof e
- | VernacExactProof c -> unsupported_attributes atts; vernac_exact_proof c
+ with_def_attributes ~atts vernac_definition ~pstate discharge kind lid d
+ | VernacStartTheoremProof (k,l) ->
+ with_def_attributes ~atts vernac_start_proof ~pstate k l
+ | VernacEndProof e ->
+ unsupported_attributes atts;
+ vernac_end_proof ?proof ?pstate e
+ | VernacExactProof c ->
+ unsupported_attributes atts;
+ vernac_require_open_proof ~pstate (vernac_exact_proof c)
| VernacAssumption ((discharge,kind),nl,l) ->
- with_def_attributes vernac_assumption ~atts discharge kind l nl
- | VernacInductive (cum, priv, finite, l) -> vernac_inductive ~atts cum priv finite l
- | VernacFixpoint (discharge, l) -> with_def_attributes vernac_fixpoint ~atts discharge l
- | VernacCoFixpoint (discharge, l) -> with_def_attributes vernac_cofixpoint ~atts discharge l
- | VernacScheme l -> unsupported_attributes atts; vernac_scheme l
- | VernacCombinedScheme (id, l) -> unsupported_attributes atts; vernac_combined_scheme id l
- | VernacUniverse l -> vernac_universe ~poly:(only_polymorphism atts) l
- | VernacConstraint l -> vernac_constraint ~poly:(only_polymorphism atts) l
+ with_def_attributes ~atts vernac_assumption ~pstate discharge kind l nl;
+ pstate
+ | VernacInductive (cum, priv, finite, l) ->
+ vernac_inductive ~atts cum priv finite l;
+ pstate
+ | VernacFixpoint (discharge, l) ->
+ with_def_attributes ~atts vernac_fixpoint ~pstate discharge l
+ | VernacCoFixpoint (discharge, l) ->
+ with_def_attributes ~atts vernac_cofixpoint ~pstate discharge l
+ | VernacScheme l ->
+ unsupported_attributes atts;
+ vernac_scheme l;
+ pstate
+ | VernacCombinedScheme (id, l) ->
+ unsupported_attributes atts;
+ vernac_combined_scheme id l;
+ pstate
+ | VernacUniverse l ->
+ vernac_universe ~poly:(only_polymorphism atts) l;
+ pstate
+ | VernacConstraint l ->
+ vernac_constraint ~poly:(only_polymorphism atts) l;
+ pstate
(* Modules *)
| VernacDeclareModule (export,lid,bl,mtyo) ->
- unsupported_attributes atts; vernac_declare_module export lid bl mtyo
+ unsupported_attributes atts;
+ vernac_declare_module export lid bl mtyo;
+ pstate
| VernacDefineModule (export,lid,bl,mtys,mexprl) ->
- unsupported_attributes atts; vernac_define_module export lid bl mtys mexprl
+ unsupported_attributes atts;
+ vernac_define_module ~pstate export lid bl mtys mexprl;
+ pstate
| VernacDeclareModuleType (lid,bl,mtys,mtyo) ->
- unsupported_attributes atts; vernac_declare_module_type lid bl mtys mtyo
+ unsupported_attributes atts;
+ vernac_declare_module_type ~pstate lid bl mtys mtyo;
+ pstate
| VernacInclude in_asts ->
- unsupported_attributes atts; vernac_include in_asts
+ unsupported_attributes atts;
+ vernac_include in_asts;
+ pstate
(* Gallina extensions *)
- | VernacBeginSection lid -> unsupported_attributes atts; vernac_begin_section lid
+ | VernacBeginSection lid ->
+ unsupported_attributes atts;
+ vernac_begin_section ~pstate lid;
+ pstate
- | VernacEndSegment lid -> unsupported_attributes atts; vernac_end_segment lid
+ | VernacEndSegment lid ->
+ unsupported_attributes atts;
+ vernac_end_segment ~pstate lid;
+ pstate
- | VernacNameSectionHypSet (lid, set) -> unsupported_attributes atts; vernac_name_sec_hyp lid set
+ | VernacNameSectionHypSet (lid, set) ->
+ unsupported_attributes atts;
+ vernac_name_sec_hyp lid set;
+ pstate
- | VernacRequire (from, export, qidl) -> unsupported_attributes atts; vernac_require from export qidl
- | VernacImport (export,qidl) -> unsupported_attributes atts; vernac_import export qidl
- | VernacCanonical qid -> unsupported_attributes atts; vernac_canonical qid
- | VernacCoercion (r,s,t) -> vernac_coercion ~atts r s t
+ | VernacRequire (from, export, qidl) ->
+ unsupported_attributes atts;
+ vernac_require from export qidl;
+ pstate
+ | VernacImport (export,qidl) ->
+ unsupported_attributes atts;
+ vernac_import export qidl;
+ pstate
+ | VernacCanonical qid ->
+ unsupported_attributes atts;
+ vernac_canonical qid;
+ pstate
+ | VernacCoercion (r,s,t) ->
+ vernac_coercion ~atts r s t;
+ pstate
| VernacIdentityCoercion ({v=id},s,t) ->
- vernac_identity_coercion ~atts id s t
+ vernac_identity_coercion ~atts id s t;
+ pstate
(* Type classes *)
| VernacInstance (sup, inst, props, info) ->
- with_def_attributes vernac_instance ~atts sup inst props info
+ snd @@ with_def_attributes ~atts (vernac_instance ~pstate sup inst props info)
| VernacDeclareInstance (sup, inst, info) ->
- with_def_attributes vernac_declare_instance ~atts sup inst info
- | VernacContext sup -> vernac_context ~poly:(only_polymorphism atts) sup
- | VernacExistingInstance insts -> with_section_locality ~atts vernac_existing_instance insts
- | VernacExistingClass id -> unsupported_attributes atts; vernac_existing_class id
+ with_def_attributes ~atts vernac_declare_instance sup inst info;
+ pstate
+ | VernacContext sup ->
+ let () = vernac_context ~pstate ~poly:(only_polymorphism atts) sup in
+ pstate
+ | VernacExistingInstance insts ->
+ with_section_locality ~atts vernac_existing_instance insts;
+ pstate
+ | VernacExistingClass id ->
+ unsupported_attributes atts;
+ vernac_existing_class id;
+ pstate
(* Solving *)
- | VernacSolveExistential (n,c) -> unsupported_attributes atts; vernac_solve_existential n c
+ | VernacSolveExistential (n,c) ->
+ unsupported_attributes atts;
+ Some (vernac_require_open_proof ~pstate (vernac_solve_existential n c))
(* Auxiliary file and library management *)
- | VernacAddLoadPath (isrec,s,alias) -> unsupported_attributes atts; vernac_add_loadpath isrec s alias
- | VernacRemoveLoadPath s -> unsupported_attributes atts; vernac_remove_loadpath s
- | VernacAddMLPath (isrec,s) -> unsupported_attributes atts; vernac_add_ml_path isrec s
- | VernacDeclareMLModule l -> with_locality ~atts vernac_declare_ml_module l
- | VernacChdir s -> unsupported_attributes atts; vernac_chdir s
+ | VernacAddLoadPath (isrec,s,alias) ->
+ unsupported_attributes atts;
+ vernac_add_loadpath isrec s alias;
+ pstate
+ | VernacRemoveLoadPath s ->
+ unsupported_attributes atts;
+ vernac_remove_loadpath s;
+ pstate
+ | VernacAddMLPath (isrec,s) ->
+ unsupported_attributes atts;
+ vernac_add_ml_path isrec s;
+ pstate
+ | VernacDeclareMLModule l ->
+ with_locality ~atts vernac_declare_ml_module l;
+ pstate
+ | VernacChdir s ->
+ unsupported_attributes atts;
+ vernac_chdir s;
+ pstate
(* State management *)
- | VernacWriteState s -> unsupported_attributes atts; vernac_write_state s
- | VernacRestoreState s -> unsupported_attributes atts; vernac_restore_state s
+ | VernacWriteState s ->
+ unsupported_attributes atts;
+ vernac_write_state s;
+ pstate
+ | VernacRestoreState s ->
+ unsupported_attributes atts;
+ vernac_restore_state s;
+ pstate
(* Commands *)
| VernacCreateHintDb (dbname,b) ->
- with_module_locality ~atts vernac_create_hintdb dbname b
+ with_module_locality ~atts vernac_create_hintdb dbname b;
+ pstate
| VernacRemoveHints (dbnames,ids) ->
- with_module_locality ~atts vernac_remove_hints dbnames ids
+ with_module_locality ~atts vernac_remove_hints dbnames ids;
+ pstate
| VernacHints (dbnames,hints) ->
- vernac_hints ~atts dbnames hints
+ vernac_hints ~atts dbnames hints;
+ pstate
| VernacSyntacticDefinition (id,c,b) ->
- with_module_locality ~atts vernac_syntactic_definition id c b
+ with_module_locality ~atts vernac_syntactic_definition id c b;
+ pstate
| VernacArguments (qid, args, more_implicits, nargs, flags) ->
- with_section_locality ~atts vernac_arguments qid args more_implicits nargs flags
- | VernacReserve bl -> unsupported_attributes atts; vernac_reserve bl
- | VernacGeneralizable gen -> with_locality ~atts vernac_generalizable gen
- | VernacSetOpacity qidl -> with_locality ~atts vernac_set_opacity qidl
- | VernacSetStrategy l -> with_locality ~atts vernac_set_strategy l
- | VernacSetOption (export, key,v) -> vernac_set_option ~local:(only_locality atts) export key v
- | VernacUnsetOption (export, key) -> vernac_unset_option ~local:(only_locality atts) export key
- | VernacRemoveOption (key,v) -> unsupported_attributes atts; vernac_remove_option key v
- | VernacAddOption (key,v) -> unsupported_attributes atts; vernac_add_option key v
- | VernacMemOption (key,v) -> unsupported_attributes atts; vernac_mem_option key v
- | VernacPrintOption key -> unsupported_attributes atts; vernac_print_option key
+ with_section_locality ~atts vernac_arguments qid args more_implicits nargs flags;
+ pstate
+ | VernacReserve bl ->
+ unsupported_attributes atts;
+ vernac_reserve bl;
+ pstate
+ | VernacGeneralizable gen ->
+ with_locality ~atts vernac_generalizable gen;
+ pstate
+ | VernacSetOpacity qidl ->
+ with_locality ~atts vernac_set_opacity qidl;
+ pstate
+ | VernacSetStrategy l ->
+ with_locality ~atts vernac_set_strategy l;
+ pstate
+ | VernacSetOption (export, key,v) ->
+ vernac_set_option ~local:(only_locality atts) export key v;
+ pstate
+ | VernacUnsetOption (export, key) ->
+ vernac_unset_option ~local:(only_locality atts) export key;
+ pstate
+ | VernacRemoveOption (key,v) ->
+ unsupported_attributes atts;
+ vernac_remove_option key v;
+ pstate
+ | VernacAddOption (key,v) ->
+ unsupported_attributes atts;
+ vernac_add_option key v;
+ pstate
+ | VernacMemOption (key,v) ->
+ unsupported_attributes atts;
+ vernac_mem_option key v;
+ pstate
+ | VernacPrintOption key ->
+ unsupported_attributes atts;
+ vernac_print_option key;
+ pstate
| VernacCheckMayEval (r,g,c) ->
- Feedback.msg_notice @@ vernac_check_may_eval ~atts r g c
- | VernacDeclareReduction (s,r) -> with_locality ~atts vernac_declare_reduction s r
+ Feedback.msg_notice @@
+ vernac_check_may_eval ~pstate ~atts r g c;
+ pstate
+ | VernacDeclareReduction (s,r) ->
+ with_locality ~atts vernac_declare_reduction s r;
+ pstate
| VernacGlobalCheck c ->
unsupported_attributes atts;
- Feedback.msg_notice @@ vernac_global_check c
+ Feedback.msg_notice @@ vernac_global_check c;
+ pstate
| VernacPrint p ->
- let sigma, env = Pfedit.get_current_context () in
- Feedback.msg_notice @@ vernac_print ~atts env sigma p
- | VernacSearch (s,g,r) -> unsupported_attributes atts; vernac_search ~atts s g r
+ Feedback.msg_notice @@ vernac_print ~pstate ~atts p;
+ pstate
+ | VernacSearch (s,g,r) ->
+ unsupported_attributes atts;
+ vernac_search ~pstate ~atts s g r;
+ pstate
| VernacLocate l -> unsupported_attributes atts;
- Feedback.msg_notice @@ vernac_locate l
- | VernacRegister (qid, r) -> unsupported_attributes atts; vernac_register qid r
- | VernacPrimitive (id, prim, typopt) -> unsupported_attributes atts; ComAssumption.do_primitive id prim typopt
- | VernacComments l -> unsupported_attributes atts;
- Flags.if_verbose Feedback.msg_info (str "Comments ok\n")
+ Feedback.msg_notice @@ vernac_locate ~pstate l;
+ pstate
+ | VernacRegister (qid, r) ->
+ unsupported_attributes atts;
+ vernac_register ~pstate qid r;
+ pstate
+ | VernacPrimitive (id, prim, typopt) ->
+ unsupported_attributes atts;
+ ComAssumption.do_primitive id prim typopt;
+ pstate
+ | VernacComments l ->
+ unsupported_attributes atts;
+ Flags.if_verbose Feedback.msg_info (str "Comments ok\n");
+ pstate
(* Proof management *)
- | VernacFocus n -> unsupported_attributes atts; vernac_focus n
- | VernacUnfocus -> unsupported_attributes atts; vernac_unfocus ()
- | VernacUnfocused -> unsupported_attributes atts;
- Feedback.msg_notice @@ vernac_unfocused ()
- | VernacBullet b -> unsupported_attributes atts; vernac_bullet b
- | VernacSubproof n -> unsupported_attributes atts; vernac_subproof n
- | VernacEndSubproof -> unsupported_attributes atts; vernac_end_subproof ()
- | VernacShow s -> unsupported_attributes atts;
- Feedback.msg_notice @@ vernac_show s
- | VernacCheckGuard -> unsupported_attributes atts;
- Feedback.msg_notice @@ vernac_check_guard ()
- | VernacProof (tac, using) -> unsupported_attributes atts;
+ | VernacFocus n ->
+ unsupported_attributes atts;
+ Option.map (vernac_focus n) pstate
+ | VernacUnfocus ->
+ unsupported_attributes atts;
+ Option.map (vernac_unfocus ()) pstate
+ | VernacUnfocused ->
+ unsupported_attributes atts;
+ Option.iter (fun pstate -> Feedback.msg_notice @@ vernac_unfocused ~pstate) pstate;
+ pstate
+ | VernacBullet b ->
+ unsupported_attributes atts;
+ Option.map (vernac_bullet b) pstate
+ | VernacSubproof n ->
+ unsupported_attributes atts;
+ Option.map (vernac_subproof n) pstate
+ | VernacEndSubproof ->
+ unsupported_attributes atts;
+ Option.map (vernac_end_subproof ()) pstate
+ | VernacShow s ->
+ unsupported_attributes atts;
+ Feedback.msg_notice @@ vernac_show ~pstate s;
+ pstate
+ | VernacCheckGuard ->
+ unsupported_attributes atts;
+ Feedback.msg_notice @@
+ vernac_require_open_proof ~pstate (vernac_check_guard);
+ pstate
+ | VernacProof (tac, using) ->
+ unsupported_attributes atts;
let using = Option.append using (Proof_using.get_default_proof_using ()) in
let tacs = if Option.is_empty tac then "tac:no" else "tac:yes" in
let usings = if Option.is_empty using then "using:no" else "using:yes" in
Aux_file.record_in_aux_at "VernacProof" (tacs^" "^usings);
- Option.iter vernac_set_end_tac tac;
- Option.iter vernac_set_used_variables using
- | VernacProofMode mn -> unsupported_attributes atts; ()
+ let pstate =
+ vernac_require_open_proof ~pstate (fun ~pstate ->
+ let pstate = Option.cata (vernac_set_end_tac ~pstate) pstate tac in
+ Option.cata (vernac_set_used_variables ~pstate) pstate using)
+ in Some pstate
+ | VernacProofMode mn ->
+ unsupported_attributes atts;
+ pstate
(* Extensions *)
| VernacExtend (opn,args) ->
(* XXX: Here we are returning the state! :) *)
- let _st : Vernacstate.t = Vernacextend.call ~atts opn args ~st in
- ()
+ let st : Vernacstate.t = Vernacextend.call ~atts opn args ~st in
+ st.Vernacstate.proof
(** A global default timeout, controlled by option "Set Default Timeout n".
Use "Unset Default Timeout" to deactivate it (or set it to 0). *)
@@ -2333,12 +2562,18 @@ let () =
let current_timeout = ref None
-let vernac_timeout f =
+let vernac_timeout (f : 'a -> 'b) (x : 'a) : 'b =
match !current_timeout, !default_timeout with
- | Some n, _ | None, Some n ->
- let f () = f (); current_timeout := None in
- Control.timeout n f () Timeout
- | None, None -> f ()
+ | Some n, _
+ | None, Some n ->
+ let f v =
+ let res = f v in
+ current_timeout := None;
+ res
+ in
+ Control.timeout n f x Timeout
+ | None, None ->
+ f x
let restore_timeout () = current_timeout := None
@@ -2354,84 +2589,87 @@ let test_mode = ref false
(* XXX STATE: this type hints that restoring the state should be the
caller's responsibility *)
-let with_fail st b f =
- if not b
- then f ()
- else begin try
- (* If the command actually works, ignore its effects on the state.
+let with_fail ~st f =
+ try
+ (* If the command actually works, ignore its effects on the state.
* Note that error has to be printed in the right state, hence
* within the purified function *)
- try f (); raise HasNotFailed
- with
- | HasNotFailed as e -> raise e
- | e ->
- let e = CErrors.push e in
- raise (HasFailed (CErrors.iprint
- (ExplainErr.process_vernac_interp_error ~allow_uncaught:false e)))
- with e when CErrors.noncritical e ->
- (* Restore the previous state XXX Careful here with the cache! *)
- Vernacstate.invalidate_cache ();
- Vernacstate.unfreeze_interp_state st;
- let (e, _) = CErrors.push e in
- match e with
- | HasNotFailed ->
- user_err ~hdr:"Fail" (str "The command has not failed!")
- | HasFailed msg ->
- if not !Flags.quiet || !test_mode then Feedback.msg_info
- (str "The command has indeed failed with message:" ++ fnl () ++ msg)
- | _ -> assert false
- end
-
-let interp ?(verbosely=true) ?proof ~st {CAst.loc;v=c} =
- let rec control = function
+ try let _ = f () in raise HasNotFailed
+ with
+ | HasNotFailed as e -> raise e
+ | e ->
+ let e = CErrors.push e in
+ raise (HasFailed (CErrors.iprint
+ (ExplainErr.process_vernac_interp_error ~allow_uncaught:false e)))
+ with e when CErrors.noncritical e ->
+ (* Restore the previous state XXX Careful here with the cache! *)
+ Vernacstate.invalidate_cache ();
+ Vernacstate.unfreeze_interp_state st;
+ let (e, _) = CErrors.push e in
+ match e with
+ | HasNotFailed ->
+ user_err ~hdr:"Fail" (str "The command has not failed!")
+ | HasFailed msg ->
+ if not !Flags.quiet || !test_mode then Feedback.msg_info
+ (str "The command has indeed failed with message:" ++ fnl () ++ msg)
+ | _ -> assert false
+
+let interp ?(verbosely=true) ?proof ~st {CAst.loc;v=c} : Proof_global.t option =
+ let rec control ~st = function
| VernacExpr (atts, v) ->
- aux ~atts v
- | VernacFail v -> with_fail st true (fun () -> control v)
+ aux ~atts ~st v
+ | VernacFail v ->
+ with_fail ~st (fun () -> ignore(control ~st v));
+ st.Vernacstate.proof
| VernacTimeout (n,v) ->
current_timeout := Some n;
- control v
+ control ~st v
| VernacRedirect (s, {v}) ->
- Topfmt.with_output_to_file s control v
- | VernacTime (batch, com) ->
+ Topfmt.with_output_to_file s (control ~st) v
+ | VernacTime (batch, ({v} as com)) ->
let header = if batch then Topfmt.pr_cmd_header com else Pp.mt () in
- System.with_time ~batch ~header control com.CAst.v;
+ System.with_time ~batch ~header (control ~st) v;
- and aux ~atts : _ -> unit =
+ and aux ~atts ~st : _ -> Proof_global.t option =
function
| VernacLoad (_,fname) ->
unsupported_attributes atts;
- vernac_load control fname
+ vernac_load ~st control fname
| c ->
(* NB: we keep polymorphism and program in the attributes, we're
just parsing them to do our option magic. *)
try
- vernac_timeout begin fun () ->
- if verbosely
- then Flags.verbosely (interp ?proof ~atts ~st) c
- else Flags.silently (interp ?proof ~atts ~st) c;
- end
- with
- | reraise when
- (match reraise with
- | Timeout -> true
- | e -> CErrors.noncritical e)
- ->
- let e = CErrors.push reraise in
- let e = locate_if_not_already ?loc e in
- let () = restore_timeout () in
- iraise e
+ vernac_timeout begin fun st ->
+ let pstate : Proof_global.t option =
+ if verbosely
+ then Flags.verbosely (interp ?proof ~atts ~st) c
+ else Flags.silently (interp ?proof ~atts ~st) c
+ in
+ pstate
+ end st
+ with
+ | reraise when
+ (match reraise with
+ | Timeout -> true
+ | e -> CErrors.noncritical e)
+ ->
+ let e = CErrors.push reraise in
+ let e = locate_if_not_already ?loc e in
+ let () = restore_timeout () in
+ iraise e
in
if verbosely
- then Flags.verbosely control c
- else control c
+ then Flags.verbosely (control ~st) c
+ else (control ~st) c
(* Be careful with the cache here in case of an exception. *)
let interp ?verbosely ?proof ~st cmd =
Vernacstate.unfreeze_interp_state st;
try
- interp ?verbosely ?proof ~st cmd;
+ let pstate = interp ?verbosely ?proof ~st cmd in
+ Vernacstate.Proof_global.set pstate;
Vernacstate.freeze_interp_state ~marshallable:false
with exn ->
let exn = CErrors.push exn in
diff --git a/vernac/vernacentries.mli b/vernac/vernacentries.mli
index f43cec48e9..71cc29b6e1 100644
--- a/vernac/vernacentries.mli
+++ b/vernac/vernacentries.mli
@@ -33,15 +33,17 @@ val interp :
val make_cases : string -> string list list
-(* XXX STATE: this type hints that restoring the state should be the
- caller's responsibility *)
-val with_fail : Vernacstate.t -> bool -> (unit -> unit) -> unit
+(** [with_fail ~st f] runs [f ()] and expects it to fail, otherwise it fails. *)
+val with_fail : st:Vernacstate.t -> (unit -> 'a) -> unit
val command_focus : unit Proof.focus_kind
val interp_redexp_hook : (Environ.env -> Evd.evar_map -> Genredexpr.raw_red_expr ->
Evd.evar_map * Redexpr.red_expr) Hook.t
+(** Helper *)
+val vernac_require_open_proof : pstate:Proof_global.t option -> (pstate:Proof_global.t -> 'a) -> 'a
+
(* Flag set when the test-suite is called. Its only effect to display
verbose information for `Fail` *)
val test_mode : bool ref
diff --git a/vernac/vernacstate.ml b/vernac/vernacstate.ml
index c691dc8559..77f54361da 100644
--- a/vernac/vernacstate.ml
+++ b/vernac/vernacstate.ml
@@ -28,10 +28,10 @@ module Parser = struct
end
type t = {
- parsing: Parser.state;
- system : States.state; (* summary + libstack *)
- proof : Proof_global.t; (* proof state *)
- shallow : bool; (* is the state trimmed down (libstack) *)
+ parsing : Parser.state;
+ system : States.state; (* summary + libstack *)
+ proof : Proof_global.t option; (* proof state *)
+ shallow : bool (* is the state trimmed down (libstack) *)
}
let s_cache = ref None
@@ -55,14 +55,14 @@ let do_if_not_cached rf f v =
let freeze_interp_state ~marshallable =
{ system = update_cache s_cache (States.freeze ~marshallable);
- proof = update_cache s_proof (Proof_global.freeze ~marshallable);
+ proof = !s_proof;
shallow = false;
parsing = Parser.cur_state ();
}
let unfreeze_interp_state { system; proof; parsing } =
do_if_not_cached s_cache States.unfreeze system;
- do_if_not_cached s_proof Proof_global.unfreeze proof;
+ s_proof := proof;
Pcoq.unfreeze parsing
let make_shallow st =
@@ -71,3 +71,75 @@ let make_shallow st =
system = States.replace_lib st.system @@ Lib.drop_objects lib;
shallow = true;
}
+
+(* Compatibility module *)
+module Proof_global = struct
+
+ let get () = !s_proof
+ let set x = s_proof := x
+
+ let freeze ~marshallable:_ = get ()
+ let unfreeze x = s_proof := Some x
+
+ exception NoCurrentProof
+
+ let () =
+ CErrors.register_handler begin function
+ | NoCurrentProof ->
+ CErrors.user_err Pp.(str "No focused proof (No proof-editing in progress).")
+ | _ -> raise CErrors.Unhandled
+ end
+
+ open Proof_global
+
+ let cc f = match !s_proof with
+ | None -> raise NoCurrentProof
+ | Some x -> f x
+
+ let dd f = match !s_proof with
+ | None -> raise NoCurrentProof
+ | Some x -> s_proof := Some (f x)
+
+ let there_are_pending_proofs () = !s_proof <> None
+ let get_open_goals () = cc get_open_goals
+
+ let set_terminator x = dd (set_terminator x)
+ let give_me_the_proof_opt () = Option.map give_me_the_proof !s_proof
+ let give_me_the_proof () = cc give_me_the_proof
+ let get_current_proof_name () = cc get_current_proof_name
+
+ let simple_with_current_proof f =
+ dd (simple_with_current_proof f)
+
+ let with_current_proof f =
+ let pf, res = cc (with_current_proof f) in
+ s_proof := Some pf; res
+
+ let install_state s = s_proof := Some s
+
+ let return_proof ?allow_partial () =
+ cc (return_proof ?allow_partial)
+
+ let close_future_proof ~opaque ~feedback_id pf =
+ cc (fun st -> close_future_proof ~opaque ~feedback_id st pf)
+
+ let close_proof ~opaque ~keep_body_ucst_separate f =
+ cc (close_proof ~opaque ~keep_body_ucst_separate f)
+
+ let discard_all () = s_proof := None
+ let update_global_env () = dd update_global_env
+
+ let get_current_context () = cc Pfedit.get_current_context
+
+ let get_all_proof_names () =
+ try cc get_all_proof_names
+ with NoCurrentProof -> []
+
+ let copy_terminators ~src ~tgt =
+ match src, tgt with
+ | None, None -> None
+ | Some _ , None -> None
+ | None, Some x -> Some x
+ | Some src, Some tgt -> Some (copy_terminators ~src ~tgt)
+
+end
diff --git a/vernac/vernacstate.mli b/vernac/vernacstate.mli
index 581c23386a..b79f97796f 100644
--- a/vernac/vernacstate.mli
+++ b/vernac/vernacstate.mli
@@ -19,10 +19,10 @@ module Parser : sig
end
type t = {
- parsing: Parser.state;
- system : States.state; (* summary + libstack *)
- proof : Proof_global.t; (* proof state *)
- shallow : bool; (* is the state trimmed down (libstack) *)
+ parsing : Parser.state;
+ system : States.state; (* summary + libstack *)
+ proof : Proof_global.t option; (* proof state *)
+ shallow : bool (* is the state trimmed down (libstack) *)
}
val freeze_interp_state : marshallable:bool -> t
@@ -32,3 +32,53 @@ val make_shallow : t -> t
(* WARNING: Do not use, it will go away in future releases *)
val invalidate_cache : unit -> unit
+
+(* Compatibility module: Do Not Use *)
+module Proof_global : sig
+
+ open Proof_global
+
+ (* Low-level stuff *)
+ val get : unit -> t option
+ val set : t option -> unit
+
+ val freeze : marshallable:bool -> t option
+ val unfreeze : t -> unit
+
+ exception NoCurrentProof
+
+ val there_are_pending_proofs : unit -> bool
+ val get_open_goals : unit -> int
+
+ val set_terminator : proof_terminator -> unit
+ val give_me_the_proof : unit -> Proof.t
+ val give_me_the_proof_opt : unit -> Proof.t option
+ val get_current_proof_name : unit -> Names.Id.t
+
+ val simple_with_current_proof :
+ (unit Proofview.tactic -> Proof.t -> Proof.t) -> unit
+
+ val with_current_proof :
+ (unit Proofview.tactic -> Proof.t -> Proof.t * 'a) -> 'a
+
+ val install_state : t -> unit
+
+ val return_proof : ?allow_partial:bool -> unit -> closed_proof_output
+
+ val close_future_proof :
+ opaque:opacity_flag ->
+ feedback_id:Stateid.t ->
+ closed_proof_output Future.computation -> closed_proof
+
+ val close_proof : opaque:opacity_flag -> keep_body_ucst_separate:bool -> Future.fix_exn -> closed_proof
+
+ val discard_all : unit -> unit
+ val update_global_env : unit -> unit
+
+ val get_current_context : unit -> Evd.evar_map * Environ.env
+
+ val get_all_proof_names : unit -> Names.Id.t list
+
+ val copy_terminators : src:t option -> tgt:t option -> t option
+
+end