aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authoraspiwack2010-04-22 19:20:00 +0000
committeraspiwack2010-04-22 19:20:00 +0000
commitaa99fc9ed78a0246d11d53dde502773a915b1022 (patch)
treed2ead3a9cf896fff6a49cfef72b6d5a52e928b41
parentf77d428c11bf47c20b8ea67d8ed7dce6af106bcd (diff)
Here comes the commit, announced long ago, of the new tactic engine.
This is a fairly large commit (around 140 files and 7000 lines of code impacted), it will cause some troubles for sure (I've listed the know regressions below, there is bound to be more). At this state of developpement it brings few features to the user, as the old tactics were ported with no change. Changes are on the side of the developer mostly. Here comes a list of the major changes. I will stay brief, but the code is hopefully well documented so that it is reasonably easy to infer the details from it. Feature developer-side: * Primitives for a "real" refine tactic (generating a goal for each evar). * Abstract type of tactics, goals and proofs * Tactics can act on several goals (formally all the focused goals). An interesting consequence of this is that the tactical (. ; [ . | ... ]) can be separated in two tacticals (. ; .) and ( [ . | ... ] ) (although there is a conflict for this particular syntax). We can also imagine a tactic to reorder the goals. * Possibility for a tactic to pass a value to following tactics (a typical example is an intro function which tells the following tactics which name it introduced). * backtracking primitives for tactics (it is now possible to implement a tactical '+' with (a+b);c equivalent to (a;c+b;c) (itself equivalent to (a;c||b;c)). This is a valuable tool to implement tactics like "auto" without nowing of the implementation of tactics. * A notion of proof modes, which allows to dynamically change the parser for tactics. It is controlled at user level with the keywords Set Default Proof Mode (this is the proof mode which is loaded at the start of each proof) and Proof Mode (switches the proof mode of the current proof) to control them. * A new primitive Evd.fold_undefined which operates like an Evd.fold, except it only goes through the evars whose body is Evar_empty. This is a common operation throughout the code, some of the fold-and-test-if-empty occurences have been replaced by fold_undefined. For now, it is only implemented as a fold-and-test, but we expect to have some optimisations coming some day, as there can be a lot of evars in an evar_map with this new implementation (I've observed a couple of thousands), whereas there are rarely more than a dozen undefined ones. Folding being a linear operation, this might result in a significant speed-up. * The declarative mode has been moved into the plugins. This is made possible by the proof mode feature. I tried to document it so that it can serve as a tutorial for a tactic mode plugin. Features user-side: * Unfocus does not go back to the root of the proof if several Focus-s have been performed. It only goes back to the point where it was last focused. * experimental (non-documented) support of keywords BeginSubproof/EndSubproof: BeginSubproof focuses on first goal, one can unfocus only with EndSubproof, and only if the proof is completed for that goal. * experimental (non-documented) support for bullets ('+', '-' and '*') they act as hierarchical BeginSubproof/EndSubproof: First time one uses '+' (for instance) it focuses on first goal, when the subproof is completed, one can use '+' again which unfocuses and focuses on next first goal. Meanwhile, one cas use '*' (for instance) to focus more deeply. Known regressions: * The xml plugin had some functions related to proof trees. As the structure of proof changed significantly, they do not work anymore. * I do not know how to implement info or show script in this new engine. Actually I don't even know what they were suppose to actually mean in earlier versions either. I wager they would require some calm thinking before going back to work. * Declarative mode not entirely working (in particular proofs by induction need to be restored). * A bug in the inversion tactic (observed in some contributions) * A bug in Program (observed in some contributions) * Minor change in the 'old' type of tactics causing some contributions to fail. * Compilation time takes about 10-15% longer for unknown reasons (I suspect it might be linked to the fact that I don't perform any reduction at QED-s, and also to some linear operations on evar_map-s (see Evd.fold_undefined above)). git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@12961 85f007b7-540e-0410-9357-904b9bb8a0f7
-rw-r--r--Makefile.common5
-rw-r--r--dev/base_include6
-rw-r--r--dev/printers.mllib9
-rw-r--r--dev/top_printers.ml28
-rw-r--r--ide/coq.ml41
-rw-r--r--ide/coqide.ml24
-rw-r--r--kernel/cbytegen.ml17
-rw-r--r--kernel/cooking.ml2
-rw-r--r--kernel/term.mli4
-rw-r--r--lib/bstack.ml75
-rw-r--r--lib/bstack.mli22
-rw-r--r--lib/edit.ml134
-rw-r--r--lib/edit.mli63
-rw-r--r--lib/lib.mllib4
-rw-r--r--lib/store.ml63
-rw-r--r--lib/store.mli27
-rw-r--r--parsing/egrammar.ml5
-rw-r--r--parsing/egrammar.mli2
-rw-r--r--parsing/g_decl_mode.ml4252
-rw-r--r--parsing/g_proofs.ml44
-rw-r--r--parsing/g_vernac.ml455
-rw-r--r--parsing/grammar.mllib1
-rw-r--r--parsing/highparsing.mllib1
-rw-r--r--parsing/parsing.mllib1
-rw-r--r--parsing/pcoq.ml42
-rw-r--r--parsing/pcoq.mli1
-rw-r--r--parsing/ppvernac.ml13
-rw-r--r--parsing/printer.ml79
-rw-r--r--parsing/printer.mli8
-rw-r--r--parsing/tactic_printer.ml89
-rw-r--r--parsing/tactic_printer.mli1
-rw-r--r--parsing/vernacextend.ml420
-rw-r--r--plugins/cc/ccalgo.ml10
-rw-r--r--plugins/cc/cctac.ml4
-rw-r--r--plugins/decl_mode/decl_expr.mli (renamed from proofs/decl_expr.mli)0
-rw-r--r--plugins/decl_mode/decl_interp.ml (renamed from tactics/decl_interp.ml)0
-rw-r--r--plugins/decl_mode/decl_interp.mli (renamed from tactics/decl_interp.mli)0
-rw-r--r--plugins/decl_mode/decl_mode.ml (renamed from proofs/decl_mode.ml)101
-rw-r--r--plugins/decl_mode/decl_mode.mli (renamed from proofs/decl_mode.mli)12
-rw-r--r--plugins/decl_mode/decl_mode_plugin.mllib7
-rw-r--r--plugins/decl_mode/decl_proof_instr.ml (renamed from tactics/decl_proof_instr.ml)150
-rw-r--r--plugins/decl_mode/decl_proof_instr.mli (renamed from tactics/decl_proof_instr.mli)10
-rw-r--r--plugins/decl_mode/g_decl_mode.ml4407
-rw-r--r--plugins/decl_mode/ppdecl_proof.ml (renamed from parsing/ppdecl_proof.ml)0
-rw-r--r--plugins/decl_mode/ppdecl_proof.mli (renamed from parsing/ppdecl_proof.mli)0
-rw-r--r--plugins/firstorder/ground.ml28
-rw-r--r--plugins/funind/functional_principles_proofs.ml2
-rw-r--r--plugins/funind/indfun_common.ml45
-rw-r--r--plugins/funind/indfun_common.mli7
-rw-r--r--plugins/funind/invfun.ml2
-rw-r--r--plugins/funind/recdef.ml10
-rw-r--r--plugins/pluginsbyte.itarget1
-rw-r--r--plugins/pluginsopt.itarget1
-rw-r--r--plugins/quote/quote.ml1
-rw-r--r--plugins/ring/ring.ml1
-rw-r--r--plugins/rtauto/refl_tauto.ml5
-rw-r--r--plugins/setoid_ring/newring.ml46
-rw-r--r--plugins/subtac/eterm.ml6
-rw-r--r--plugins/subtac/subtac.ml4
-rw-r--r--plugins/subtac/subtac_obligations.ml1
-rw-r--r--plugins/xml/dumptree.ml428
-rw-r--r--plugins/xml/proof2aproof.ml82
-rw-r--r--plugins/xml/proofTree2Xml.ml49
-rw-r--r--plugins/xml/xmlcommand.ml29
-rw-r--r--pretyping/evarutil.ml22
-rw-r--r--pretyping/evarutil.mli4
-rw-r--r--pretyping/evd.ml59
-rw-r--r--pretyping/evd.mli11
-rw-r--r--pretyping/pretyping.mllib1
-rw-r--r--pretyping/typeclasses.ml15
-rw-r--r--pretyping/typeclasses.mli4
-rw-r--r--proofs/clenv.ml (renamed from pretyping/clenv.ml)26
-rw-r--r--proofs/clenv.mli (renamed from pretyping/clenv.mli)16
-rw-r--r--proofs/clenvtac.ml1
-rw-r--r--proofs/evar_refiner.ml18
-rw-r--r--proofs/evar_refiner.mli2
-rw-r--r--proofs/goal.ml572
-rw-r--r--proofs/goal.mli228
-rw-r--r--proofs/logic.ml287
-rw-r--r--proofs/logic.mli3
-rw-r--r--proofs/pfedit.ml369
-rw-r--r--proofs/pfedit.mli57
-rw-r--r--proofs/proof.ml294
-rw-r--r--proofs/proof.mli133
-rw-r--r--proofs/proof_global.ml295
-rw-r--r--proofs/proof_global.mli89
-rw-r--r--proofs/proof_trees.ml107
-rw-r--r--proofs/proof_trees.mli48
-rw-r--r--proofs/proof_type.ml13
-rw-r--r--proofs/proof_type.mli12
-rw-r--r--proofs/proofs.mllib9
-rw-r--r--proofs/proofview.ml491
-rw-r--r--proofs/proofview.mli203
-rw-r--r--proofs/refiner.ml619
-rw-r--r--proofs/refiner.mli83
-rw-r--r--proofs/tacmach.ml51
-rw-r--r--proofs/tacmach.mli38
-rw-r--r--proofs/tactic_debug.ml12
-rw-r--r--tactics/auto.ml7
-rw-r--r--tactics/autorewrite.ml6
-rw-r--r--tactics/class_tactics.ml4141
-rw-r--r--tactics/eauto.ml420
-rw-r--r--tactics/eqdecide.ml41
-rw-r--r--tactics/evar_tactics.ml4
-rw-r--r--tactics/hipattern.ml41
-rw-r--r--tactics/hipattern.mli1
-rw-r--r--tactics/leminv.ml46
-rw-r--r--tactics/refine.ml2
-rw-r--r--tactics/rewrite.ml418
-rw-r--r--tactics/tacinterp.ml19
-rw-r--r--tactics/tacinterp.mli2
-rw-r--r--tactics/tacticals.ml2
-rw-r--r--tactics/tacticals.mli2
-rw-r--r--tactics/tactics.ml7
-rw-r--r--tactics/tactics.mllib2
-rw-r--r--theories/Classes/RelationClasses.v2
-rw-r--r--theories/Init/Prelude.v2
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v2
-rw-r--r--theories/Numbers/Cyclic/Int31/Cyclic31.v2
-rw-r--r--theories/Program/Tactics.v6
-rw-r--r--theories/Reals/RiemannInt.v2
-rw-r--r--toplevel/auto_ind_decl.ml4
-rw-r--r--toplevel/himsg.ml7
-rw-r--r--toplevel/lemmas.mli2
-rw-r--r--toplevel/vernacentries.ml244
-rw-r--r--toplevel/vernacentries.mli11
-rw-r--r--toplevel/vernacexpr.ml17
-rw-r--r--toplevel/whelp.ml43
128 files changed, 3850 insertions, 2954 deletions
diff --git a/Makefile.common b/Makefile.common
index e3a94f56de..4d5e00d251 100644
--- a/Makefile.common
+++ b/Makefile.common
@@ -73,7 +73,7 @@ SRCDIRS:=\
omega romega micromega quote ring dp \
setoid_ring xml extraction fourier \
cc funind firstorder field subtac \
- rtauto groebner syntax)
+ rtauto groebner syntax decl_mode)
# Order is relevent here because kernel and checker contain files
# with the same name
@@ -188,10 +188,12 @@ OTHERSYNTAXCMA:=$(addprefix plugins/syntax/, \
r_syntax_plugin.cma \
ascii_syntax_plugin.cma \
string_syntax_plugin.cma )
+DECLMODECMA:=plugins/decl_mode/decl_mode_plugin.cma
PLUGINSCMA:=$(OMEGACMA) $(ROMEGACMA) $(MICROMEGACMA) \
$(QUOTECMA) $(RINGCMA) $(NEWRINGCMA) $(DPCMA) $(FIELDCMA) \
$(FOURIERCMA) $(EXTRACTIONCMA) $(XMLCMA) \
+ $(DECLMODECMA) \
$(CCCMA) $(FOCMA) $(SUBTACCMA) $(RTAUTOCMA) \
$(FUNINDCMA) $(GBCMA) $(NATSYNTAXCMA) $(OTHERSYNTAXCMA)
@@ -326,7 +328,6 @@ ALLMODS:=$(call vo_to_mod,$(ALLVO))
LIBFILES:=$(THEORIESVO) $(PLUGINSVO)
LIBFILESLIGHT:=$(THEORIESLIGHTVO)
-
###########################################################################
# Miscellaneous
###########################################################################
diff --git a/dev/base_include b/dev/base_include
index 3a31230f19..debc074de9 100644
--- a/dev/base_include
+++ b/dev/base_include
@@ -110,11 +110,9 @@ open Topconstr
open Prettyp
open Search
-open Clenvtac
open Evar_refiner
open Logic
open Pfedit
-open Proof_trees
open Proof_type
open Redexpr
open Refiner
@@ -193,12 +191,12 @@ let constbody_of_string s =
Option.get b.const_body;;
(* Get the current goal *)
-
+(*
let getgoal x = top_goal_of_pftreestate (Pfedit.get_pftreestate x);;
let get_nth_goal n = nth_goal_of_pftreestate n (Pfedit.get_pftreestate ());;
let current_goal () = get_nth_goal 1;;
-
+*)
let pf_e gl s =
Constrintern.interp_constr (project gl) (pf_env gl) (parse_constr s);;
diff --git a/dev/printers.mllib b/dev/printers.mllib
index e8ec10c5cd..8ada6769cc 100644
--- a/dev/printers.mllib
+++ b/dev/printers.mllib
@@ -12,8 +12,7 @@ Hashcons
Dyn
System
Envars
-Bstack
-Edit
+Store
Gset
Gmap
Fset
@@ -93,7 +92,6 @@ Coercion
Unification
Cases
Pretyping
-Clenv
Lexer
Ppextend
@@ -108,12 +106,15 @@ Syntax_def
Implicit_quantifiers
Smartlocate
Constrintern
-Proof_trees
Tacexpr
Proof_type
+Goal
Logic
Refiner
Evar_refiner
+Proofview
+Proof
+Proof_global
Pfedit
Tactic_debug
Decl_mode
diff --git a/dev/top_printers.ml b/dev/top_printers.ml
index b481e9f169..21a690f953 100644
--- a/dev/top_printers.ml
+++ b/dev/top_printers.ml
@@ -16,14 +16,11 @@ open Libnames
open Nameops
open Sign
open Univ
-open Proof_trees
open Environ
open Printer
open Tactic_printer
-open Refiner
open Term
open Termops
-open Clenv
open Cerrors
open Evd
open Goptions
@@ -103,21 +100,24 @@ let pp_transparent_state s = pp (pr_transparent_state s)
(* proof printers *)
let ppmetas metas = pp(pr_metaset metas)
let ppevm evd = pp(pr_evar_map evd)
+(* spiwack: deactivated until a replacement is found
let ppclenv clenv = pp(pr_clenv clenv)
let ppgoal g = pp(db_pr_goal g)
let pppftreestate p = pp(print_pftreestate p)
+*)
-let pr_gls gls =
- hov 0 (pr_evar_map (sig_sig gls) ++ fnl () ++ db_pr_goal (sig_it gls))
+(* let ppgoal g = pp(db_pr_goal g) *)
+(* let pr_gls gls = *)
+(* hov 0 (pr_evar_defs (sig_sig gls) ++ fnl () ++ db_pr_goal (sig_it gls)) *)
-let pr_glls glls =
- hov 0 (pr_evar_map (sig_sig glls) ++ fnl () ++
- prlist_with_sep pr_fnl db_pr_goal (sig_it glls))
+(* let pr_glls glls = *)
+(* hov 0 (pr_evar_defs (sig_sig glls) ++ fnl () ++ *)
+(* prlist_with_sep pr_fnl db_pr_goal (sig_it glls)) *)
-let ppsigmagoal g = pp(pr_goal (sig_it g))
-let prgls gls = pp(pr_gls gls)
-let prglls glls = pp(pr_glls glls)
-let pproof p = pp(print_proof Evd.empty empty_named_context p)
+(* let ppsigmagoal g = pp(pr_goal (sig_it g)) *)
+(* let prgls gls = pp(pr_gls gls) *)
+(* let prglls glls = pp(pr_glls glls) *)
+(* let pproof p = pp(print_proof Evd.empty empty_named_context p) *)
let ppuni u = pp(pr_uni u)
@@ -402,7 +402,7 @@ let _ =
with
e -> Pp.pp (Cerrors.explain_exn e)
let _ =
- extend_vernac_command_grammar "PrintConstr"
+ extend_vernac_command_grammar "PrintConstr" None
[[GramTerminal "PrintConstr";
GramNonTerminal
(dummy_loc,ConstrArgType,Extend.Aentry ("constr","constr"),
@@ -419,7 +419,7 @@ let _ =
with
e -> Pp.pp (Cerrors.explain_exn e)
let _ =
- extend_vernac_command_grammar "PrintPureConstr"
+ extend_vernac_command_grammar "PrintPureConstr" None
[[GramTerminal "PrintPureConstr";
GramNonTerminal
(dummy_loc,ConstrArgType,Extend.Aentry ("constr","constr"),
diff --git a/ide/coq.ml b/ide/coq.ml
index 87f0c35e65..9b0afc7bed 100644
--- a/ide/coq.ml
+++ b/ide/coq.ml
@@ -19,7 +19,6 @@ open Printer
open Environ
open Evarutil
open Evd
-open Decl_mode
open Hipattern
open Tacmach
open Reductionops
@@ -90,11 +89,6 @@ let version () =
let is_in_loadpath coqtop dir =
Library.is_in_load_paths (System.physical_path_of_string dir)
-let is_in_proof_mode () =
- match Decl_mode.get_current_mode () with
- Decl_mode.Mode_none -> false
- | _ -> true
-
let user_error_loc l s =
raise (Stdpp.Exc_located (l, Util.UserError ("CoqIde", s)))
@@ -169,11 +163,6 @@ let rec attribute_of_vernac_command = function
| VernacSolve _ -> [SolveCommand]
| VernacSolveExistential _ -> [SolveCommand]
- (* MMode *)
- | VernacDeclProof -> [SolveCommand]
- | VernacReturn -> [SolveCommand]
- | VernacProofInstr _ -> [SolveCommand]
-
(* Auxiliary file and library management *)
| VernacRequireFrom _ -> []
| VernacAddLoadPath _ -> []
@@ -241,9 +230,15 @@ let rec attribute_of_vernac_command = function
| VernacProof (Tacexpr.TacId []) -> [OtherStatePreservingCommand]
| VernacProof _ -> []
+ | VernacProofMode _ -> []
+
+ | VernacSubproof _ -> [SolveCommand]
+ | VernacEndSubproof _ -> [SolveCommand]
+
(* Toplevel control *)
| VernacToplevelControl _ -> []
+
(* Extensions *)
| VernacExtend ("Subtac_Obligations", _) -> [GoalStartingCommand]
| VernacExtend _ -> []
@@ -508,13 +503,13 @@ let hyp_next_tac sigma env (id,_,ast) =
("inversion clear "^id_s), ("inversion_clear "^id_s^".\n")
]
-let concl_next_tac concl =
+let concl_next_tac sigma concl =
let expand s = (s,s^".\n") in
List.map expand ([
"intro";
"intros";
"intuition"
- ] @ (if Hipattern.is_equality_type concl.Evd.evar_concl then [
+ ] @ (if Hipattern.is_equality_type (Goal.V82.concl sigma concl) then [
"reflexivity";
"discriminate";
"symmetry"
@@ -538,30 +533,26 @@ let concl_next_tac concl =
let goals coqtop =
PrintOpt.enforce_hack ();
let pfts = Pfedit.get_pftreestate () in
- let sigma = Tacmach.evc_of_pftreestate pfts in
- let (all_goals,_) = Refiner.frontier (Refiner.proof_of_pftreestate pfts) in
+ let { it=all_goals ; sigma=sigma } = Proof.V82.subgoals pfts in
if all_goals = [] then
begin
Message (
- match Decl_mode.get_end_command pfts with
- | Some c -> "Subproof completed, now type "^c^".\n"
- | None ->
- let exl = Evarutil.non_instantiated sigma in
- if exl = [] then "Proof Completed.\n" else
- ("No more subgoals but non-instantiated existential variables:\n"^
- string_of_ppcmds (pr_evars_int 1 exl)))
+ let exl = Evarutil.non_instantiated sigma in
+ if exl = [] then "Proof Completed.\n" else
+ ("No more subgoals but non-instantiated existential variables:\n"^
+ string_of_ppcmds (pr_evars_int 1 exl)))
end
else
begin
let process_goal g =
- let env = Evd.evar_env g in
+ let env = Goal.V82.env sigma g in
let ccl =
- string_of_ppcmds (pr_ltype_env_at_top env g.Evd.evar_concl) in
+ string_of_ppcmds (pr_ltype_env_at_top env (Goal.V82.concl sigma g)) in
let process_hyp h_env d acc =
(string_of_ppcmds (pr_var_decl h_env d), hyp_next_tac sigma h_env d)::acc in
let hyps =
List.rev (Environ.fold_named_context process_hyp env ~init:[]) in
- (hyps,(ccl,concl_next_tac g))
+ (hyps,(ccl,concl_next_tac sigma g))
in
Goals (List.map process_goal all_goals)
end
diff --git a/ide/coqide.ml b/ide/coqide.ml
index 4fa0c28f51..5d604a62bf 100644
--- a/ide/coqide.ml
+++ b/ide/coqide.ml
@@ -730,14 +730,10 @@ object(self)
(String.make previous_line_spaces ' ')
end
+
method show_goals =
try
- match Decl_mode.get_current_mode () with
- Decl_mode.Mode_none -> ()
- | Decl_mode.Mode_tactic ->
- Ideproof.display (Ideproof.mode_tactic (fun _ _ -> ())) proof_view (Coq.goals Coq.dummy_coqtop)
- | Decl_mode.Mode_proof ->
- Ideproof.display Ideproof.mode_cesar proof_view (Coq.goals Coq.dummy_coqtop)
+ Ideproof.display (Ideproof.mode_tactic (fun _ _ -> ())) proof_view (Coq.goals Coq.dummy_coqtop)
with e ->
prerr_endline ("Don't worry be happy despite: "^Printexc.to_string e)
@@ -747,15 +743,10 @@ object(self)
if not full_goal_done then
begin
try
- match Decl_mode.get_current_mode () with
- Decl_mode.Mode_none -> ()
- | Decl_mode.Mode_tactic ->
- Ideproof.display
- (Ideproof.mode_tactic (fun s () -> ignore (self#insert_this_phrase_on_success
- true true false ("progress "^s) s)))
- proof_view (Coq.goals Coq.dummy_coqtop)
- | Decl_mode.Mode_proof ->
- Ideproof.display Ideproof.mode_cesar proof_view (Coq.goals Coq.dummy_coqtop)
+ Ideproof.display
+ (Ideproof.mode_tactic (fun s () -> ignore (self#insert_this_phrase_on_success
+ true true false ("progress "^s) s)))
+ proof_view (Coq.goals Coq.dummy_coqtop)
with e -> prerr_endline (Printexc.to_string e)
end
@@ -787,9 +778,8 @@ object(self)
try
full_goal_done <- false;
prerr_endline "Send_to_coq starting now";
- Decl_mode.clear_daimon_flag ();
let r = Coq.interp Coq.dummy_coqtop verbosely phrase in
- let is_complete = not (Decl_mode.get_daimon_flag ()) in
+ let is_complete = true in
let msg = read_stdout () in
sync display_output msg;
Some (is_complete,r)
diff --git a/kernel/cbytegen.ml b/kernel/cbytegen.ml
index e7859962eb..68ee15ab7c 100644
--- a/kernel/cbytegen.ml
+++ b/kernel/cbytegen.ml
@@ -335,7 +335,7 @@ let rec str_const c =
| App(f,args) ->
begin
match kind_of_term f with
- | Construct((kn,j),i) -> (* arnaud: Construct(((kn,j),i) as cstr) -> *)
+ | Construct((kn,j),i) ->
begin
let oib = lookup_mind kn !global_env in
let oip = oib.mind_packets.(j) in
@@ -405,7 +405,7 @@ let rec str_const c =
| _ -> Bconstr c
end
| Ind ind -> Bstrconst (Const_ind ind)
- | Construct ((kn,j),i) -> (*arnaud: Construct ((kn,j),i as cstr) -> *)
+ | Construct ((kn,j),i) ->
begin
(* spiwack: tries first to apply the run-time compilation
behavior of the constructor, as in 2/ above *)
@@ -664,19 +664,6 @@ and compile_str_cst reloc sc sz cont =
(* spiwack : compilation of constants with their arguments.
Makes a special treatment with 31-bit integer addition *)
and compile_const =
-(*arnaud: let code_construct kn cont =
- let f_cont =
- let else_lbl = Label.create () in
- Kareconst(2, else_lbl):: Kacc 0:: Kpop 1::
- Kaddint31:: Kreturn 0:: Klabel else_lbl::
- (* works as comp_app with nargs = 2 and tailcall cont [Kreturn 0]*)
- Kgetglobal (get_allias !global_env kn)::
- Kappterm(2, 2):: [] (* = discard_dead_code [Kreturn 0] *)
- in
- let lbl = Label.create () in
- fun_code := [Ksequence (add_grab 2 lbl f_cont, !fun_code)];
- Kclosure(lbl, 0)::cont
- in *)
fun reloc-> fun kn -> fun args -> fun sz -> fun cont ->
let nargs = Array.length args in
(* spiwack: checks if there is a specific way to compile the constant
diff --git a/kernel/cooking.ml b/kernel/cooking.ml
index c971ed2993..7605df9eb2 100644
--- a/kernel/cooking.ml
+++ b/kernel/cooking.ml
@@ -99,7 +99,7 @@ let expmod_constr modlist c =
in
if modlist = empty_modlist then c
- else under_outer_cast nf_betaiota (substrec c)
+ else substrec c
let abstract_constant_type =
List.fold_left (fun c d -> mkNamedProd_wo_LetIn d c)
diff --git a/kernel/term.mli b/kernel/term.mli
index 1ffe5eeb5f..ff71c3594e 100644
--- a/kernel/term.mli
+++ b/kernel/term.mli
@@ -161,9 +161,9 @@ val mkFix : fixpoint -> constr
(* If [funnames = [|f1,.....fn|]]
[typarray = [|t1,...tn|]]
[bodies = [b1,.....bn]] \par\noindent
- then [mkCoFix (i, (typsarray, funnames, bodies))]
+ then [mkCoFix (i, (funnames, typarray, bodies))]
constructs the ith function of the block
-
+
[CoFixpoint f1 = b1
with f2 = b2
...
diff --git a/lib/bstack.ml b/lib/bstack.ml
deleted file mode 100644
index 4191ccdb15..0000000000
--- a/lib/bstack.ml
+++ /dev/null
@@ -1,75 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(* $Id$ *)
-
-(* Queues of a given length *)
-
-open Util
-
-(* - size is the count of elements actually in the queue
- - depth is the the amount of elements pushed in the queue (and not popped)
- in particular, depth >= size always and depth > size if the queue overflowed
- (and forgot older elements)
- *)
-
-type 'a t = {mutable pos : int;
- mutable size : int;
- mutable depth : int;
- stack : 'a array}
-
-let create depth e =
- {pos = 0;
- size = 1;
- depth = 1;
- stack = Array.create depth e}
-
-(*
-let set_depth bs n = bs.depth <- n
-*)
-
-let incr_pos bs =
- bs.pos <- if bs.pos = Array.length bs.stack - 1 then 0 else bs.pos + 1
-
-let incr_size bs =
- if bs.size < Array.length bs.stack then bs.size <- bs.size + 1
-
-let decr_pos bs =
- bs.pos <- if bs.pos = 0 then Array.length bs.stack - 1 else bs.pos - 1
-
-let push bs e =
- incr_pos bs;
- incr_size bs;
- bs.depth <- bs.depth + 1;
- bs.stack.(bs.pos) <- e
-
-let pop bs =
- if bs.size > 1 then begin
- bs.size <- bs.size - 1;
- bs.depth <- bs.depth - 1;
- let oldpos = bs.pos in
- decr_pos bs;
- (* Release the memory at oldpos, by copying what is at new pos *)
- bs.stack.(oldpos) <- bs.stack.(bs.pos)
- end
-
-let top bs =
- if bs.size >= 1 then bs.stack.(bs.pos)
- else error "Nothing on the stack"
-
-let app_push bs f =
- if bs.size = 0 then error "Nothing on the stack"
- else push bs (f (bs.stack.(bs.pos)))
-
-let app_repl bs f =
- if bs.size = 0 then error "Nothing on the stack"
- else bs.stack.(bs.pos) <- f (bs.stack.(bs.pos))
-
-let depth bs = bs.depth
-
-let size bs = bs.size
diff --git a/lib/bstack.mli b/lib/bstack.mli
deleted file mode 100644
index cef8e1d9e5..0000000000
--- a/lib/bstack.mli
+++ /dev/null
@@ -1,22 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(*i $Id$ i*)
-
-(* Bounded stacks. If the depth is [None], then there is no depth limit. *)
-
-type 'a t
-
-val create : int -> 'a -> 'a t
-val push : 'a t -> 'a -> unit
-val app_push : 'a t -> ('a -> 'a) -> unit
-val app_repl : 'a t -> ('a -> 'a) -> unit
-val pop : 'a t -> unit
-val top : 'a t -> 'a
-val depth : 'a t -> int
-val size : 'a t -> int
diff --git a/lib/edit.ml b/lib/edit.ml
deleted file mode 100644
index fd870a21ba..0000000000
--- a/lib/edit.ml
+++ /dev/null
@@ -1,134 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(* $Id$ *)
-
-open Pp
-open Util
-
-type ('a,'b,'c) t = {
- mutable focus : 'a option;
- mutable last_focused_stk : 'a list;
- buf : ('a, 'b Bstack.t * 'c) Hashtbl.t }
-
-let empty () = {
- focus = None;
- last_focused_stk = [];
- buf = Hashtbl.create 17 }
-
-let focus e nd =
- if not (Hashtbl.mem e.buf nd) then invalid_arg "Edit.focus";
- begin match e.focus with
- | Some foc when foc <> nd ->
- e.last_focused_stk <- foc::(list_except foc e.last_focused_stk);
- | _ -> ()
- end;
- e.focus <- Some nd
-
-let unfocus e =
- match e.focus with
- | None -> invalid_arg "Edit.unfocus"
- | Some foc ->
- begin
- e.last_focused_stk <- foc::(list_except foc e.last_focused_stk);
- e.focus <- None
- end
-
-let last_focused e =
- match e.last_focused_stk with
- | [] -> None
- | f::_ -> Some f
-
-let restore_last_focus e =
- match e.last_focused_stk with
- | [] -> ()
- | f::_ -> focus e f
-
-let focusedp e =
- match e.focus with
- | None -> false
- | _ -> true
-
-let read e =
- match e.focus with
- | None -> None
- | Some d ->
- let (bs,c) = Hashtbl.find e.buf d in
- Some(d,Bstack.top bs,c)
-
-let mutate e f =
- match e.focus with
- | None -> invalid_arg "Edit.mutate"
- | Some d ->
- let (bs,c) = Hashtbl.find e.buf d in
- Bstack.app_push bs (f c)
-
-let rev_mutate e f =
- match e.focus with
- | None -> invalid_arg "Edit.rev_mutate"
- | Some d ->
- let (bs,c) = Hashtbl.find e.buf d in
- Bstack.app_repl bs (f c)
-
-let undo e n =
- match e.focus with
- | None -> invalid_arg "Edit.undo"
- | Some d ->
- let (bs,_) = Hashtbl.find e.buf d in
- if n >= Bstack.size bs then
- errorlabstrm "Edit.undo" (str"Undo stack exhausted");
- repeat n Bstack.pop bs
-
-(* Return the depth of the focused proof of [e] stack, this is used to
- put informations in coq prompt (in emacs mode). *)
-let depth e =
- match e.focus with
- | None -> invalid_arg "Edit.depth"
- | Some d ->
- let (bs,_) = Hashtbl.find e.buf d in
- Bstack.depth bs
-
-(* Undo focused proof of [e] to reach depth [n] *)
-let undo_todepth e n =
- match e.focus with
- | None ->
- if n <> 0
- then errorlabstrm "Edit.undo_todepth" (str"No proof in progress")
- else () (* if there is no proof in progress, then n must be zero *)
- | Some d ->
- let (bs,_) = Hashtbl.find e.buf d in
- let ucnt = Bstack.depth bs - n in
- if ucnt >= Bstack.size bs then
- errorlabstrm "Edit.undo_todepth" (str"Undo stack would be exhausted");
- repeat ucnt Bstack.pop bs
-
-let create e (d,b,c,usize) =
- if Hashtbl.mem e.buf d then
- errorlabstrm "Edit.create"
- (str"Already editing something of that name");
- let bs = Bstack.create usize b in
- Hashtbl.add e.buf d (bs,c)
-
-let delete e d =
- if not(Hashtbl.mem e.buf d) then
- errorlabstrm "Edit.delete" (str"No such editor");
- Hashtbl.remove e.buf d;
- e.last_focused_stk <- (list_except d e.last_focused_stk);
- match e.focus with
- | Some d' -> if d = d' then (e.focus <- None ; (restore_last_focus e))
- | None -> ()
-
-let dom e =
- let l = ref [] in
- Hashtbl.iter (fun x _ -> l := x :: !l) e.buf;
- !l
-
-let clear e =
- e.focus <- None;
- e.last_focused_stk <- [];
- Hashtbl.clear e.buf
diff --git a/lib/edit.mli b/lib/edit.mli
deleted file mode 100644
index d13d9c6f69..0000000000
--- a/lib/edit.mli
+++ /dev/null
@@ -1,63 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(*i $Id$ i*)
-
-(* The type of editors.
- * An editor is a finite map, ['a -> 'b], which knows how to apply
- * modification functions to the value in the range, and how to
- * focus on a member of the range.
- * It also supports the notion of a limited-depth undo, and that certain
- * modification actions do not push onto the undo stack, since they are
- * reversible. *)
-
-type ('a,'b,'c) t
-
-val empty : unit -> ('a,'b,'c) t
-
-(* sets the focus to the specified domain element *)
-val focus : ('a,'b,'c) t -> 'a -> unit
-
-(* unsets the focus which must not already be unfocused *)
-val unfocus : ('a,'b,'c) t -> unit
-
-(* gives the last focused element or [None] if none *)
-val last_focused : ('a,'b,'c) t -> 'a option
-
-(* are we focused ? *)
-val focusedp : ('a,'b,'c) t -> bool
-
-(* If we are focused, then return the current domain,range pair. *)
-val read : ('a,'b,'c) t -> ('a * 'b * 'c) option
-
-(* mutates the currently-focused range element, pushing its
- * old value onto the undo stack
- *)
-val mutate : ('a,'b,'c) t -> ('c -> 'b -> 'b) -> unit
-
-(* mutates the currently-focused range element, in place. *)
-val rev_mutate : ('a,'b,'c) t -> ('c -> 'b -> 'b) -> unit
-
-(* Pops the specified number of elements off of the undo stack, *
- reinstating the last popped element. The undo stack is independently
- managed for each range element. *)
-val undo : ('a,'b,'c) t -> int -> unit
-
-val create : ('a,'b,'c) t -> 'a * 'b * 'c * int -> unit
-val delete : ('a,'b,'c) t -> 'a -> unit
-
-val dom : ('a,'b,'c) t -> 'a list
-
-val clear : ('a,'b,'c) t -> unit
-
-(* [depth e] Returns the depth of the focused proof stack of [e], this
- is used to put informations in coq prompt (in emacs mode). *)
-val depth : ('a,'b,'c) t -> int
-
-(* [undo_todepth e n] Undoes focused proof of [e] to reach depth [n] *)
-val undo_todepth : ('a,'b,'c) t -> int -> unit
diff --git a/lib/lib.mllib b/lib/lib.mllib
index 1743ce26ca..86d29c4be6 100644
--- a/lib/lib.mllib
+++ b/lib/lib.mllib
@@ -10,8 +10,6 @@ Hashcons
Dyn
System
Envars
-Bstack
-Edit
Gset
Gmap
Fset
@@ -26,4 +24,4 @@ Rtree
Heap
Option
Dnet
-
+Store
diff --git a/lib/store.ml b/lib/store.ml
new file mode 100644
index 0000000000..8f1309531e
--- /dev/null
+++ b/lib/store.ml
@@ -0,0 +1,63 @@
+(***********************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
+(* \VV/ *************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(***********************************************************************)
+
+(* $Id$ *)
+
+(*** This module implements an "untyped store", in this particular case we
+ see it as an extensible record whose fields are left unspecified. ***)
+
+(* We give a short implementation of a universal type. This is mostly equivalent
+ to what is proposed by module Dyn.ml, except that it requires no explicit tag. *)
+module type Universal = sig
+ type t
+
+ type 'a etype = {
+ put : 'a -> t ;
+ get : t -> 'a option
+ }
+
+ val embed : unit -> 'a etype
+end
+
+(* We use a dynamic "name" allocator. But if we needed to serialise stores, we
+might want something static to avoid troubles with plugins order. *)
+
+let next =
+ let count = ref 0 in fun () ->
+ let n = !count in
+ incr count;
+ n
+
+type t = Obj.t Util.Intmap.t
+
+module Field = struct
+ type 'a field = {
+ set : 'a -> t -> t ;
+ get : t -> 'a option ;
+ remove : t -> t
+ }
+ type 'a t = 'a field
+end
+
+open Field
+
+let empty = Util.Intmap.empty
+
+let field () =
+ let fid = next () in
+ let set a s =
+ Util.Intmap.add fid (Obj.repr a) s
+ in
+ let get s =
+ try Some (Obj.obj (Util.Intmap.find fid s))
+ with _ -> None
+ in
+ let remove s =
+ Util.Intmap.remove fid s
+ in
+ { set = set ; get = get ; remove = remove }
diff --git a/lib/store.mli b/lib/store.mli
new file mode 100644
index 0000000000..0caeb2abba
--- /dev/null
+++ b/lib/store.mli
@@ -0,0 +1,27 @@
+(***********************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
+(* \VV/ *************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(***********************************************************************)
+
+(* $Id$ *)
+
+(*** This module implements an "untyped store", in this particular case we
+ see it as an extensible record whose fields are left unspecified. ***)
+
+type t
+
+module Field : sig
+ type 'a field = {
+ set : 'a -> t -> t ;
+ get : t -> 'a option ;
+ remove : t -> t
+ }
+ type 'a t = 'a field
+end
+
+val empty : t
+
+val field : unit -> 'a Field.field
diff --git a/parsing/egrammar.ml b/parsing/egrammar.ml
index 023ec0f3c7..b5ee1ae60c 100644
--- a/parsing/egrammar.ml
+++ b/parsing/egrammar.ml
@@ -222,12 +222,13 @@ let extend_tactic_grammar s gl =
let vernac_exts = ref []
let get_extend_vernac_grammars () = !vernac_exts
-let extend_vernac_command_grammar s gl =
+let extend_vernac_command_grammar s nt gl =
+ let nt = Option.default Vernac_.command nt in
vernac_exts := (s,gl) :: !vernac_exts;
let univ = get_univ "vernac" in
let mkact loc l = VernacExtend (s,List.map snd l) in
let rules = List.map (make_rule univ mkact make_prod_item) gl in
- Gram.extend Vernac_.command None [(None, None, List.rev rules)]
+ Gram.extend nt None [(None, None, List.rev rules)]
(**********************************************************************)
(** Grammar declaration for Tactic Notation (Coq level) *)
diff --git a/parsing/egrammar.mli b/parsing/egrammar.mli
index 1228b40cf3..a45ea9549c 100644
--- a/parsing/egrammar.mli
+++ b/parsing/egrammar.mli
@@ -59,7 +59,7 @@ val extend_tactic_grammar :
string -> grammar_prod_item list list -> unit
val extend_vernac_command_grammar :
- string -> grammar_prod_item list list -> unit
+ string -> vernac_expr Gram.Entry.e option -> grammar_prod_item list list -> unit
val get_extend_vernac_grammars :
unit -> (string * grammar_prod_item list list) list
diff --git a/parsing/g_decl_mode.ml4 b/parsing/g_decl_mode.ml4
deleted file mode 100644
index e73e54e7f5..0000000000
--- a/parsing/g_decl_mode.ml4
+++ /dev/null
@@ -1,252 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(*i camlp4deps: "parsing/grammar.cma" i*)
-(*i camlp4use: "pa_extend.cmo q_MLast.cmo" i*)
-
-(* $Id$ *)
-
-
-open Decl_expr
-open Names
-open Term
-open Genarg
-open Pcoq
-
-open Pcoq.Constr
-open Pcoq.Tactic
-open Pcoq.Vernac_
-
-let none_is_empty = function
- None -> []
- | Some l -> l
-
-GEXTEND Gram
-GLOBAL: proof_instr;
- thesis :
- [[ "thesis" -> Plain
- | "thesis"; "for"; i=ident -> (For i)
- ]];
- statement :
- [[ i=ident ; ":" ; c=constr -> {st_label=Name i;st_it=c}
- | i=ident -> {st_label=Anonymous;
- st_it=Topconstr.CRef (Libnames.Ident (loc, i))}
- | c=constr -> {st_label=Anonymous;st_it=c}
- ]];
- constr_or_thesis :
- [[ t=thesis -> Thesis t ] |
- [ c=constr -> This c
- ]];
- statement_or_thesis :
- [
- [ t=thesis -> {st_label=Anonymous;st_it=Thesis t} ]
- |
- [ i=ident ; ":" ; cot=constr_or_thesis -> {st_label=Name i;st_it=cot}
- | i=ident -> {st_label=Anonymous;
- st_it=This (Topconstr.CRef (Libnames.Ident (loc, i)))}
- | c=constr -> {st_label=Anonymous;st_it=This c}
- ]
- ];
- justification_items :
- [[ -> Some []
- | "by"; l=LIST1 constr SEP "," -> Some l
- | "by"; "*" -> None ]]
- ;
- justification_method :
- [[ -> None
- | "using"; tac = tactic -> Some tac ]]
- ;
- simple_cut_or_thesis :
- [[ ls = statement_or_thesis;
- j = justification_items;
- taco = justification_method
- -> {cut_stat=ls;cut_by=j;cut_using=taco} ]]
- ;
- simple_cut :
- [[ ls = statement;
- j = justification_items;
- taco = justification_method
- -> {cut_stat=ls;cut_by=j;cut_using=taco} ]]
- ;
- elim_type:
- [[ IDENT "induction" -> ET_Induction
- | IDENT "cases" -> ET_Case_analysis ]]
- ;
- block_type :
- [[ IDENT "claim" -> B_claim
- | IDENT "focus" -> B_focus
- | IDENT "proof" -> B_proof
- | et=elim_type -> B_elim et ]]
- ;
- elim_obj:
- [[ IDENT "on"; c=constr -> Real c
- | IDENT "of"; c=simple_cut -> Virtual c ]]
- ;
- elim_step:
- [[ IDENT "consider" ;
- h=consider_vars ; IDENT "from" ; c=constr -> Pconsider (c,h)
- | IDENT "per"; et=elim_type; obj=elim_obj -> Pper (et,obj)
- | IDENT "suffices"; ls=suff_clause;
- j = justification_items;
- taco = justification_method
- -> Psuffices {cut_stat=ls;cut_by=j;cut_using=taco} ]]
- ;
- rew_step :
- [[ "~=" ; c=simple_cut -> (Rhs,c)
- | "=~" ; c=simple_cut -> (Lhs,c)]]
- ;
- cut_step:
- [[ "then"; tt=elim_step -> Pthen tt
- | "then"; c=simple_cut_or_thesis -> Pthen (Pcut c)
- | IDENT "thus"; tt=rew_step -> Pthus (let s,c=tt in Prew (s,c))
- | IDENT "thus"; c=simple_cut_or_thesis -> Pthus (Pcut c)
- | IDENT "hence"; c=simple_cut_or_thesis -> Phence (Pcut c)
- | tt=elim_step -> tt
- | tt=rew_step -> let s,c=tt in Prew (s,c);
- | IDENT "have"; c=simple_cut_or_thesis -> Pcut c;
- | IDENT "claim"; c=statement -> Pclaim c;
- | IDENT "focus"; IDENT "on"; c=statement -> Pfocus c;
- | "end"; bt = block_type -> Pend bt;
- | IDENT "escape" -> Pescape ]]
- ;
- (* examiner s'il est possible de faire R _ et _ R pour R une relation qcq*)
- loc_id:
- [[ id=ident -> fun x -> (loc,(id,x)) ]];
- hyp:
- [[ id=loc_id -> id None ;
- | id=loc_id ; ":" ; c=constr -> id (Some c)]]
- ;
- consider_vars:
- [[ name=hyp -> [Hvar name]
- | name=hyp; ","; v=consider_vars -> (Hvar name) :: v
- | name=hyp;
- IDENT "such"; IDENT "that"; h=consider_hyps -> (Hvar name)::h
- ]]
- ;
- consider_hyps:
- [[ st=statement; IDENT "and"; h=consider_hyps -> Hprop st::h
- | st=statement; IDENT "and";
- IDENT "consider" ; v=consider_vars -> Hprop st::v
- | st=statement -> [Hprop st]
- ]]
- ;
- assume_vars:
- [[ name=hyp -> [Hvar name]
- | name=hyp; ","; v=assume_vars -> (Hvar name) :: v
- | name=hyp;
- IDENT "such"; IDENT "that"; h=assume_hyps -> (Hvar name)::h
- ]]
- ;
- assume_hyps:
- [[ st=statement; IDENT "and"; h=assume_hyps -> Hprop st::h
- | st=statement; IDENT "and";
- IDENT "we"; IDENT "have" ; v=assume_vars -> Hprop st::v
- | st=statement -> [Hprop st]
- ]]
- ;
- assume_clause:
- [[ IDENT "we" ; IDENT "have" ; v=assume_vars -> v
- | h=assume_hyps -> h ]]
- ;
- suff_vars:
- [[ name=hyp; IDENT "to"; IDENT "show" ; c = constr_or_thesis ->
- [Hvar name],c
- | name=hyp; ","; v=suff_vars ->
- let (q,c) = v in ((Hvar name) :: q),c
- | name=hyp;
- IDENT "such"; IDENT "that"; h=suff_hyps ->
- let (q,c) = h in ((Hvar name) :: q),c
- ]];
- suff_hyps:
- [[ st=statement; IDENT "and"; h=suff_hyps ->
- let (q,c) = h in (Hprop st::q),c
- | st=statement; IDENT "and";
- IDENT "to" ; IDENT "have" ; v=suff_vars ->
- let (q,c) = v in (Hprop st::q),c
- | st=statement; IDENT "to"; IDENT "show" ; c = constr_or_thesis ->
- [Hprop st],c
- ]]
- ;
- suff_clause:
- [[ IDENT "to" ; IDENT "have" ; v=suff_vars -> v
- | h=suff_hyps -> h ]]
- ;
- let_vars:
- [[ name=hyp -> [Hvar name]
- | name=hyp; ","; v=let_vars -> (Hvar name) :: v
- | name=hyp; IDENT "be";
- IDENT "such"; IDENT "that"; h=let_hyps -> (Hvar name)::h
- ]]
- ;
- let_hyps:
- [[ st=statement; IDENT "and"; h=let_hyps -> Hprop st::h
- | st=statement; IDENT "and"; "let"; v=let_vars -> Hprop st::v
- | st=statement -> [Hprop st]
- ]];
- given_vars:
- [[ name=hyp -> [Hvar name]
- | name=hyp; ","; v=given_vars -> (Hvar name) :: v
- | name=hyp; IDENT "such"; IDENT "that"; h=given_hyps -> (Hvar name)::h
- ]]
- ;
- given_hyps:
- [[ st=statement; IDENT "and"; h=given_hyps -> Hprop st::h
- | st=statement; IDENT "and"; IDENT "given"; v=given_vars -> Hprop st::v
- | st=statement -> [Hprop st]
- ]];
- suppose_vars:
- [[name=hyp -> [Hvar name]
- |name=hyp; ","; v=suppose_vars -> (Hvar name) :: v
- |name=hyp; OPT[IDENT "be"];
- IDENT "such"; IDENT "that"; h=suppose_hyps -> (Hvar name)::h
- ]]
- ;
- suppose_hyps:
- [[ st=statement_or_thesis; IDENT "and"; h=suppose_hyps -> Hprop st::h
- | st=statement_or_thesis; IDENT "and"; IDENT "we"; IDENT "have";
- v=suppose_vars -> Hprop st::v
- | st=statement_or_thesis -> [Hprop st]
- ]]
- ;
- suppose_clause:
- [[ IDENT "we"; IDENT "have"; v=suppose_vars -> v;
- | h=suppose_hyps -> h ]]
- ;
- intro_step:
- [[ IDENT "suppose" ; h=assume_clause -> Psuppose h
- | IDENT "suppose" ; IDENT "it"; IDENT "is" ; c=pattern LEVEL "0" ;
- po=OPT[ "with"; p=LIST1 hyp SEP ","-> p ] ;
- ho=OPT[ IDENT "and" ; h=suppose_clause -> h ] ->
- Pcase (none_is_empty po,c,none_is_empty ho)
- | "let" ; v=let_vars -> Plet v
- | IDENT "take"; witnesses = LIST1 constr SEP "," -> Ptake witnesses
- | IDENT "assume"; h=assume_clause -> Passume h
- | IDENT "given"; h=given_vars -> Pgiven h
- | IDENT "define"; id=ident; args=LIST0 hyp;
- "as"; body=constr -> Pdefine(id,args,body)
- | IDENT "reconsider"; id=ident; "as" ; typ=constr -> Pcast (This id,typ)
- | IDENT "reconsider"; t=thesis; "as" ; typ=constr -> Pcast (Thesis t ,typ)
- ]]
- ;
- emphasis :
- [[ -> 0
- | "*" -> 1
- | "**" -> 2
- | "***" -> 3
- ]]
- ;
- bare_proof_instr:
- [[ c = cut_step -> c ;
- | i = intro_step -> i ]]
- ;
- proof_instr :
- [[ e=emphasis;i=bare_proof_instr -> {emph=e;instr=i}]]
- ;
-END;;
-
-
diff --git a/parsing/g_proofs.ml4 b/parsing/g_proofs.ml4
index 39e577b88c..27ad1e9645 100644
--- a/parsing/g_proofs.ml4
+++ b/parsing/g_proofs.ml4
@@ -40,6 +40,7 @@ GEXTEND Gram
[ [ IDENT "Goal"; c = lconstr -> VernacGoal c
| IDENT "Proof" -> VernacProof (Tacexpr.TacId [])
| IDENT "Proof"; "with"; ta = tactic -> VernacProof ta
+ | IDENT "Proof" ; IDENT "Mode" ; mn = string -> VernacProofMode mn
| IDENT "Abort" -> VernacAbort None
| IDENT "Abort"; IDENT "All" -> VernacAbortAll
| IDENT "Abort"; id = identref -> VernacAbort (Some id)
@@ -66,6 +67,9 @@ GEXTEND Gram
| IDENT "Focus" -> VernacFocus None
| IDENT "Focus"; n = natural -> VernacFocus (Some n)
| IDENT "Unfocus" -> VernacUnfocus
+ | IDENT "BeginSubproof" -> VernacSubproof None
+ | IDENT "BeginSubproof"; n = natural -> VernacSubproof (Some n)
+ | IDENT "EndSubproof" -> VernacEndSubproof
| IDENT "Show" -> VernacShow (ShowGoal None)
| IDENT "Show"; n = natural -> VernacShow (ShowGoal (Some n))
| IDENT "Show"; IDENT "Implicit"; IDENT "Arguments"; n = OPT natural ->
diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4
index 008a6ac510..db683b9a91 100644
--- a/parsing/g_vernac.ml4
+++ b/parsing/g_vernac.ml4
@@ -19,7 +19,6 @@ open Topconstr
open Extend
open Vernacexpr
open Pcoq
-open Decl_mode
open Tactic
open Decl_kinds
open Genarg
@@ -40,7 +39,6 @@ let _ = List.iter (fun s -> Lexer.add_token ("",s)) vernac_kw
let check_command = Gram.Entry.create "vernac:check_command"
let tactic_mode = Gram.Entry.create "vernac:tactic_command"
-let proof_mode = Gram.Entry.create "vernac:proof_command"
let noedit_mode = Gram.Entry.create "vernac:noedit_command"
let class_rawexpr = Gram.Entry.create "vernac:class_rawexpr"
@@ -50,13 +48,23 @@ let decl_notation = Gram.Entry.create "vernac:decl_notation"
let typeclass_context = Gram.Entry.create "vernac:typeclass_context"
let record_field = Gram.Entry.create "vernac:record_field"
let of_type_with_opt_coercion = Gram.Entry.create "vernac:of_type_with_opt_coercion"
+let subgoal_command = Gram.Entry.create "proof_mode:subgoal_command"
let instance_name = Gram.Entry.create "vernac:instance_name"
-let get_command_entry () =
- match Decl_mode.get_current_mode () with
- Mode_proof -> proof_mode
- | Mode_tactic -> tactic_mode
- | Mode_none -> noedit_mode
+let command_entry = ref noedit_mode
+let set_command_entry e = command_entry := e
+let get_command_entry () = !command_entry
+
+
+(* Registers the Classic Proof Mode (which uses [tactic_mode] as a parser for
+ proof editing and changes nothing else). Then sets it as the default proof mode. *)
+let set_tactic_mode () = set_command_entry tactic_mode
+let set_noedit_mode () = set_command_entry noedit_mode
+let _ = Proof_global.register_proof_mode {Proof_global.
+ name = "Classic" ;
+ set = set_tactic_mode ;
+ reset = set_noedit_mode
+ }
let default_command_entry =
Gram.Entry.of_parser "command_entry"
@@ -64,7 +72,7 @@ let default_command_entry =
let no_hook _ _ = ()
GEXTEND Gram
- GLOBAL: vernac gallina_ext tactic_mode proof_mode noedit_mode;
+ GLOBAL: vernac gallina_ext tactic_mode noedit_mode subgoal_command;
vernac: FIRST
[ [ IDENT "Time"; v = vernac -> VernacTime v
| IDENT "Timeout"; n = natural; v = vernac -> VernacTimeout(n,v)
@@ -89,25 +97,25 @@ GEXTEND Gram
| -> locality_flag := None ] ]
;
noedit_mode:
- [ [ c = subgoal_command -> c None] ]
+ [ [ c = subgoal_command -> c None None] ]
;
tactic_mode:
[ [ gln = OPT[n=natural; ":" -> n];
- tac = subgoal_command -> tac gln ] ]
+ tac = subgoal_command -> tac gln None
+ | b = bullet; tac = subgoal_command -> tac None (Some b)] ]
+ ;
+ bullet:
+ [ [ "-" -> Dash
+ | "*" -> Star
+ | "+" -> Plus ] ]
;
- subgoal_command:
- [ [ c = check_command; "." -> c
+ subgoal_command:
+ [ [ c = check_command; "." -> fun g _ -> c g
| tac = Tactic.tactic;
use_dft_tac = [ "." -> false | "..." -> true ] ->
- (fun g ->
- let g = match g with Some gl -> gl | _ -> 1 in
- VernacSolve(g,tac,use_dft_tac)) ] ]
- ;
- proof_mode:
- [ [ instr = proof_instr; "." -> VernacProofInstr instr ] ]
- ;
- proof_mode: LAST
- [ [ c=subgoal_command -> c (Some 1) ] ]
+ (fun g b ->
+ let g = Option.default 1 g in
+ VernacSolve(g,b,tac,use_dft_tac)) ] ]
;
located_vernac:
[ [ v = vernac -> loc, v ] ]
@@ -713,10 +721,7 @@ GEXTEND Gram
| IDENT "Remove"; table = IDENT; field = IDENT; v= LIST1 option_ref_value
-> VernacRemoveOption ([table;field], v)
| IDENT "Remove"; table = IDENT; v = LIST1 option_ref_value ->
- VernacRemoveOption ([table], v)
-
- | IDENT "proof" -> VernacDeclProof
- | "return" -> VernacReturn ]]
+ VernacRemoveOption ([table], v) ]]
;
check_command: (* TODO: rapprocher Eval et Check *)
[ [ IDENT "Eval"; r = Tactic.red_expr; "in"; c = lconstr ->
diff --git a/parsing/grammar.mllib b/parsing/grammar.mllib
index 248a8ad9ad..483538da62 100644
--- a/parsing/grammar.mllib
+++ b/parsing/grammar.mllib
@@ -14,6 +14,7 @@ Hashcons
Predicate
Rtree
Option
+Store
Names
Univ
diff --git a/parsing/highparsing.mllib b/parsing/highparsing.mllib
index 3eb27abbb6..eed6caea30 100644
--- a/parsing/highparsing.mllib
+++ b/parsing/highparsing.mllib
@@ -4,4 +4,3 @@ G_prim
G_proofs
G_tactic
G_ltac
-G_decl_mode
diff --git a/parsing/parsing.mllib b/parsing/parsing.mllib
index c0c1817d1d..84a08d549a 100644
--- a/parsing/parsing.mllib
+++ b/parsing/parsing.mllib
@@ -6,7 +6,6 @@ G_xml
Ppconstr
Printer
Pptactic
-Ppdecl_proof
Tactic_printer
Printmod
Prettyp
diff --git a/parsing/pcoq.ml4 b/parsing/pcoq.ml4
index 7120f72d20..de15d8a7cb 100644
--- a/parsing/pcoq.ml4
+++ b/parsing/pcoq.ml4
@@ -373,8 +373,6 @@ module Vernac_ =
let command = gec_vernac "command"
let syntax = gec_vernac "syntax_command"
let vernac = gec_vernac "Vernac.vernac"
- let proof_instr = Gram.Entry.create "proofmode:instr"
-
let vernac_eoi = eoi_entry vernac
(* Main vernac entry *)
diff --git a/parsing/pcoq.mli b/parsing/pcoq.mli
index ed370a9953..7ed05ed5c7 100644
--- a/parsing/pcoq.mli
+++ b/parsing/pcoq.mli
@@ -247,7 +247,6 @@ module Vernac_ :
val syntax : vernac_expr Gram.Entry.e
val vernac : vernac_expr Gram.Entry.e
val vernac_eoi : vernac_expr Gram.Entry.e
- val proof_instr : Decl_expr.raw_proof_instr Gram.Entry.e
end
(* The main entry: reads an optional vernac command *)
diff --git a/parsing/ppvernac.ml b/parsing/ppvernac.ml
index 9ebf77adba..dc61aaa261 100644
--- a/parsing/ppvernac.ml
+++ b/parsing/ppvernac.ml
@@ -774,23 +774,16 @@ let rec pr_vernac = function
hov 2 (str"Include " ++
prlist_with_sep (fun () -> str " <+ ") pr_m mexprs)
(* Solving *)
- | VernacSolve (i,tac,deftac) ->
+ | VernacSolve (i,b,tac,deftac) ->
(if i = 1 then mt() else int i ++ str ": ") ++
+ (match b with None -> mt () | Some Dash -> str"-" | Some Star -> str"*" | Some Plus -> str"+") ++
pr_raw_tactic tac
- ++ (try if deftac & Pfedit.get_end_tac() <> None then str ".." else mt ()
+ ++ (try if deftac then str ".." else mt ()
with UserError _|Stdpp.Exc_located _ -> mt())
| VernacSolveExistential (i,c) ->
str"Existential " ++ int i ++ pr_lconstrarg c
- (* MMode *)
-
- | VernacProofInstr instr -> anomaly "Not implemented"
- | VernacDeclProof -> str "proof"
- | VernacReturn -> str "return"
-
- (* /MMode *)
-
(* Auxiliary file and library management *)
| VernacRequireFrom (exp,spe,f) -> hov 2
(str"Require" ++ spc() ++ pr_require_token exp ++
diff --git a/parsing/printer.ml b/parsing/printer.ml
index d9dced7916..cd07c4e15c 100644
--- a/parsing/printer.ml
+++ b/parsing/printer.ml
@@ -22,13 +22,14 @@ open Libnames
open Nametab
open Evd
open Proof_type
-open Decl_mode
open Refiner
open Pfedit
open Ppconstr
open Constrextern
open Tacexpr
+open Store.Field
+
let emacs_str s alts =
match !Flags.print_emacs, !Flags.print_emacs_safechar with
| true, true -> alts
@@ -265,18 +266,13 @@ let pr_subgoal_metas metas env=
hv 0 (prlist_with_sep mt pr_one metas)
(* display complete goal *)
-let default_pr_goal g =
- let env = evar_unfiltered_env g in
+let default_pr_goal gs =
+ let (g,sigma) = Goal.V82.nf_evar (project gs) (sig_it gs) in
+ let env = Goal.V82.unfiltered_env sigma g in
let preamb,thesis,penv,pc =
- if g.evar_extra = None then
- mt (), mt (),
- pr_context_of env,
- pr_ltype_env_at_top env g.evar_concl
- else
- (str " *** Declarative Mode ***" ++ fnl ()++fnl ()),
- (str "thesis := " ++ fnl ()),
- pr_context_of env,
- pr_ltype_env_at_top env g.evar_concl
+ mt (), mt (),
+ pr_context_of env,
+ pr_ltype_env_at_top env (Goal.V82.concl sigma g)
in
preamb ++
str" " ++ hv 0 (penv ++ fnl () ++
@@ -285,9 +281,10 @@ let default_pr_goal g =
thesis ++ str " " ++ pc) ++ fnl ()
(* display the conclusion of a goal *)
-let pr_concl n g =
- let env = evar_env g in
- let pc = pr_ltype_env_at_top env g.evar_concl in
+let pr_concl n sigma g =
+ let (g,sigma) = Goal.V82.nf_evar sigma g in
+ let env = Goal.V82.env sigma g in
+ let pc = pr_ltype_env_at_top env (Goal.V82.concl sigma g) in
str (emacs_str (String.make 1 (Char.chr 253)) "") ++
str "subgoal " ++ int n ++ str " is:" ++ cut () ++ str" " ++ pc
@@ -313,12 +310,12 @@ let rec pr_evars_int i = function
str (string_of_existential ev) ++ str " : " ++ pegl)) ++
fnl () ++ pei
-let default_pr_subgoal n =
+let default_pr_subgoal n sigma =
let rec prrec p = function
| [] -> error "No such goal."
| g::rest ->
if p = 1 then
- let pg = default_pr_goal g in
+ let pg = default_pr_goal { sigma=sigma ; it=g } in
v 0 (str "subgoal " ++ int n ++ str " is:" ++ cut () ++ pg)
else
prrec (p-1) rest
@@ -343,17 +340,17 @@ let default_pr_subgoals close_cmd sigma = function
str "variables:" ++ fnl () ++ (hov 0 pei))
end
| [g] ->
- let pg = default_pr_goal g in
+ let pg = default_pr_goal { it = g ; sigma = sigma } in
v 0 (str ("1 "^"subgoal") ++cut () ++ pg)
| g1::rest ->
let rec pr_rec n = function
| [] -> (mt ())
| g::rest ->
- let pc = pr_concl n g in
+ let pc = pr_concl n sigma g in
let prest = pr_rec (n+1) rest in
(cut () ++ pc ++ prest)
in
- let pg1 = default_pr_goal g1 in
+ let pg1 = default_pr_goal { it = g1 ; sigma = sigma } in
let prest = pr_rec 2 rest in
v 0 (int(List.length rest+1) ++ str" subgoals" ++ cut ()
++ pg1 ++ prest ++ fnl ())
@@ -365,8 +362,8 @@ let default_pr_subgoals close_cmd sigma = function
type printer_pr = {
pr_subgoals : string option -> evar_map -> goal list -> std_ppcmds;
- pr_subgoal : int -> goal list -> std_ppcmds;
- pr_goal : goal -> std_ppcmds;
+ pr_subgoal : int -> evar_map -> goal list -> std_ppcmds;
+ pr_goal : goal sigma -> std_ppcmds;
}
let default_printer_pr = {
@@ -387,25 +384,29 @@ let pr_goal x = !printer_pr.pr_goal x
(**********************************************************************)
let pr_open_subgoals () =
- let pfts = get_pftreestate () in
- let gls = fst (frontier (proof_of_pftreestate pfts)) in
- match focus() with
- | 0 ->
- let sigma = (top_goal_of_pftreestate pfts).sigma in
- let close_cmd = Decl_mode.get_end_command pfts in
- pr_subgoals close_cmd sigma gls
- | n ->
- assert (n > List.length gls);
- if List.length gls < 2 then
- pr_subgoal n gls
- else
- (* LEM TODO: this way of saying how many subgoals has to be abstracted out*)
- v 0 (int(List.length gls) ++ str" subgoals" ++ cut () ++
- pr_subgoal n gls)
+ let p = Proof_global.give_me_the_proof () in
+ let { Evd.it = goals ; sigma = sigma } = Proof.V82.subgoals p in
+ begin match goals with
+ | [] -> let { Evd.it = bgoals ; sigma = bsigma } = Proof.V82.background_subgoals p in
+ begin match bgoals with
+ | [] -> pr_subgoals None sigma goals
+ | _ -> pr_subgoals None bsigma bgoals ++ fnl () ++ fnl () ++
+ str"This subproof is complete, but there are still unfocused goals:"
+ (* spiwack: to stay compatible with the proof general and coqide,
+ I use print the message after the goal. It would be better to have
+ something like:
+ str"This subproof is complete, but there are still unfocused goals:"
+ ++ fnl () ++ fnl () ++ pr_subgoals None bsigma bgoals
+ instead. But it doesn't quite work.
+ *)
+ end
+ | _ -> pr_subgoals None sigma goals
+ end
let pr_nth_open_subgoal n =
- let pf = proof_of_pftreestate (get_pftreestate ()) in
- pr_subgoal n (fst (frontier pf))
+ let pf = get_pftreestate () in
+ let { it=gls ; sigma=sigma } = Proof.V82.subgoals pf in
+ pr_subgoal n sigma gls
(* Elementary tactics *)
diff --git a/parsing/printer.mli b/parsing/printer.mli
index 1797eaf22b..2c1586abf7 100644
--- a/parsing/printer.mli
+++ b/parsing/printer.mli
@@ -101,9 +101,9 @@ val pr_transparent_state : transparent_state -> std_ppcmds
(* Proofs *)
-val pr_goal : goal -> std_ppcmds
+val pr_goal : goal sigma -> std_ppcmds
val pr_subgoals : string option -> evar_map -> goal list -> std_ppcmds
-val pr_subgoal : int -> goal list -> std_ppcmds
+val pr_subgoal : int -> evar_map -> goal list -> std_ppcmds
val pr_open_subgoals : unit -> std_ppcmds
val pr_nth_open_subgoal : int -> std_ppcmds
@@ -130,8 +130,8 @@ val pr_assumptionset : env -> Term.types Environ.ContextObjectMap.t ->std_ppcmds
type printer_pr = {
pr_subgoals : string option -> evar_map -> goal list -> std_ppcmds;
- pr_subgoal : int -> goal list -> std_ppcmds;
- pr_goal : goal -> std_ppcmds;
+ pr_subgoal : int -> evar_map -> goal list -> std_ppcmds;
+ pr_goal : goal sigma -> std_ppcmds;
};;
val set_printer_pr : printer_pr -> unit
diff --git a/parsing/tactic_printer.ml b/parsing/tactic_printer.ml
index c09b3431e5..bf554acf69 100644
--- a/parsing/tactic_printer.ml
+++ b/parsing/tactic_printer.ml
@@ -14,8 +14,6 @@ open Sign
open Evd
open Tacexpr
open Proof_type
-open Proof_trees
-open Decl_expr
open Logic
open Printer
@@ -26,16 +24,12 @@ let pr_tactic = function
| t ->
Pptactic.pr_tactic (Global.env()) t
-let pr_proof_instr instr =
- Ppdecl_proof.pr_proof_instr (Global.env()) instr
-
let pr_rule = function
| Prim r -> hov 0 (pr_prim_rule r)
| Nested(cmpd,_) ->
begin
match cmpd with
| Tactic (texp,_) -> hov 0 (pr_tactic texp)
- | Proof_instr (_,instr) -> hov 0 (pr_proof_instr instr)
end
| Daimon -> str "<Daimon>"
| Decl_proof _ -> str "proof"
@@ -62,33 +56,23 @@ let pr_rule_dot_fnl = function
exception Different
-(* We remove from the var context of env what is already in osign *)
-let thin_sign osign sign =
- Sign.fold_named_context
- (fun (id,c,ty as d) sign ->
- try
- if Sign.lookup_named id osign = (id,c,ty) then sign
- else raise Different
- with Not_found | Different -> Environ.push_named_context_val d sign)
- sign ~init:Environ.empty_named_context_val
-
-let rec print_proof _sigma osign pf =
- let hyps = Environ.named_context_of_val pf.goal.evar_hyps in
- let hyps' = thin_sign osign hyps in
+let rec print_proof sigma osign pf =
+ (* spiwack: [osign] is currently ignored, not sure if this function is even used. *)
+ let hyps = Environ.named_context_of_val (Goal.V82.hyps sigma pf.goal) in
match pf.ref with
| None ->
- hov 0 (pr_goal {pf.goal with evar_hyps=hyps'})
+ hov 0 (pr_goal {sigma = sigma; it=pf.goal })
| Some(r,spfl) ->
hov 0
- (hov 0 (pr_goal {pf.goal with evar_hyps=hyps'}) ++
+ (hov 0 (pr_goal {sigma = sigma; it=pf.goal }) ++
spc () ++ str" BY " ++
hov 0 (pr_rule r) ++ fnl () ++
str" " ++
- hov 0 (prlist_with_sep pr_fnl (print_proof _sigma hyps) spfl))
+ hov 0 (prlist_with_sep pr_fnl (print_proof sigma hyps) spfl))
-let pr_change gl =
+let pr_change sigma gl =
str"change " ++
- pr_lconstr_env (Global.env_of_context gl.evar_hyps) gl.evar_concl ++ str"."
+ pr_lconstr_env (Goal.V82.env sigma gl) (Goal.V82.concl sigma gl) ++ str"."
let print_decl_script tac_printer ?(nochange=true) sigma pf =
let rec print_prf pf =
@@ -97,36 +81,10 @@ let print_decl_script tac_printer ?(nochange=true) sigma pf =
(if nochange then
(str"<Your Proof Text here>")
else
- pr_change pf.goal)
+ pr_change sigma pf.goal)
++ fnl ()
| Some (Daimon,[]) -> str "(* Some proof has been skipped here *)"
| Some (Prim Change_evars,[subpf]) -> print_prf subpf
- | Some (Nested(Proof_instr (opened,instr),_) as rule,subprfs) ->
- begin
- match instr.instr,subprfs with
- Pescape,[{ref=Some(_,subsubprfs)}] ->
- hov 7
- (pr_rule_dot_fnl rule ++
- prlist_with_sep pr_fnl tac_printer subsubprfs) ++ fnl () ++
- if opened then mt () else str "return."
- | Pclaim _,[body;cont] ->
- hov 2 (pr_rule_dot_fnl rule ++ print_prf body) ++ fnl () ++
- (if opened then mt () else str "end claim." ++ fnl ()) ++
- print_prf cont
- | Pfocus _,[body;cont] ->
- hov 2 (pr_rule_dot_fnl rule ++ print_prf body) ++
- fnl () ++
- (if opened then mt () else str "end focus." ++ fnl ()) ++
- print_prf cont
- | (Psuppose _ |Pcase (_,_,_)),[body;cont] ->
- hov 2 (pr_rule_dot_fnl rule ++ print_prf body) ++ fnl () ++
- print_prf cont
- | _,[next] ->
- pr_rule_dot_fnl rule ++ print_prf next
- | _,[] ->
- pr_rule_dot rule
- | _,_ -> anomaly "unknown branching instruction"
- end
| _ -> anomaly "Not Applicable" in
print_prf pf
@@ -137,12 +95,12 @@ let print_script ?(nochange=true) sigma pf =
(if nochange then
(str"<Your Tactic Text here>")
else
- pr_change pf.goal)
+ pr_change sigma pf.goal)
++ fnl ()
| Some(Decl_proof opened,script) ->
assert (List.length script = 1);
begin
- if nochange then (mt ()) else (pr_change pf.goal ++ fnl ())
+ if nochange then (mt ()) else (pr_change sigma pf.goal ++ fnl ())
end ++
begin
hov 0 (str "proof." ++ fnl () ++
@@ -153,10 +111,10 @@ let print_script ?(nochange=true) sigma pf =
if opened then mt () else (str "end proof." ++ fnl ())
end
| Some(Daimon,spfl) ->
- ((if nochange then (mt ()) else (pr_change pf.goal ++ fnl ())) ++
+ ((if nochange then (mt ()) else (pr_change sigma pf.goal ++ fnl ())) ++
prlist_with_sep pr_fnl print_prf spfl )
| Some(rule,spfl) ->
- ((if nochange then (mt ()) else (pr_change pf.goal ++ fnl ())) ++
+ ((if nochange then (mt ()) else (pr_change sigma pf.goal ++ fnl ())) ++
pr_rule_dot_fnl rule ++
prlist_with_sep pr_fnl print_prf spfl ) in
print_prf pf
@@ -168,13 +126,12 @@ let print_treescript ?(nochange=true) sigma pf =
match pf.ref with
| None ->
if nochange then
- if pf.goal.evar_extra=None then str"<Your Tactic Text here>"
- else str"<Your Proof Text here>"
- else pr_change pf.goal
+ str"<Your Proof Text here>"
+ else pr_change sigma pf.goal
| Some(Decl_proof opened,script) ->
assert (List.length script = 1);
begin
- if nochange then mt () else pr_change pf.goal ++ fnl ()
+ if nochange then mt () else pr_change sigma pf.goal ++ fnl ()
end ++
hov 0
begin str "proof." ++ fnl () ++
@@ -184,16 +141,16 @@ let print_treescript ?(nochange=true) sigma pf =
if opened then mt () else (str "end proof." ++ fnl ())
end
| Some(Daimon,spfl) ->
- (if nochange then mt () else pr_change pf.goal ++ fnl ()) ++
+ (if nochange then mt () else pr_change sigma pf.goal ++ fnl ()) ++
prlist_with_sep pr_fnl (print_script ~nochange sigma) spfl
| Some(r,spfl) ->
let indent = if List.length spfl >= 2 then 1 else 0 in
- (if nochange then mt () else pr_change pf.goal ++ fnl ()) ++
+ (if nochange then mt () else pr_change sigma pf.goal ++ fnl ()) ++
hv indent (pr_rule_dot_fnl r ++ prlist_with_sep fnl print_prf spfl)
in hov 0 (print_prf pf)
let rec print_info_script sigma osign pf =
- let {evar_hyps=sign; evar_concl=cl} = pf.goal in
+ let sign = Goal.V82.hyps sigma pf.goal in
match pf.ref with
| None -> (mt ())
| Some(r,spfl) ->
@@ -214,12 +171,4 @@ let rec print_info_script sigma osign pf =
let format_print_info_script sigma osign pf =
hov 0 (print_info_script sigma osign pf)
-let print_subscript sigma sign pf =
- if is_tactic_proof pf then
- format_print_info_script sigma sign (subproof_of_proof pf)
- else
- format_print_info_script sigma sign pf
-
-let _ = Refiner.set_info_printer print_subscript
-let _ = Refiner.set_proof_printer print_proof
diff --git a/parsing/tactic_printer.mli b/parsing/tactic_printer.mli
index d46f19c64c..96cbeb9a84 100644
--- a/parsing/tactic_printer.mli
+++ b/parsing/tactic_printer.mli
@@ -21,7 +21,6 @@ open Proof_type
val print_proof : evar_map -> named_context -> proof_tree -> std_ppcmds
val pr_rule : rule -> std_ppcmds
val pr_tactic : tactic_expr -> std_ppcmds
-val pr_proof_instr : Decl_expr.proof_instr -> Pp.std_ppcmds
val print_script :
?nochange:bool -> evar_map -> proof_tree -> std_ppcmds
val print_treescript :
diff --git a/parsing/vernacextend.ml4 b/parsing/vernacextend.ml4
index e8a3094b9a..e1997b878b 100644
--- a/parsing/vernacextend.ml4
+++ b/parsing/vernacextend.ml4
@@ -48,9 +48,10 @@ let make_clauses s l =
let mlexpr_of_clause =
mlexpr_of_list
- (fun (a,b,c) -> mlexpr_of_list make_prod_item (GramTerminal a::b))
+ (fun (a,b,c) -> mlexpr_of_list make_prod_item
+ (Option.List.cons (Option.map (fun a -> GramTerminal a) a) b))
-let declare_command loc s cl =
+let declare_command loc s nt cl =
let gl = mlexpr_of_clause cl in
let icl = make_clauses s cl in
<:str_item<
@@ -59,7 +60,7 @@ let declare_command loc s cl =
open Extrawit;
try Vernacinterp.vinterp_add $mlexpr_of_string s$ (fun [ $list:icl$ ])
with e -> Pp.pp (Cerrors.explain_exn e);
- Egrammar.extend_vernac_command_grammar $mlexpr_of_string s$ $gl$;
+ Egrammar.extend_vernac_command_grammar $mlexpr_of_string s$ $nt$ $gl$;
end
>>
@@ -71,13 +72,22 @@ EXTEND
[ [ "VERNAC"; "COMMAND"; "EXTEND"; s = UIDENT;
OPT "|"; l = LIST1 rule SEP "|";
"END" ->
- declare_command loc s l ] ]
+ declare_command loc s <:expr<None>> l
+ | "VERNAC"; nt = LIDENT ; "EXTEND"; s = UIDENT;
+ OPT "|"; l = LIST1 rule SEP "|";
+ "END" ->
+ declare_command loc s <:expr<Some $lid:nt$>> l ] ]
;
+ (* spiwack: comment-by-guessing: it seems that the isolated string (which
+ otherwise could have been another argument) is not passed to the
+ VernacExtend interpreter function to discriminate between the clauses. *)
rule:
[ [ "["; s = STRING; l = LIST0 args; "]"; "->"; "["; e = Pcaml.expr; "]"
->
if s = "" then Util.user_err_loc (loc,"",Pp.str"Command name is empty.");
- (s,l,<:expr< fun () -> $e$ >>)
+ (Some s,l,<:expr< fun () -> $e$ >>)
+ | "[" ; "-" ; l = LIST1 args ; "]" ; "->" ; "[" ; e = Pcaml.expr ; "]" ->
+ (None,l,<:expr< fun () -> $e$ >>)
] ]
;
args:
diff --git a/plugins/cc/ccalgo.ml b/plugins/cc/ccalgo.ml
index 9cc6f9de93..9b5fca4da1 100644
--- a/plugins/cc/ccalgo.ml
+++ b/plugins/cc/ccalgo.ml
@@ -668,11 +668,11 @@ let __eps__ = id_of_string "_eps_"
let new_state_var typ state =
let id = pf_get_new_id __eps__ state.gls in
- state.gls<-
- {state.gls with it =
- {state.gls.it with evar_hyps =
- Environ.push_named_context_val (id,None,typ)
- state.gls.it.evar_hyps}};
+ let {it=gl ; sigma=sigma} = state.gls in
+ let new_hyps =
+ Environ.push_named_context_val (id,None,typ) (Goal.V82.hyps sigma gl) in
+ let gls = Goal.V82.new_goal_with sigma gl new_hyps in
+ state.gls<- gls;
id
let complete_one_class state i=
diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml
index 4e6ea8022e..ccfa2a0a70 100644
--- a/plugins/cc/cctac.ml
+++ b/plugins/cc/cctac.ml
@@ -204,9 +204,9 @@ let rec make_prb gls depth additionnal_terms =
neg_hyps:=(cid,nh):: !neg_hyps
| `Rule patts -> add_quant state id true patts
| `Nrule patts -> add_quant state id false patts
- end) (Environ.named_context_of_val gls.it.evar_hyps);
+ end) (Environ.named_context_of_val (Goal.V82.hyps gls.sigma gls.it));
begin
- match atom_of_constr env sigma gls.it.evar_concl with
+ match atom_of_constr env sigma (pf_concl gls) with
`Eq (t,a,b) -> add_disequality state Goal a b
| `Other g ->
List.iter
diff --git a/proofs/decl_expr.mli b/plugins/decl_mode/decl_expr.mli
index 20a95dabff..20a95dabff 100644
--- a/proofs/decl_expr.mli
+++ b/plugins/decl_mode/decl_expr.mli
diff --git a/tactics/decl_interp.ml b/plugins/decl_mode/decl_interp.ml
index 2b583af407..2b583af407 100644
--- a/tactics/decl_interp.ml
+++ b/plugins/decl_mode/decl_interp.ml
diff --git a/tactics/decl_interp.mli b/plugins/decl_mode/decl_interp.mli
index bd0859382b..bd0859382b 100644
--- a/tactics/decl_interp.mli
+++ b/plugins/decl_mode/decl_interp.mli
diff --git a/proofs/decl_mode.ml b/plugins/decl_mode/decl_mode.ml
index ba675327c6..90b0085ffa 100644
--- a/proofs/decl_mode.ml
+++ b/plugins/decl_mode/decl_mode.ml
@@ -13,32 +13,17 @@ open Term
open Evd
open Util
+
let daimon_flag = ref false
let set_daimon_flag () = daimon_flag:=true
let clear_daimon_flag () = daimon_flag:=false
let get_daimon_flag () = !daimon_flag
-type command_mode =
- Mode_tactic
- | Mode_proof
- | Mode_none
-let mode_of_pftreestate pts =
- let goal = sig_it (Refiner.top_goal_of_pftreestate pts) in
- if goal.evar_extra = None then
- Mode_tactic
- else
- Mode_proof
-let get_current_mode () =
- try
- mode_of_pftreestate (Pfedit.get_pftreestate ())
- with _ -> Mode_none
-
-let check_not_proof_mode str =
- if get_current_mode () = Mode_proof then
- error str
+(* Information associated to goals. *)
+open Store.Field
type split_tree=
Skip_patt of Idset.t * split_tree
@@ -72,53 +57,51 @@ type stack_info =
type pm_info =
{ pm_stack : stack_info list}
+let info = Store.field ()
+
+
+(* Current proof mode *)
+
+type command_mode =
+ Mode_tactic
+ | Mode_proof
+ | Mode_none
+
+let mode_of_pftreestate pts =
+ (* spiwack: it was "top_goal_..." but this should be fine *)
+ let { it = goals ; sigma = sigma } = Proof.V82.subgoals pts in
+ let goal = List.hd goals in
+ if info.get (Goal.V82.extra sigma goal) = None then
+ Mode_tactic
+ else
+ Mode_proof
+
+let get_current_mode () =
+ try
+ mode_of_pftreestate (Pfedit.get_pftreestate ())
+ with _ -> Mode_none
+
+let check_not_proof_mode str =
+ if get_current_mode () = Mode_proof then
+ error str
-let pm_in,pm_out = Dyn.create "pm_info"
+let get_info sigma gl=
+ match info.get (Goal.V82.extra sigma gl) with
+ | None -> invalid_arg "get_info"
+ | Some pm -> pm
-let get_info gl=
- match gl.evar_extra with
- None -> invalid_arg "get_info"
- | Some extra ->
- try pm_out extra with _ -> invalid_arg "get_info"
+let try_get_info sigma gl =
+ info.get (Goal.V82.extra sigma gl)
let get_stack pts =
- let info = get_info (sig_it (Refiner.nth_goal_of_pftreestate 1 pts)) in
- info.pm_stack
+ let { it = goals ; sigma = sigma } = Proof.V82.subgoals pts in
+ let info = get_info sigma (List.hd goals) in
+ info.pm_stack
let get_top_stack pts =
- let info = get_info (sig_it (Refiner.top_goal_of_pftreestate pts)) in
- info.pm_stack
-
-let get_end_command pts =
- match mode_of_pftreestate pts with
- Mode_proof ->
- Some
- begin
- match get_top_stack pts with
- [] -> "\"end proof\""
- | Claim::_ -> "\"end claim\""
- | Focus_claim::_-> "\"end focus\""
- | (Suppose_case :: Per (et,_,_,_) :: _
- | Per (et,_,_,_) :: _ ) ->
- begin
- match et with
- Decl_expr.ET_Case_analysis ->
- "\"end cases\" or start a new case"
- | Decl_expr.ET_Induction ->
- "\"end induction\" or start a new case"
- end
- | _ -> anomaly "lonely suppose"
- end
- | Mode_tactic ->
- begin
- try
- ignore
- (Refiner.up_until_matching_rule Proof_trees.is_proof_instr pts);
- Some "\"return\""
- with Not_found -> None
- end
- | Mode_none ->
- error "no proof in progress"
+ let { it = gl ; sigma = sigma } = Proof.V82.top_goal pts in
+ let info = get_info sigma gl in
+ info.pm_stack
let get_last env =
try
diff --git a/proofs/decl_mode.mli b/plugins/decl_mode/decl_mode.mli
index 734dc04421..fe485ce9b6 100644
--- a/proofs/decl_mode.mli
+++ b/plugins/decl_mode/decl_mode.mli
@@ -22,7 +22,7 @@ type command_mode =
| Mode_proof
| Mode_none
-val mode_of_pftreestate : pftreestate -> command_mode
+val mode_of_pftreestate : Proof.proof -> command_mode
val get_current_mode : unit -> command_mode
@@ -61,14 +61,14 @@ type stack_info =
type pm_info =
{pm_stack : stack_info list }
-val pm_in : pm_info -> Dyn.t
+val info : pm_info Store.Field.t
-val get_info : Proof_type.goal -> pm_info
+val get_info : Evd.evar_map -> Proof_type.goal -> pm_info
-val get_end_command : pftreestate -> string option
+val try_get_info : Evd.evar_map -> Proof_type.goal -> pm_info option
-val get_stack : pftreestate -> stack_info list
+val get_stack : Proof.proof -> stack_info list
-val get_top_stack : pftreestate -> stack_info list
+val get_top_stack : Proof.proof -> stack_info list
val get_last: Environ.env -> identifier
diff --git a/plugins/decl_mode/decl_mode_plugin.mllib b/plugins/decl_mode/decl_mode_plugin.mllib
new file mode 100644
index 0000000000..dce989bbcd
--- /dev/null
+++ b/plugins/decl_mode/decl_mode_plugin.mllib
@@ -0,0 +1,7 @@
+Decl_expr
+Decl_mode
+Decl_interp
+Decl_proof_instr
+Ppdecl_proof
+G_decl_mode
+Decl_mode_plugin_mod
diff --git a/tactics/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml
index 9c58f06ee9..cd695f239d 100644
--- a/tactics/decl_proof_instr.ml
+++ b/plugins/decl_mode/decl_proof_instr.ml
@@ -14,7 +14,6 @@ open Evd
open Refiner
open Proof_type
-open Proof_trees
open Tacmach
open Tacinterp
open Decl_expr
@@ -35,7 +34,7 @@ open Goptions
(* Strictness option *)
-let get_its_info gls = get_info gls.it
+let get_its_info gls = get_info gls.sigma gls.it
let get_strictness,set_strictness =
let strictness = ref false in
@@ -51,18 +50,20 @@ let _ =
let tcl_change_info_gen info_gen =
(fun gls ->
- let gl =sig_it gls in
- {it=[{gl with evar_extra=info_gen}];sigma=sig_sig gls},
- function
- [pftree] ->
- {pftree with
- goal=gl;
- ref=Some (Prim Change_evars,[pftree])}
- | _ -> anomaly "change_info : Wrong number of subtrees")
+ let concl = pf_concl gls in
+ let hyps = Goal.V82.hyps (project gls) (sig_it gls) in
+ let extra = Goal.V82.extra (project gls) (sig_it gls) in
+ let (gl,ev,sigma) = Goal.V82.mk_goal (project gls) hyps concl (info_gen extra) in
+ let sigma = Goal.V82.partial_solution sigma (sig_it gls) ev in
+ { it = [gl] ; sigma= sigma } )
+
+open Store.Field
-let tcl_change_info info gls = tcl_change_info_gen (Some (pm_in info)) gls
+let tcl_change_info info gls =
+ let info_gen = Decl_mode.info.set info in
+ tcl_change_info_gen info_gen gls
-let tcl_erase_info gls = tcl_change_info_gen None gls
+let tcl_erase_info gls = tcl_change_info_gen (Decl_mode.info.remove) gls
let special_whd gl=
let infos=Closure.create_clos_infos Closure.betadeltaiota (pf_env gl) in
@@ -77,7 +78,7 @@ let is_good_inductive env ind =
oib.mind_nrealargs = 0 && not (Inductiveops.mis_is_recursive (ind,mib,oib))
let check_not_per pts =
- if not (Proof_trees.is_complete_proof (proof_of_pftreestate pts)) then
+ if not (Proof.is_done pts) then
match get_stack pts with
Per (_,_,_,_)::_ ->
error "You are inside a proof per cases/induction.\n\
@@ -112,32 +113,23 @@ let assert_postpone id t =
(* start a proof *)
+let proof_focus = Proof.new_focus_kind ()
+let proof_cond = Proof.no_cond proof_focus
+
let start_proof_tac gls=
- let gl=sig_it gls in
let info={pm_stack=[]} in
- {it=[{gl with evar_extra=Some (pm_in info)}];sigma=sig_sig gls},
- function
- [pftree] ->
- {pftree with
- goal=gl;
- ref=Some (Decl_proof true,[pftree])}
- | _ -> anomaly "Dem : Wrong number of subtrees"
+ tcl_change_info info gls
let go_to_proof_mode () =
- Pfedit.mutate
- (fun pts -> nth_unproven 1 (solve_pftreestate start_proof_tac pts))
+ Pfedit.by start_proof_tac ;
+ let p = Proof_global.give_me_the_proof () in
+ Proof.focus proof_cond 1 p
(* closing gaps *)
let daimon_tac gls =
set_daimon_flag ();
- ({it=[];sigma=sig_sig gls},
- function
- [] ->
- {open_subgoals=0;
- goal=sig_it gls;
- ref=Some (Daimon,[])}
- | _ -> anomaly "Daimon: Wrong number of subtrees")
+ {it=[];sigma=sig_sig gls}
let daimon _ pftree =
set_daimon_flag ();
@@ -145,7 +137,8 @@ let daimon _ pftree =
open_subgoals=0;
ref=Some (Daimon,[])}
-let daimon_subtree = map_pftreestate (fun _ -> frontier_mapi daimon )
+let daimon_subtree =
+ fun _ -> Util.anomaly "Todo: Decl_proof_instr.daimon_subtree"
(* marking closed blocks *)
@@ -159,32 +152,17 @@ let mark_rule_as_done = function
Decl_proof true -> Decl_proof false
| Decl_proof false ->
anomaly "already marked as done"
- | Nested(Proof_instr (lock_focus,instr),spfl) ->
- if lock_focus then
- Nested(Proof_instr (false,instr),spfl)
- else
- anomaly "already marked as done"
| _ -> anomaly "mark_rule_as_done"
-let mark_proof_tree_as_done pt =
- match pt.ref with
- None -> anomaly "mark_proof_tree_as_done"
- | Some (r,spfl) ->
- {pt with ref= Some (mark_rule_as_done r,spfl)}
-
-let mark_as_done pts =
- map_pftreestate
- (fun _ -> mark_proof_tree_as_done)
- (up_to_matching_rule is_focussing_command pts)
(* post-instruction focus management *)
-let goto_current_focus pts = up_until_matching_rule is_focussing_command pts
+let goto_current_focus pts =
+ Proof.unfocus proof_focus pts
let goto_current_focus_or_top pts =
- try
- up_until_matching_rule is_focussing_command pts
- with Not_found -> top_of_tree pts
+ try goto_current_focus pts
+ with Util.UserError _ -> ()
(* return *)
@@ -194,22 +172,21 @@ let close_tactic_mode pts =
with Not_found ->
error "\"return\" cannot be used outside of Declarative Proof Mode." in
let pts2 = daimon_subtree pts1 in
- let pts3 = mark_as_done pts2 in
- goto_current_focus pts3
+ goto_current_focus pts2
-let return_from_tactic_mode () = Pfedit.mutate close_tactic_mode
+let return_from_tactic_mode () =
+ Util.anomaly "Todo: Decl_proof_instr.return_from_tactic_mode"
(* end proof/claim *)
let close_block bt pts =
- let stack =
- if Proof_trees.is_complete_proof (proof_of_pftreestate pts) then
- get_top_stack pts
+ if Proof.no_focused_goal pts then
+ goto_current_focus pts
else
- get_stack pts in
- match bt,stack with
+ let stack =get_stack pts in
+ match bt,stack with
B_claim, Claim::_ | B_focus, Focus_claim::_ | B_proof, [] ->
- daimon_subtree (goto_current_focus pts)
+ (goto_current_focus pts)
| _, Claim::_ ->
error "\"end claim\" expected."
| _, Focus_claim::_ ->
@@ -223,23 +200,24 @@ let close_block bt pts =
| ET_Induction -> error "\"end induction\" expected."
end
| _,_ -> anomaly "Lonely suppose on stack."
+
(* utility for suppose / suppose it is *)
let close_previous_case pts =
if
- Proof_trees.is_complete_proof (proof_of_pftreestate pts)
+ Proof.is_done pts
then
match get_top_stack pts with
Per (et,_,_,_) :: _ -> anomaly "Weird case occured ..."
| Suppose_case :: Per (et,_,_,_) :: _ ->
- goto_current_focus (mark_as_done pts)
+ goto_current_focus (pts)
| _ -> error "Not inside a proof per cases or induction."
else
match get_stack pts with
- Per (et,_,_,_) :: _ -> pts
+ Per (et,_,_,_) :: _ -> ()
| Suppose_case :: Per (et,_,_,_) :: _ ->
- goto_current_focus (mark_as_done (daimon_subtree pts))
+ goto_current_focus ((pts))
| _ -> error "Not inside a proof per cases or induction."
(* Proof instructions *)
@@ -252,7 +230,7 @@ let filter_hyps f gls =
tclIDTAC
else
tclTRY (clear [id]) in
- tclMAP filter_aux (Environ.named_context_of_val gls.it.evar_hyps) gls
+ tclMAP filter_aux (pf_hyps gls) gls
let local_hyp_prefix = id_of_string "___"
@@ -1446,59 +1424,59 @@ let rec preprocess pts instr =
| Pconsider (_,_) | Pcast (_,_) | Pgiven _ | Ptake _
| Pdefine (_,_,_) | Pper _ | Prew _ ->
check_not_per pts;
- true,pts
+ true
| Pescape ->
check_not_per pts;
- true,pts
+ true
| Pcase _ | Psuppose _ | Pend (B_elim _) ->
- true,close_previous_case pts
+ close_previous_case pts ;
+ true
| Pend bt ->
- false,close_block bt pts
+ close_block bt pts ;
+ false
let rec postprocess pts instr =
match instr with
Phence i | Pthus i | Pthen i -> postprocess pts i
| Pcut _ | Psuffices _ | Passume _ | Plet _ | Pconsider (_,_) | Pcast (_,_)
- | Pgiven _ | Ptake _ | Pdefine (_,_,_) | Prew (_,_) -> pts
+ | Pgiven _ | Ptake _ | Pdefine (_,_,_) | Prew (_,_) -> ()
| Pclaim _ | Pfocus _ | Psuppose _ | Pcase _ | Pper _
- | Pescape -> nth_unproven 1 pts
+ | Pescape -> Proof.focus proof_cond 1 pts
| Pend (B_elim ET_Induction) ->
begin
- let pf = proof_of_pftreestate pts in
- let (pfterm,_) = extract_open_pftreestate pts in
- let env = Evd.evar_env (goal_of_proof pf) in
+ let pfterm = List.hd (Proof.partial_proof pts) in
+ let { it = gls ; sigma = sigma } = Proof.V82.subgoals pts in
+ let env = Goal.V82.env sigma (List.hd gls) in
try
Inductiveops.control_only_guard env pfterm;
- goto_current_focus_or_top (mark_as_done pts)
+ goto_current_focus_or_top pts
with
Type_errors.TypeError(env,
Type_errors.IllFormedRecBody(_,_,_,_,_)) ->
anomaly "\"end induction\" generated an ill-formed fixpoint"
end
| Pend _ ->
- goto_current_focus_or_top (mark_as_done pts)
+ goto_current_focus_or_top (pts)
let do_instr raw_instr pts =
- let has_tactic,pts1 = preprocess pts raw_instr.instr in
- let pts2 =
+ let has_tactic = preprocess pts raw_instr.instr in
+ begin
if has_tactic then
- let gl = nth_goal_of_pftreestate 1 pts1 in
+ let { it=gls ; sigma=sigma } = Proof.V82.subgoals pts in
+ let gl = { it=List.hd gls ; sigma=sigma } in
let env= pf_env gl in
- let sigma= project gl in
let ist = {ltacvars = ([],[]); ltacrecvars = [];
gsigma = sigma; genv = env} in
let glob_instr = intern_proof_instr ist raw_instr in
let instr =
interp_proof_instr (get_its_info gl) sigma env glob_instr in
- let lock_focus = is_focussing_instr instr.instr in
- let marker= Proof_instr (lock_focus,instr) in
- solve_nth_pftreestate 1
- (abstract_operation marker (tclTHEN (eval_instr instr) clean_tmp)) pts1
- else pts1 in
- postprocess pts2 raw_instr.instr
+ Pfedit.by (tclTHEN (eval_instr instr) clean_tmp)
+ else () end;
+ postprocess pts raw_instr.instr
let proof_instr raw_instr =
- Pfedit.mutate (do_instr raw_instr)
+ let p = Proof_global.give_me_the_proof () in
+ do_instr raw_instr p
(*
diff --git a/tactics/decl_proof_instr.mli b/plugins/decl_mode/decl_proof_instr.mli
index 1cfcfedf14..0b22b7ac53 100644
--- a/tactics/decl_proof_instr.mli
+++ b/plugins/decl_mode/decl_proof_instr.mli
@@ -21,20 +21,16 @@ val register_automation_tac: tactic -> unit
val automation_tac : tactic
-val daimon_subtree: pftreestate -> pftreestate
+val daimon_subtree: Proof.proof -> Proof.proof
val concl_refiner:
Termops.meta_type_map -> constr -> Proof_type.goal sigma -> constr
-val do_instr: Decl_expr.raw_proof_instr -> pftreestate -> pftreestate
+val do_instr: Decl_expr.raw_proof_instr -> Proof.proof -> unit
val proof_instr: Decl_expr.raw_proof_instr -> unit
val tcl_change_info : Decl_mode.pm_info -> tactic
-val mark_proof_tree_as_done : Proof_type.proof_tree -> Proof_type.proof_tree
-
-val mark_as_done : pftreestate -> pftreestate
-
val execute_cases :
Names.name ->
Decl_mode.per_info ->
@@ -74,7 +70,7 @@ val register_dep_subcase :
val thesis_for : Term.constr ->
Term.constr -> Decl_mode.per_info -> Environ.env -> Term.constr
-val close_previous_case : pftreestate -> pftreestate
+val close_previous_case : Proof.proof -> unit
val pop_stacks :
(Names.identifier *
diff --git a/plugins/decl_mode/g_decl_mode.ml4 b/plugins/decl_mode/g_decl_mode.ml4
new file mode 100644
index 0000000000..cdfc57b374
--- /dev/null
+++ b/plugins/decl_mode/g_decl_mode.ml4
@@ -0,0 +1,407 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i camlp4deps: "parsing/grammar.cma" i*)
+(*i camlp4use: "pa_extend.cmo q_MLast.cmo" i*)
+
+(* $Id$ *)
+
+(* arnaud: veiller à l'aspect tutorial des commentaires *)
+
+open Pp
+
+open Decl_expr
+open Names
+open Term
+open Genarg
+open Pcoq
+
+open Pcoq.Constr
+open Pcoq.Tactic
+open Pcoq.Vernac_
+
+let pr_goal gs =
+ let (g,sigma) = Goal.V82.nf_evar (Tacmach.project gs) (Evd.sig_it gs) in
+ let env = Goal.V82.unfiltered_env sigma g in
+ let preamb,thesis,penv,pc =
+ (str " *** Declarative Mode ***" ++ fnl ()++fnl ()),
+ (str "thesis := " ++ fnl ()),
+ Printer.pr_context_of env,
+ Printer.pr_ltype_env_at_top env (Goal.V82.concl sigma g)
+ in
+ preamb ++
+ str" " ++ hv 0 (penv ++ fnl () ++
+ str (Printer.emacs_str (String.make 1 (Char.chr 253)) "") ++
+ str "============================" ++ fnl () ++
+ thesis ++ str " " ++ pc) ++ fnl ()
+
+(* arnaud: rebrancher ça
+let pr_open_subgoals () =
+ let p = Proof_global.give_me_the_proof () in
+ let { Evd.it = goals ; sigma = sigma } = Proof.V82.subgoals p in
+ let close_cmd = Decl_mode.get_end_command p in
+ pr_subgoals close_cmd sigma goals
+*)
+
+let pr_proof_instr instr =
+ Util.anomaly "Cannot print a proof_instr"
+ (* arnaud: Il nous faut quelque chose de type extr_genarg_printer si on veut aller
+ dans cette direction
+ Ppdecl_proof.pr_proof_instr (Global.env()) instr
+ *)
+let pr_raw_proof_instr instr =
+ Util.anomaly "Cannot print a raw proof_instr"
+let pr_glob_proof_instr instr =
+ Util.anomaly "Cannot print a non-interpreted proof_instr"
+
+let interp_proof_instr _ { Evd.it = gl ; sigma = sigma }=
+ Decl_interp.interp_proof_instr
+ (Decl_mode.get_info sigma gl)
+ (sigma)
+ (Goal.V82.env sigma gl)
+
+let vernac_decl_proof () =
+ let pf = Proof_global.give_me_the_proof () in
+ if Proof.is_done pf then
+ Util.error "Nothing left to prove here."
+ else
+ begin
+ Decl_proof_instr.go_to_proof_mode () ;
+ Proof_global.set_proof_mode "Declarative" ;
+ Vernacentries.print_subgoals ()
+ end
+
+(* spiwack: some bureaucracy is not performed here *)
+let vernac_return () =
+ Decl_proof_instr.return_from_tactic_mode () ;
+ Proof_global.set_proof_mode "Declarative" ;
+ Vernacentries.print_subgoals ()
+
+let vernac_proof_instr instr =
+ Decl_proof_instr.proof_instr instr;
+ Vernacentries.print_subgoals ()
+
+(* We create a new parser entry [proof_mode]. The Declarative proof mode
+ will replace the normal parser entry for tactics with this one. *)
+let proof_mode = Gram.Entry.create "vernac:proof_command"
+(* Auxiliary grammar entry. *)
+let proof_instr = Gram.Entry.create "proofmode:instr"
+
+(* Before we can write an new toplevel command (see below)
+ which takes a [proof_instr] as argument, we need to declare
+ how to parse it, print it, globalise it and interprete it.
+ Normally we could do that easily through ARGUMENT EXTEND,
+ but as the parsing is fairly complicated we will do it manually to
+ indirect through the [proof_instr] grammar entry. *)
+(* spiwack: proposal: doing that directly from argextend.ml4, maybe ? *)
+
+(* [Genarg.create_arg] creates a new embedding into Genarg. *)
+let (wit_proof_instr,globwit_proof_instr,rawwit_proof_instr) =
+ Genarg.create_arg "proof_instr"
+let _ = Tacinterp.add_interp_genarg "proof_instr"
+ begin
+ begin fun e x -> (* declares the globalisation function *)
+ Genarg.in_gen globwit_proof_instr
+ (Decl_interp.intern_proof_instr e (Genarg.out_gen rawwit_proof_instr x))
+ end,
+ begin fun ist gl x -> (* declares the interpretation function *)
+ Genarg.in_gen wit_proof_instr
+ (interp_proof_instr ist gl (Genarg.out_gen globwit_proof_instr x))
+ end,
+ begin fun _ x -> x end (* declares the substitution function, irrelevant in our case *)
+ end
+
+let _ = Pptactic.declare_extra_genarg_pprule
+ (rawwit_proof_instr, pr_raw_proof_instr)
+ (globwit_proof_instr, pr_glob_proof_instr)
+ (wit_proof_instr, pr_proof_instr)
+
+(* We use the VERNAC EXTEND facility with a custom non-terminal
+ to populate [proof_mode] with a new toplevel interpreter.
+ The "-" indicates that the rule does not start with a distinguished
+ string. *)
+VERNAC proof_mode EXTEND ProofInstr
+ [ - proof_instr(instr) ] -> [ vernac_proof_instr instr ]
+END
+
+(* It is useful to use GEXTEND directly to call grammar entries that have been
+ defined previously VERNAC EXTEND. In this case we allow, in proof mode,
+ the use of commands like Check or Print. VERNAC EXTEND does quite a bit of
+ bureaucracy for us, but it is not needed in this sort of case, and it would require
+ to have an ARGUMENT EXTEND version of the "proof_mode" grammar entry. *)
+GEXTEND Gram
+ GLOBAL: proof_mode ;
+
+ proof_mode: LAST
+ [ [ c=G_vernac.subgoal_command -> c (Some 1) None ] ]
+ ;
+END
+
+(* We register a new proof mode here *)
+
+let _ =
+ Proof_global.register_proof_mode { Proof_global.
+ name = "Declarative" ; (* name for identifying and printing *)
+ (* function [set] goes from No Proof Mode to
+ Declarative Proof Mode performing side effects *)
+ set = begin fun () ->
+ (* We set the command non terminal to
+ [proof_mode] (which we just defined). *)
+ G_vernac.set_command_entry proof_mode ;
+ (* We substitute the goal printer, by the one we built
+ for the proof mode. *)
+ Printer.set_printer_pr { Printer.default_printer_pr with
+ Printer.pr_goal = pr_goal }
+ end ;
+ (* function [reset] goes back to No Proof Mode from
+ Declarative Proof Mode *)
+ reset = begin fun () ->
+ (* We restore the command non terminal to
+ [noedit_mode]. *)
+ G_vernac.set_command_entry G_vernac.noedit_mode ;
+ (* We restore the goal printer to default *)
+ Printer.set_printer_pr Printer.default_printer_pr
+ end
+ }
+
+(* Two new vernacular commands *)
+VERNAC COMMAND EXTEND DeclProof
+ [ "proof" ] -> [ vernac_decl_proof () ]
+END
+VERNAC COMMAND EXTEND DeclReturn
+ [ "return" ] -> [ vernac_return () ]
+END
+
+let none_is_empty = function
+ None -> []
+ | Some l -> l
+
+GEXTEND Gram
+GLOBAL: proof_instr;
+ thesis :
+ [[ "thesis" -> Plain
+ | "thesis"; "for"; i=ident -> (For i)
+ ]];
+ statement :
+ [[ i=ident ; ":" ; c=constr -> {st_label=Name i;st_it=c}
+ | i=ident -> {st_label=Anonymous;
+ st_it=Topconstr.CRef (Libnames.Ident (loc, i))}
+ | c=constr -> {st_label=Anonymous;st_it=c}
+ ]];
+ constr_or_thesis :
+ [[ t=thesis -> Thesis t ] |
+ [ c=constr -> This c
+ ]];
+ statement_or_thesis :
+ [
+ [ t=thesis -> {st_label=Anonymous;st_it=Thesis t} ]
+ |
+ [ i=ident ; ":" ; cot=constr_or_thesis -> {st_label=Name i;st_it=cot}
+ | i=ident -> {st_label=Anonymous;
+ st_it=This (Topconstr.CRef (Libnames.Ident (loc, i)))}
+ | c=constr -> {st_label=Anonymous;st_it=This c}
+ ]
+ ];
+ justification_items :
+ [[ -> Some []
+ | "by"; l=LIST1 constr SEP "," -> Some l
+ | "by"; "*" -> None ]]
+ ;
+ justification_method :
+ [[ -> None
+ | "using"; tac = tactic -> Some tac ]]
+ ;
+ simple_cut_or_thesis :
+ [[ ls = statement_or_thesis;
+ j = justification_items;
+ taco = justification_method
+ -> {cut_stat=ls;cut_by=j;cut_using=taco} ]]
+ ;
+ simple_cut :
+ [[ ls = statement;
+ j = justification_items;
+ taco = justification_method
+ -> {cut_stat=ls;cut_by=j;cut_using=taco} ]]
+ ;
+ elim_type:
+ [[ IDENT "induction" -> ET_Induction
+ | IDENT "cases" -> ET_Case_analysis ]]
+ ;
+ block_type :
+ [[ IDENT "claim" -> B_claim
+ | IDENT "focus" -> B_focus
+ | IDENT "proof" -> B_proof
+ | et=elim_type -> B_elim et ]]
+ ;
+ elim_obj:
+ [[ IDENT "on"; c=constr -> Real c
+ | IDENT "of"; c=simple_cut -> Virtual c ]]
+ ;
+ elim_step:
+ [[ IDENT "consider" ;
+ h=consider_vars ; IDENT "from" ; c=constr -> Pconsider (c,h)
+ | IDENT "per"; et=elim_type; obj=elim_obj -> Pper (et,obj)
+ | IDENT "suffices"; ls=suff_clause;
+ j = justification_items;
+ taco = justification_method
+ -> Psuffices {cut_stat=ls;cut_by=j;cut_using=taco} ]]
+ ;
+ rew_step :
+ [[ "~=" ; c=simple_cut -> (Rhs,c)
+ | "=~" ; c=simple_cut -> (Lhs,c)]]
+ ;
+ cut_step:
+ [[ "then"; tt=elim_step -> Pthen tt
+ | "then"; c=simple_cut_or_thesis -> Pthen (Pcut c)
+ | IDENT "thus"; tt=rew_step -> Pthus (let s,c=tt in Prew (s,c))
+ | IDENT "thus"; c=simple_cut_or_thesis -> Pthus (Pcut c)
+ | IDENT "hence"; c=simple_cut_or_thesis -> Phence (Pcut c)
+ | tt=elim_step -> tt
+ | tt=rew_step -> let s,c=tt in Prew (s,c);
+ | IDENT "have"; c=simple_cut_or_thesis -> Pcut c;
+ | IDENT "claim"; c=statement -> Pclaim c;
+ | IDENT "focus"; IDENT "on"; c=statement -> Pfocus c;
+ | "end"; bt = block_type -> Pend bt;
+ | IDENT "escape" -> Pescape ]]
+ ;
+ (* examiner s'il est possible de faire R _ et _ R pour R une relation qcq*)
+ loc_id:
+ [[ id=ident -> fun x -> (loc,(id,x)) ]];
+ hyp:
+ [[ id=loc_id -> id None ;
+ | id=loc_id ; ":" ; c=constr -> id (Some c)]]
+ ;
+ consider_vars:
+ [[ name=hyp -> [Hvar name]
+ | name=hyp; ","; v=consider_vars -> (Hvar name) :: v
+ | name=hyp;
+ IDENT "such"; IDENT "that"; h=consider_hyps -> (Hvar name)::h
+ ]]
+ ;
+ consider_hyps:
+ [[ st=statement; IDENT "and"; h=consider_hyps -> Hprop st::h
+ | st=statement; IDENT "and";
+ IDENT "consider" ; v=consider_vars -> Hprop st::v
+ | st=statement -> [Hprop st]
+ ]]
+ ;
+ assume_vars:
+ [[ name=hyp -> [Hvar name]
+ | name=hyp; ","; v=assume_vars -> (Hvar name) :: v
+ | name=hyp;
+ IDENT "such"; IDENT "that"; h=assume_hyps -> (Hvar name)::h
+ ]]
+ ;
+ assume_hyps:
+ [[ st=statement; IDENT "and"; h=assume_hyps -> Hprop st::h
+ | st=statement; IDENT "and";
+ IDENT "we"; IDENT "have" ; v=assume_vars -> Hprop st::v
+ | st=statement -> [Hprop st]
+ ]]
+ ;
+ assume_clause:
+ [[ IDENT "we" ; IDENT "have" ; v=assume_vars -> v
+ | h=assume_hyps -> h ]]
+ ;
+ suff_vars:
+ [[ name=hyp; IDENT "to"; IDENT "show" ; c = constr_or_thesis ->
+ [Hvar name],c
+ | name=hyp; ","; v=suff_vars ->
+ let (q,c) = v in ((Hvar name) :: q),c
+ | name=hyp;
+ IDENT "such"; IDENT "that"; h=suff_hyps ->
+ let (q,c) = h in ((Hvar name) :: q),c
+ ]];
+ suff_hyps:
+ [[ st=statement; IDENT "and"; h=suff_hyps ->
+ let (q,c) = h in (Hprop st::q),c
+ | st=statement; IDENT "and";
+ IDENT "to" ; IDENT "have" ; v=suff_vars ->
+ let (q,c) = v in (Hprop st::q),c
+ | st=statement; IDENT "to"; IDENT "show" ; c = constr_or_thesis ->
+ [Hprop st],c
+ ]]
+ ;
+ suff_clause:
+ [[ IDENT "to" ; IDENT "have" ; v=suff_vars -> v
+ | h=suff_hyps -> h ]]
+ ;
+ let_vars:
+ [[ name=hyp -> [Hvar name]
+ | name=hyp; ","; v=let_vars -> (Hvar name) :: v
+ | name=hyp; IDENT "be";
+ IDENT "such"; IDENT "that"; h=let_hyps -> (Hvar name)::h
+ ]]
+ ;
+ let_hyps:
+ [[ st=statement; IDENT "and"; h=let_hyps -> Hprop st::h
+ | st=statement; IDENT "and"; "let"; v=let_vars -> Hprop st::v
+ | st=statement -> [Hprop st]
+ ]];
+ given_vars:
+ [[ name=hyp -> [Hvar name]
+ | name=hyp; ","; v=given_vars -> (Hvar name) :: v
+ | name=hyp; IDENT "such"; IDENT "that"; h=given_hyps -> (Hvar name)::h
+ ]]
+ ;
+ given_hyps:
+ [[ st=statement; IDENT "and"; h=given_hyps -> Hprop st::h
+ | st=statement; IDENT "and"; IDENT "given"; v=given_vars -> Hprop st::v
+ | st=statement -> [Hprop st]
+ ]];
+ suppose_vars:
+ [[name=hyp -> [Hvar name]
+ |name=hyp; ","; v=suppose_vars -> (Hvar name) :: v
+ |name=hyp; OPT[IDENT "be"];
+ IDENT "such"; IDENT "that"; h=suppose_hyps -> (Hvar name)::h
+ ]]
+ ;
+ suppose_hyps:
+ [[ st=statement_or_thesis; IDENT "and"; h=suppose_hyps -> Hprop st::h
+ | st=statement_or_thesis; IDENT "and"; IDENT "we"; IDENT "have";
+ v=suppose_vars -> Hprop st::v
+ | st=statement_or_thesis -> [Hprop st]
+ ]]
+ ;
+ suppose_clause:
+ [[ IDENT "we"; IDENT "have"; v=suppose_vars -> v;
+ | h=suppose_hyps -> h ]]
+ ;
+ intro_step:
+ [[ IDENT "suppose" ; h=assume_clause -> Psuppose h
+ | IDENT "suppose" ; IDENT "it"; IDENT "is" ; c=pattern LEVEL "0" ;
+ po=OPT[ "with"; p=LIST1 hyp SEP ","-> p ] ;
+ ho=OPT[ IDENT "and" ; h=suppose_clause -> h ] ->
+ Pcase (none_is_empty po,c,none_is_empty ho)
+ | "let" ; v=let_vars -> Plet v
+ | IDENT "take"; witnesses = LIST1 constr SEP "," -> Ptake witnesses
+ | IDENT "assume"; h=assume_clause -> Passume h
+ | IDENT "given"; h=given_vars -> Pgiven h
+ | IDENT "define"; id=ident; args=LIST0 hyp;
+ "as"; body=constr -> Pdefine(id,args,body)
+ | IDENT "reconsider"; id=ident; "as" ; typ=constr -> Pcast (This id,typ)
+ | IDENT "reconsider"; t=thesis; "as" ; typ=constr -> Pcast (Thesis t ,typ)
+ ]]
+ ;
+ emphasis :
+ [[ -> 0
+ | "*" -> 1
+ | "**" -> 2
+ | "***" -> 3
+ ]]
+ ;
+ bare_proof_instr:
+ [[ c = cut_step -> c ;
+ | i = intro_step -> i ]]
+ ;
+ proof_instr :
+ [[ e=emphasis;i=bare_proof_instr -> {emph=e;instr=i}]]
+ ;
+END;;
+
+
diff --git a/parsing/ppdecl_proof.ml b/plugins/decl_mode/ppdecl_proof.ml
index 40c712cdff..40c712cdff 100644
--- a/parsing/ppdecl_proof.ml
+++ b/plugins/decl_mode/ppdecl_proof.ml
diff --git a/parsing/ppdecl_proof.mli b/plugins/decl_mode/ppdecl_proof.mli
index fd6fb66376..fd6fb66376 100644
--- a/parsing/ppdecl_proof.mli
+++ b/plugins/decl_mode/ppdecl_proof.mli
diff --git a/plugins/firstorder/ground.ml b/plugins/firstorder/ground.ml
index 8a0f02d27e..7f4acb8564 100644
--- a/plugins/firstorder/ground.ml
+++ b/plugins/firstorder/ground.ml
@@ -18,32 +18,6 @@ open Tactics
open Tacticals
open Libnames
-(*
-let old_search=ref !Auto.searchtable
-
-(* I use this solution as a means to know whether hints have changed,
-but this prevents the GC from collecting the previous table,
-resulting in some limited space wasting*)
-
-let update_flags ()=
- if not ( !Auto.searchtable == !old_search ) then
- begin
- old_search:=!Auto.searchtable;
- let predref=ref Names.KNpred.empty in
- let f p_a_t =
- match p_a_t.Auto.code with
- Auto.Unfold_nth (ConstRef kn)->
- predref:=Names.KNpred.add kn !predref
- | _ ->() in
- let g _ l=List.iter f l in
- let h _ hdb=Auto.Hint_db.iter g hdb in
- Util.Stringmap.iter h !Auto.searchtable;
- red_flags:=
- Closure.RedFlags.red_add_transparent
- Closure.betaiotazeta (Names.Idpred.full,!predref)
- end
-*)
-
let update_flags ()=
let predref=ref Names.Cpred.empty in
let f coe=
@@ -61,7 +35,7 @@ let ground_tac solver startseq gl=
update_flags ();
let rec toptac skipped seq gl=
if Tacinterp.get_debug()=Tactic_debug.DebugOn 0
- then Pp.msgnl (Printer.pr_goal (sig_it gl));
+ then Pp.msgnl (Printer.pr_goal gl);
tclORELSE (axiom_tac seq.gl seq)
begin
try
diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml
index e2cad94495..bef89909de 100644
--- a/plugins/funind/functional_principles_proofs.ml
+++ b/plugins/funind/functional_principles_proofs.ml
@@ -35,7 +35,7 @@ let observennl strm =
let do_observe_tac s tac g =
try let v = tac g in (* msgnl (goal ++ fnl () ++ (str s)++(str " ")++(str "finished")); *) v
with e ->
- let goal = begin try (Printer.pr_goal (sig_it g)) with _ -> assert false end in
+ let goal = begin try (Printer.pr_goal g) with _ -> assert false end in
msgnl (str "observation "++ s++str " raised exception " ++
Cerrors.explain_exn e ++ str " on goal " ++ goal );
raise e;;
diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml
index 0f048f59a0..ff7089613f 100644
--- a/plugins/funind/indfun_common.ml
+++ b/plugins/funind/indfun_common.ml
@@ -180,48 +180,9 @@ let save with_clean id const (locality,kind) hook =
-
-let extract_pftreestate pts =
- let pfterm,subgoals = Refiner.extract_open_pftreestate pts in
- let tpfsigma = Refiner.evc_of_pftreestate pts in
- let exl = Evarutil.non_instantiated tpfsigma in
- if subgoals <> [] or exl <> [] then
- Util.errorlabstrm "extract_proof"
- (if subgoals <> [] then
- str "Attempt to save an incomplete proof"
- else
- str "Attempt to save a proof with existential variables still non-instantiated");
- let env = Global.env_of_context (Refiner.proof_of_pftreestate pts).Proof_type.goal.Evd.evar_hyps in
- env,tpfsigma,pfterm
-
-
-let nf_betaiotazeta =
- let clos_norm_flags flgs env sigma t =
- Closure.norm_val (Closure.create_clos_infos flgs env) (Closure.inject (Reductionops.nf_evar sigma t)) in
- clos_norm_flags Closure.betaiotazeta
-
-let nf_betaiota =
- let clos_norm_flags flgs env sigma t =
- Closure.norm_val (Closure.create_clos_infos flgs env) (Closure.inject (Reductionops.nf_evar sigma t)) in
- clos_norm_flags Closure.betaiota
-
-let cook_proof do_reduce =
- let pfs = Pfedit.get_pftreestate ()
-(* and ident = Pfedit.get_current_proof_name () *)
- and (ident,strength,concl,hook) = Pfedit.current_proof_statement () in
- let env,sigma,pfterm = extract_pftreestate pfs in
- let pfterm =
- if do_reduce
- then nf_betaiota env sigma pfterm
- else pfterm
- in
- (ident,
- ({ const_entry_body = pfterm;
- const_entry_type = Some concl;
- const_entry_opaque = false;
- const_entry_boxed = false},
- strength, hook))
-
+let cook_proof _ =
+ let (id,(entry,_,strength,hook)) = Pfedit.cook_proof (fun _ -> ()) in
+ (id,(entry,strength,hook))
let new_save_named opacity =
let id,(const,persistence,hook) = cook_proof true in
diff --git a/plugins/funind/indfun_common.mli b/plugins/funind/indfun_common.mli
index 6f6607fccc..c48dff0c63 100644
--- a/plugins/funind/indfun_common.mli
+++ b/plugins/funind/indfun_common.mli
@@ -50,15 +50,8 @@ val jmeq_refl : unit -> Term.constr
(* [save_named] is a copy of [Command.save_named] but uses
[nf_betaiotazeta] instead of [nf_betaiotaevar_preserving_vm_cast]
-
-
-
- DON'T USE IT if you cannot ensure that there is no VMcast in the proof
-
*)
-(* val nf_betaiotazeta : Reductionops.reduction_function *)
-
val new_save_named : bool -> unit
val save : bool -> identifier -> Entries.definition_entry -> Decl_kinds.goal_kind ->
diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml
index 8c22265d66..28de815ca0 100644
--- a/plugins/funind/invfun.ml
+++ b/plugins/funind/invfun.ml
@@ -60,7 +60,7 @@ let observennl strm =
let do_observe_tac s tac g =
- let goal = begin try (Printer.pr_goal (sig_it g)) with _ -> assert false end in
+ let goal = begin try (Printer.pr_goal g) with _ -> assert false end in
try
let v = tac g in msgnl (goal ++ fnl () ++ s ++(str " ")++(str "finished")); v
with e ->
diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml
index 1eae097189..8dc660b421 100644
--- a/plugins/funind/recdef.ml
+++ b/plugins/funind/recdef.ml
@@ -85,7 +85,7 @@ let rec print_debug_queue e =
let do_observe_tac s tac g =
- let goal = Printer.pr_goal (sig_it g) in
+ let goal = Printer.pr_goal g in
let lmsg = (str "recdef ") ++ (str s) in
Queue.add (lmsg,goal) debug_queue;
try
@@ -376,7 +376,7 @@ let rec mk_intros_and_continue thin_intros (extra_eqn:bool)
(fun g1 ->
let ty_teq = pf_type_of g1 (mkVar teq) in
let teq_lhs,teq_rhs =
- let _,args = try destApp ty_teq with _ -> Pp.msgnl (Printer.pr_goal (sig_it g1) ++ fnl () ++ pr_id teq ++ str ":" ++ Printer.pr_lconstr ty_teq); assert false in
+ let _,args = try destApp ty_teq with _ -> Pp.msgnl (Printer.pr_goal g1 ++ fnl () ++ pr_id teq ++ str ":" ++ Printer.pr_lconstr ty_teq); assert false in
args.(1),args.(2)
in
cont_function (mkVar teq::eqs) (replace_term teq_lhs teq_rhs expr) g1
@@ -884,9 +884,9 @@ let whole_start (concl_tac:tactic) nb_args is_mes func input_type relation rec_a
end
let get_current_subgoals_types () =
- let pts = get_pftreestate () in
- let _,subs = extract_open_pftreestate pts in
- List.map snd ((* List.sort (fun (x,_) (y,_) -> x -y ) *)subs )
+ let p = Proof_global.give_me_the_proof () in
+ let { Evd.it=sgs ; sigma=sigma } = Proof.V82.subgoals p in
+ List.map (Goal.V82.abstract_type sigma) sgs
let build_and_l l =
let and_constr = Coqlib.build_coq_and () in
diff --git a/plugins/pluginsbyte.itarget b/plugins/pluginsbyte.itarget
index 7ca8020dc1..c0237ecf70 100644
--- a/plugins/pluginsbyte.itarget
+++ b/plugins/pluginsbyte.itarget
@@ -21,3 +21,4 @@ syntax/r_syntax_plugin.cma
syntax/string_syntax_plugin.cma
syntax/z_syntax_plugin.cma
quote/quote_plugin.cma
+decl_mode/decl_mode_plugin.cma
diff --git a/plugins/pluginsopt.itarget b/plugins/pluginsopt.itarget
index 5206271157..26b3f90842 100644
--- a/plugins/pluginsopt.itarget
+++ b/plugins/pluginsopt.itarget
@@ -21,3 +21,4 @@ syntax/r_syntax_plugin.cmxa
syntax/string_syntax_plugin.cmxa
syntax/z_syntax_plugin.cmxa
quote/quote_plugin.cmxa
+decl_mode/decl_mode_plugin.cmxa
diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml
index 2e4d07d633..ce65a45ec8 100644
--- a/plugins/quote/quote.ml
+++ b/plugins/quote/quote.ml
@@ -111,7 +111,6 @@ open Pattern
open Matching
open Tacmach
open Tactics
-open Proof_trees
open Tacexpr
(*i*)
diff --git a/plugins/ring/ring.ml b/plugins/ring/ring.ml
index 1e3765da6c..fdf3a9bd9d 100644
--- a/plugins/ring/ring.ml
+++ b/plugins/ring/ring.ml
@@ -21,7 +21,6 @@ open Reductionops
open Tacticals
open Tacexpr
open Tacmach
-open Proof_trees
open Printer
open Equality
open Vernacinterp
diff --git a/plugins/rtauto/refl_tauto.ml b/plugins/rtauto/refl_tauto.ml
index 23cb07050a..06313e8fe7 100644
--- a/plugins/rtauto/refl_tauto.ml
+++ b/plugins/rtauto/refl_tauto.ml
@@ -267,14 +267,13 @@ open Pp
let rtauto_tac gls=
Coqlib.check_required_library ["Coq";"rtauto";"Rtauto"];
let gamma={next=1;env=[]} in
- let gl=gls.it.evar_concl in
+ let gl=pf_concl gls in
let _=
if Retyping.get_sort_family_of
(pf_env gls) (Tacmach.project gls) gl <> InProp
then errorlabstrm "rtauto" (Pp.str "goal should be in Prop") in
let glf=make_form gamma gls gl in
- let hyps=make_hyps gamma gls [gl]
- (Environ.named_context_of_val gls.it.evar_hyps) in
+ let hyps=make_hyps gamma gls [gl] (pf_hyps gls) in
let formula=
List.fold_left (fun gl (_,f)-> Arrow (f,gl)) glf hyps in
let search_fun =
diff --git a/plugins/setoid_ring/newring.ml4 b/plugins/setoid_ring/newring.ml4
index 490ae89db0..d6dece458b 100644
--- a/plugins/setoid_ring/newring.ml4
+++ b/plugins/setoid_ring/newring.ml4
@@ -188,8 +188,10 @@ let ltac_record flds =
let carg c = TacDynamic(dummy_loc,Pretyping.constr_in c)
let dummy_goal env =
- {Evd.it = Evd.make_evar (named_context_val env) mkProp;
- Evd.sigma = Evd.empty}
+ let (gl,_,sigma) =
+ Goal.V82.mk_goal Evd.empty (named_context_val env) mkProp Store.empty in
+ {Evd.it = gl;
+ Evd.sigma = sigma}
let exec_tactic env n f args =
let lid = list_tabulate(fun i -> id_of_string("x"^string_of_int i)) n in
diff --git a/plugins/subtac/eterm.ml b/plugins/subtac/eterm.ml
index f1bdd64092..bfa2d47413 100644
--- a/plugins/subtac/eterm.ml
+++ b/plugins/subtac/eterm.ml
@@ -33,6 +33,10 @@ type oblinfo =
ev_tac: tactic option;
ev_deps: Intset.t }
+(* spiwack: Store field for internalizing ev_tac in evar_infos' evar_extra. *)
+open Store.Field
+let evar_tactic = Store.field ()
+
(** Substitute evar references in t using De Bruijn indices,
where n binders were passed through. *)
@@ -210,7 +214,7 @@ let eterm_obligations env name isevars evm fs ?status t ty =
| Some s -> s, None
| None -> Define true, None
in
- let tac = match ev.evar_extra with
+ let tac = match evar_tactic.get ev.evar_extra with
| Some t ->
if Dyn.tag t = "tactic" then
Some (Tacinterp.interp
diff --git a/plugins/subtac/subtac.ml b/plugins/subtac/subtac.ml
index 0eba0f6336..e649ad97be 100644
--- a/plugins/subtac/subtac.ml
+++ b/plugins/subtac/subtac.ml
@@ -50,7 +50,7 @@ open Tacinterp
open Tacexpr
let solve_tccs_in_type env id isevars evm c typ =
- if not (evm = Evd.empty) then
+ if not (Evd.is_empty evm) then
let stmt_id = Nameops.add_suffix id "_stmt" in
let obls, _, c', t' = eterm_obligations env stmt_id !isevars evm 0 ~status:Expand c typ in
match Subtac_obligations.add_definition stmt_id ~term:c' typ obls with
@@ -246,5 +246,5 @@ let subtac (loc, command) =
raise e)
| e ->
- msg_warning (str "Uncatched exception: " ++ Cerrors.explain_exn e);
+ msg_warning (str "Uncaught exception: " ++ Cerrors.explain_exn e);
raise e
diff --git a/plugins/subtac/subtac_obligations.ml b/plugins/subtac/subtac_obligations.ml
index 1424618f00..181d0fa72d 100644
--- a/plugins/subtac/subtac_obligations.ml
+++ b/plugins/subtac/subtac_obligations.ml
@@ -489,6 +489,7 @@ and solve_obligation_by_tac prg obls i tac =
| Stdpp.Exc_located(_, Refiner.FailError (_, s))
| Refiner.FailError (_, s) ->
user_err_loc (obl.obl_location, "solve_obligation", Lazy.force s)
+ | Util.Anomaly _ as e -> raise e
| e -> false
and solve_prg_obligations prg tac =
diff --git a/plugins/xml/dumptree.ml4 b/plugins/xml/dumptree.ml4
index 9419ba597e..d20f9f9b67 100644
--- a/plugins/xml/dumptree.ml4
+++ b/plugins/xml/dumptree.ml4
@@ -56,13 +56,11 @@ let pr_rule_xml pr = function
hov 2 (str "<cmpdrule>" ++ fnl () ++
begin match cmpd with
Tactic (texp, _) -> pr_tactic_xml texp
- | Proof_instr (_,instr) -> pr_proof_instr_xml instr
end ++ fnl ()
++ pr subtree
) ++ fnl () ++ str "</cmpdrule>"
| Daimon -> str "<daimon/>"
| Decl_proof _ -> str "<proof/>"
-(* | Change_evars -> str "<chgevars/>"*)
;;
let pr_var_decl_xml env (id,c,typ) =
@@ -115,11 +113,11 @@ let pr_subgoal_metas_xml metas env=
List.fold_left (++) (mt ()) (List.map pr_one metas)
;;
-let pr_goal_xml g =
- let env = try evar_unfiltered_env g with _ -> empty_env in
- if g.evar_extra = None then
+let pr_goal_xml sigma g =
+ let env = try Goal.V82.unfiltered_env sigma g with _ -> empty_env in
+ if Decl_mode.try_get_info sigma g = None then
(hov 2 (str "<goal>" ++ fnl () ++ str "<concl type=\"" ++
- xmlstream (pr_ltype_env_at_top env g.evar_concl) ++
+ xmlstream (pr_ltype_env_at_top env (Goal.V82.concl sigma g)) ++
str "\"/>" ++
(pr_context_xml env)) ++
fnl () ++ str "</goal>")
@@ -129,23 +127,9 @@ let pr_goal_xml g =
fnl () ++ str "</goal>")
;;
-let rec print_proof_xml sigma osign pf =
- let hyps = Environ.named_context_of_val pf.goal.evar_hyps in
- let hyps' = thin_sign osign hyps in
- match pf.ref with
- | None -> hov 2 (str "<tree>" ++ fnl () ++ (pr_goal_xml {pf.goal with evar_hyps=hyps'})) ++ fnl () ++ str "</tree>"
- | Some(r,spfl) ->
- hov 2 (str "<tree>" ++ fnl () ++
- (pr_goal_xml {pf.goal with evar_hyps=hyps'}) ++ fnl () ++ (pr_rule_xml (print_proof_xml sigma osign) r) ++
- (List.fold_left (fun x y -> x ++ fnl () ++ y) (mt ()) (List.map (print_proof_xml sigma hyps) spfl))) ++ fnl () ++ str "</tree>"
-;;
-
let print_proof_xml () =
- let pp = print_proof_xml Evd.empty Sign.empty_named_context
- (Tacmach.proof_of_pftreestate (Refiner.top_of_tree (Pfedit.get_pftreestate ())))
- in
- msgnl pp
-;;
+ Util.anomaly "Dump Tree command not supported in this version."
+
VERNAC COMMAND EXTEND DumpTree
[ "Dump" "Tree" ] -> [ print_proof_xml () ]
diff --git a/plugins/xml/proof2aproof.ml b/plugins/xml/proof2aproof.ml
index 1beabf26ca..c7b8b556ed 100644
--- a/plugins/xml/proof2aproof.ml
+++ b/plugins/xml/proof2aproof.ml
@@ -94,83 +94,9 @@ module ProofTreeHash =
let extract_open_proof sigma pf =
- let module PT = Proof_type in
- let module L = Logic in
- let evd = ref (Evd.create_evar_defs sigma) in
- let proof_tree_to_constr = ProofTreeHash.create 503 in
- let proof_tree_to_flattened_proof_tree = ProofTreeHash.create 503 in
- let unshared_constrs = ref S.empty in
- let rec proof_extractor vl node =
- let constr =
- match node with
- {PT.ref=Some(PT.Prim _,_)} as pf ->
- L.prim_extractor proof_extractor vl pf
-
- | {PT.ref=Some(PT.Nested (_,hidden_proof),spfl)} ->
- let sgl,v = Refiner.frontier hidden_proof in
- let flat_proof = v spfl in
- ProofTreeHash.add proof_tree_to_flattened_proof_tree node flat_proof ;
- proof_extractor vl flat_proof
-
- | {PT.ref=None;PT.goal=goal} ->
- let visible_rels =
- Util.map_succeed
- (fun id ->
- (* Section variables are in the [id] list but are not *)
- (* lambda abstracted in the term [vl] *)
- try let n = Logic.proof_variable_index id vl in (n,id)
- with Not_found -> failwith "caught")
-(*CSC: the above function must be modified such that when it is found *)
-(*CSC: it becomes a Rel; otherwise a Var. Then it can be already used *)
-(*CSC: as the evar_instance. Ordering the instance becomes useless (it *)
-(*CSC: will already be ordered. *)
- (Termops.ids_of_named_context
- (Environ.named_context_of_val goal.Evd.evar_hyps)) in
- let sorted_rels =
- Sort.list (fun (n1,_) (n2,_) -> n1 < n2 ) visible_rels in
- let context =
- let l =
- List.map
- (fun (_,id) -> Sign.lookup_named id
- (Environ.named_context_of_val goal.Evd.evar_hyps))
- sorted_rels in
- Environ.val_of_named_context l
- in
-(*CSC: the section variables in the right order must be added too *)
- let evar_instance = List.map (fun (n,_) -> Term.mkRel n) sorted_rels in
- (* let env = Global.env_of_context context in *)
- let evd',evar =
- Evarutil.new_evar_instance context !evd goal.Evd.evar_concl
- evar_instance in
- evd := evd' ;
- evar
-
- | _ -> Util.anomaly "Bug : a case has been forgotten in proof_extractor"
- in
- let unsharedconstr =
- let evar_nf_constr =
- nf_evar ( !evd)
- ~preserve:(function e -> S.mem e !unshared_constrs) constr
- in
- Unshare.unshare
- ~already_unshared:(function e -> S.mem e !unshared_constrs)
- evar_nf_constr
- in
-(*CSC: debugging stuff to be removed *)
-if ProofTreeHash.mem proof_tree_to_constr node then
- Pp.ppnl (Pp.(++) (Pp.str "#DUPLICATE INSERTION: ")
- (Tactic_printer.print_proof ( !evd) [] node)) ;
- ProofTreeHash.add proof_tree_to_constr node unsharedconstr ;
- unshared_constrs := S.add unsharedconstr !unshared_constrs ;
- unsharedconstr
- in
- let unshared_pf = unshare_proof_tree pf in
- let pfterm = proof_extractor [] unshared_pf in
- (pfterm, !evd, proof_tree_to_constr, proof_tree_to_flattened_proof_tree,
- unshared_pf)
-;;
+ (* Deactivated and candidate for removal. (Apr. 2010) *)
+ ()
let extract_open_pftreestate pts =
- extract_open_proof (Refiner.evc_of_pftreestate pts)
- (Tacmach.proof_of_pftreestate pts)
-;;
+ (* Deactivated and candidate for removal. (Apr. 2010) *)
+ ()
diff --git a/plugins/xml/proofTree2Xml.ml4 b/plugins/xml/proofTree2Xml.ml4
index 3f1e0a630b..dcfa997923 100644
--- a/plugins/xml/proofTree2Xml.ml4
+++ b/plugins/xml/proofTree2Xml.ml4
@@ -161,10 +161,12 @@ Pp.ppnl (Pp.(++) (Pp.str
let of_attribute = ("name",tacname)::("script",tac)::of_attribute in
(****** le but *)
- let {Evd.evar_concl=concl;
- Evd.evar_hyps=hyps}=goal in
+
+ let concl = Goal.V82.concl sigma goal in
+ let hyps = Goal.V82.hyps sigma goal in
let env = Global.env_of_context hyps in
+
let xgoal =
X.xml_nempty "Goal" [] (constr_to_xml concl sigma env) in
@@ -188,9 +190,6 @@ Pp.ppnl (Pp.(++) (Pp.str
[<(build_hyps new_hyps) ; (aux flat_proof nhyps)>]
end
- | {PT.ref=Some((PT.Nested(PT.Proof_instr (_,_),_)|PT.Decl_proof _),nodes)} ->
- Util.anomaly "Not Implemented"
-
| {PT.ref=Some(PT.Daimon,_)} ->
X.xml_empty "Hidden_open_goal" of_attribute
diff --git a/plugins/xml/xmlcommand.ml b/plugins/xml/xmlcommand.ml
index b3b2e26549..8a095bb245 100644
--- a/plugins/xml/xmlcommand.ml
+++ b/plugins/xml/xmlcommand.ml
@@ -551,34 +551,7 @@ let print_ref qid fn =
(* where dest is either None (for stdout) or (Some filename) *)
(* pretty prints via Xml.pp the proof in progress on dest *)
let show_pftreestate internal fn (kind,pftst) id =
- let pf = Tacmach.proof_of_pftreestate pftst in
- let typ = (Proof_trees.goal_of_proof pf).Evd.evar_concl in
- let val0,evar_map,proof_tree_to_constr,proof_tree_to_flattened_proof_tree,
- unshared_pf
- =
- Proof2aproof.extract_open_pftreestate pftst in
- let env = Global.env () in
- let obj =
- mk_current_proof_obj (fst kind = Decl_kinds.Local) id val0 typ evar_map env in
- let uri =
- match kind with
- Decl_kinds.Local, _ ->
- let uri =
- "cic:/" ^ String.concat "/"
- (Cic2acic.token_list_of_path (Lib.cwd ()) id Cic2acic.TVariable)
- in
- let kind_of_var = "VARIABLE","LocalFact" in
- if not internal then print_object_kind uri kind_of_var;
- uri
- | Decl_kinds.Global, _ ->
- let uri = Cic2acic.uri_of_declaration id Cic2acic.TConstant in
- if not internal then print_object_kind uri (kind_of_global_goal kind);
- uri
- in
- print_object uri obj evar_map
- (Some (Tacmach.evc_of_pftreestate pftst,unshared_pf,proof_tree_to_constr,
- proof_tree_to_flattened_proof_tree)) fn
-;;
+ Util.anomaly "Xmlcommand.show_pftreestate is not supported in this version."
let show fn =
let pftst = Pfedit.get_pftreestate () in
diff --git a/pretyping/evarutil.ml b/pretyping/evarutil.ml
index 24850b4baf..6d37cf80f5 100644
--- a/pretyping/evarutil.ml
+++ b/pretyping/evarutil.ml
@@ -150,12 +150,11 @@ let evars_to_metas sigma (emap, c) =
(* The list of non-instantiated existential declarations *)
let non_instantiated sigma =
- let listev = to_list sigma in
- List.fold_left
- (fun l (ev,evi) ->
- if evi.evar_body = Evar_empty then
- ((ev,nf_evar_info sigma evi)::l) else l)
- [] listev
+ List.rev begin
+ Evd.fold_undefined begin fun ev evi l ->
+ (ev,nf_evar_info sigma evi)::l
+ end sigma []
+ end
(**********************)
(* Creating new evars *)
@@ -478,6 +477,10 @@ type clear_dependency_error =
exception ClearDependencyError of identifier * clear_dependency_error
+open Store.Field
+
+let cleared = Store.field ()
+
let rec check_and_clear_in_constr evdref err ids c =
(* returns a new constr where all the evars have been 'cleaned'
(ie the hypotheses ids have been removed from the contexts of
@@ -539,6 +542,13 @@ let rec check_and_clear_in_constr evdref err ids c =
let ev'= e_new_evar evdref env ~src:(evar_source evk !evdref) nconcl in
evdref := Evd.define evk ev' !evdref;
let (evk',_) = destEvar ev' in
+ (* spiwack: hacking session to mark the old [evk] as having been "cleared" *)
+ let evi = Evd.find !evdref evk in
+ let extra = evi.evar_extra in
+ let extra' = cleared.set true extra in
+ let evi' = { evi with evar_extra = extra' } in
+ evdref := Evd.add !evdref evk evi' ;
+ (* spiwack: /hacking session *)
mkEvar(evk', Array.of_list nargs)
end
diff --git a/pretyping/evarutil.mli b/pretyping/evarutil.mli
index 205ca8bd64..39f8dd05ac 100644
--- a/pretyping/evarutil.mli
+++ b/pretyping/evarutil.mli
@@ -193,6 +193,10 @@ type clear_dependency_error =
exception ClearDependencyError of identifier * clear_dependency_error
+(* spiwack: marks an evar that has been "defined" by clear.
+ used by [Goal] and (indirectly) [Proofview] to handle the clear tactic gracefully*)
+val cleared : bool Store.Field.t
+
val clear_hyps_in_evi : evar_map ref -> named_context_val -> types ->
identifier list -> named_context_val * types
diff --git a/pretyping/evd.ml b/pretyping/evd.ml
index a9b42f052b..5b4a3f214d 100644
--- a/pretyping/evd.ml
+++ b/pretyping/evd.ml
@@ -51,7 +51,7 @@ type evar_info = {
evar_body : evar_body;
evar_filter : bool list;
evar_source : hole_kind located;
- evar_extra : Dyn.t option}
+ evar_extra : Store.t }
let make_evar hyps ccl = {
evar_concl = ccl;
@@ -59,7 +59,7 @@ let make_evar hyps ccl = {
evar_body = Evar_empty;
evar_filter = List.map (fun _ -> true) (named_context_of_val hyps);
evar_source = (dummy_loc,InternalHole);
- evar_extra = None
+ evar_extra = Store.empty
}
let evar_concl evi = evi.evar_concl
@@ -93,26 +93,31 @@ module ExistentialSet = Intset
(* This exception is raised by *.existential_value *)
exception NotInstantiatedEvar
+
module EvarInfoMap = struct
- type t = evar_info ExistentialMap.t
+ type t = evar_info ExistentialMap.t
- let empty = ExistentialMap.empty
+ let empty = ExistentialMap.empty
+ let is_empty = ExistentialMap.is_empty
let to_list evc = (* Workaround for change in Map.fold behavior *)
+ (* spiwack: seems to arrange the items in decreasing order.
+ Which would also be the behaviour of a naive [fold].
+ I don't understand above comment. *)
let l = ref [] in
ExistentialMap.iter (fun evk x -> l := (evk,x)::!l) evc;
!l
let dom evc = ExistentialMap.fold (fun evk _ acc -> evk::acc) evc []
let find evc k = ExistentialMap.find k evc
- let remove evc k = ExistentialMap.remove k evc
+ let remove evc k = ExistentialMap.remove k evc
let mem evc k = ExistentialMap.mem k evc
let fold = ExistentialMap.fold
let exists evc f = ExistentialMap.fold (fun k v b -> b || f k v) evc false
let add evd evk newinfo = ExistentialMap.add evk newinfo evd
- let equal = ExistentialMap.equal
+ let map = ExistentialMap.map
let define evd evk body =
let oldinfo =
@@ -182,6 +187,14 @@ module EvarInfoMap = struct
try Some (existential_value sigma ev)
with NotInstantiatedEvar -> None
+ (* Combinators on undefined evars. *)
+ let fold_undefined f = ExistentialMap.fold begin fun ev evi acc ->
+ match evar_body evi with
+ | Evar_empty -> f ev evi acc
+ | _ -> acc
+ end
+
+ let undefined evm = fold_undefined ExistentialMap.add evm empty
end
(*******************************************************************)
@@ -336,7 +349,23 @@ module EvarMap = struct
(EvarInfoMap.find sigma2 k).evar_body <> Evar_empty)
|| not (UniverseMap.equal (=) sm1 sm2))
- let merge e e' = fold (fun n v sigma -> add sigma n v) e' e
+ (* spiwack: used to workaround a bug in clenv: evar_merge
+ could merge an "old" version of an evar map into
+ a more up to date and erase evar definitions (in the case
+ of [constructor_tac] it actually erased a goal in some cases).
+ This is due to the fact that [open_constr] carry around their own
+ sigma which can be outdated by other operations. *)
+ let add_if_more_recent evd evk newinfo =
+ if newinfo.evar_body = Evar_empty && mem evd evk then
+ evd
+ else
+ add evd evk newinfo
+
+ let merge e e' = fold (fun n v sigma -> add_if_more_recent sigma n v) e' e
+
+ (* combinators on undefined values *)
+ let undefined (sigma,sm) = (EvarInfoMap.undefined sigma,sm)
+ let fold_undefined f (sigma,_) = EvarInfoMap.fold_undefined f sigma
end
@@ -441,7 +470,6 @@ let progress_evar_map d1 d2 =
(* spiwack: tentative. It might very well not be the semantics we want
for merging evar_map *)
let merge d1 d2 = {
-(* d1 with evars = EvarMap.merge d1.evars d2.evars*)
evars = EvarMap.merge d1.evars d2.evars ;
conv_pbs = List.rev_append d1.conv_pbs d2.conv_pbs ;
last_mods = ExistentialSet.union d1.last_mods d2.last_mods ;
@@ -488,7 +516,7 @@ let subst_evar_defs_light sub evd =
assert (evd.conv_pbs = []);
{ evd with
metas = Metamap.map (map_clb (subst_mps sub)) evd.metas;
- evars = ExistentialMap.map (subst_evar_info sub) (fst evd.evars), snd evd.evars
+ evars = EvarInfoMap.map (subst_evar_info sub) (fst evd.evars), (snd evd.evars)
}
let subst_evar_map = subst_evar_defs_light
@@ -536,7 +564,7 @@ let evar_declare hyps evk ty ?(src=(dummy_loc,InternalHole)) ?filter evd =
evar_body = Evar_empty;
evar_filter = filter;
evar_source = src;
- evar_extra = None} }
+ evar_extra = Store.empty } }
let is_defined_evar evd (evk,_) = EvarMap.is_defined evd.evars evk
@@ -546,12 +574,10 @@ let is_undefined_evar evd c = match kind_of_term c with
| _ -> false
let undefined_evars evd =
- let evars =
- EvarMap.fold (fun evk evi sigma -> if evi.evar_body = Evar_empty then
- EvarMap.add sigma evk evi else sigma)
- evd.evars EvarMap.empty
- in
- { evd with evars = evars }
+ let evars = EvarMap.undefined evd.evars in
+ { evd with evars = evars }
+
+let fold_undefined f evd = EvarMap.fold_undefined f evd.evars
(* extracts conversion problems that satisfy predicate p *)
(* Note: conv_pbs not satisying p are stored back in reverse order *)
@@ -694,7 +720,6 @@ let meta_with_name evd id =
strbrk "\" occurs more than once in clause.")
-(* spiwack: we should try and replace this List.fold_left by a Metamap.fold. *)
let meta_merge evd1 evd2 =
{evd2 with
metas = List.fold_left (fun m (n,v) -> Metamap.add n v m)
diff --git a/pretyping/evd.mli b/pretyping/evd.mli
index b82e6d998d..f895ead42c 100644
--- a/pretyping/evd.mli
+++ b/pretyping/evd.mli
@@ -123,7 +123,7 @@ type evar_info = {
evar_body : evar_body;
evar_filter : bool list;
evar_source : hole_kind located;
- evar_extra : Dyn.t option}
+ evar_extra : Store.t }
val eq_evar_info : evar_info -> evar_info -> bool
@@ -185,6 +185,15 @@ val evars_reset_evd : evar_map -> evar_map -> evar_map
for moving to evarutils *)
val is_undefined_evar : evar_map -> constr -> bool
val undefined_evars : evar_map -> evar_map
+(* [fold_undefined f m] iterates ("folds") function [f] over the undefined
+ evars (that is, whose value is [Evar_empty]) of map [m].
+ Making it effectively equivalent to {!Evd.fold} applies to
+ [f] and [undefined_evars m] *)
+(* spiwack: at the moment, [fold_undefined] is defined rather naively.
+ But we can hope to enable some optimisation in the future, as
+ an evar_map contains typically few undefined evars compared to all
+ its evars. *)
+val fold_undefined : (evar -> evar_info -> 'a -> 'a) -> evar_map -> 'a -> 'a
val evar_declare :
named_context_val -> evar -> types -> ?src:loc * hole_kind ->
?filter:bool list -> evar_map -> evar_map
diff --git a/pretyping/pretyping.mllib b/pretyping/pretyping.mllib
index cea33c1e98..6aa00c5f25 100644
--- a/pretyping/pretyping.mllib
+++ b/pretyping/pretyping.mllib
@@ -21,7 +21,6 @@ Typeclasses
Classops
Coercion
Unification
-Clenv
Detyping
Indrec
Cases
diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml
index 4ec5cf1f17..88d4a4d6b3 100644
--- a/pretyping/typeclasses.ml
+++ b/pretyping/typeclasses.ml
@@ -355,18 +355,15 @@ let is_implicit_arg k =
calls (e.g. when doing apply in an External hint in typeclass_instances).
Would be solved by having real evars-as-goals. *)
-let ((bool_in : bool -> Dyn.t),
- (bool_out : Dyn.t -> bool)) = Dyn.create "bool"
-
-let bool_false = bool_in false
+let resolvable = Store.field ()
+open Store.Field
let is_resolvable evi =
- match evi.evar_extra with
- Some t -> if Dyn.tag t = "bool" then bool_out t else true
- | None -> true
+ Option.default true (resolvable.get evi.evar_extra)
let mark_unresolvable evi =
- { evi with evar_extra = Some bool_false }
+ let t = resolvable.set false evi.evar_extra in
+ { evi with evar_extra = t }
let mark_unresolvables sigma =
Evd.fold (fun ev evi evs ->
@@ -374,7 +371,7 @@ let mark_unresolvables sigma =
sigma Evd.empty
let has_typeclasses evd =
- Evd.fold (fun ev evi has -> has ||
+ Evd.fold_undefined (fun ev evi has -> has ||
(evi.evar_body = Evar_empty && is_class_evar evd evi && is_resolvable evi))
evd false
diff --git a/pretyping/typeclasses.mli b/pretyping/typeclasses.mli
index 997b28c105..e53be5c0bf 100644
--- a/pretyping/typeclasses.mli
+++ b/pretyping/typeclasses.mli
@@ -77,10 +77,6 @@ val is_implicit_arg : hole_kind -> bool
val instance_constructor : typeclass -> constr list -> constr * types
-(* Use evar_extra for marking resolvable evars. *)
-val bool_in : bool -> Dyn.t
-val bool_out : Dyn.t -> bool
-
val is_resolvable : evar_info -> bool
val mark_unresolvable : evar_info -> evar_info
val mark_unresolvables : evar_map -> evar_map
diff --git a/pretyping/clenv.ml b/proofs/clenv.ml
index 7cbaf124a7..35db5e6ed1 100644
--- a/pretyping/clenv.ml
+++ b/proofs/clenv.ml
@@ -31,10 +31,10 @@ open Coercion.Default
(* Abbreviations *)
-let pf_env gls = Global.env_of_context gls.it.evar_hyps
-let pf_hyps gls = named_context_of_val gls.it.evar_hyps
+let pf_env = Refiner.pf_env
+let pf_hyps = Refiner.pf_hyps
let pf_type_of gls c = Typing.type_of (pf_env gls) gls.sigma c
-let pf_concl gl = gl.it.evar_concl
+let pf_concl = Tacmach.pf_concl
(******************************************************************)
(* Clausal environments *)
@@ -142,7 +142,7 @@ let mk_clenv_from_env environ sigma n (c,cty) =
env = environ }
let mk_clenv_from_n gls n (c,cty) =
- mk_clenv_from_env (Global.env_of_context gls.it.evar_hyps) gls.sigma n (c, cty)
+ mk_clenv_from_env (pf_env gls) gls.sigma n (c, cty)
let mk_clenv_from gls = mk_clenv_from_n gls None
@@ -250,12 +250,13 @@ let clenv_unify_meta_types ?(flags=default_unify_flags) clenv =
{ clenv with evd = w_unify_meta_types ~flags:flags clenv.env clenv.evd }
let clenv_unique_resolver allow_K ?(flags=default_unify_flags) clenv gl =
+ let concl = Goal.V82.concl clenv.evd (sig_it gl) in
if isMeta (fst (whd_stack clenv.evd clenv.templtyp.rebus)) then
- clenv_unify allow_K CUMUL ~flags:flags (clenv_type clenv) (pf_concl gl)
+ clenv_unify allow_K CUMUL ~flags:flags (clenv_type clenv) concl
(clenv_unify_meta_types ~flags:flags clenv)
else
clenv_unify allow_K CUMUL ~flags:flags
- (meta_reducible_instance clenv.evd clenv.templtyp) (pf_concl gl) clenv
+ (meta_reducible_instance clenv.evd clenv.templtyp) concl clenv
(* [clenv_pose_metas_as_evars clenv dep_mvs]
* For each dependent evar in the clause-env which does not have a value,
@@ -304,9 +305,18 @@ let evar_clenv_unique_resolver = clenv_unique_resolver
(******************************************************************)
let connect_clenv gls clenv =
+ let evd = Evd.fold begin fun ev evi acc ->
+ if evi.evar_body = Evar_empty then
+ acc
+ else
+ Evd.add acc ev evi
+ end
+ clenv.evd gls.sigma
+ in
+ let evd = evars_reset_evd evd clenv.evd in
{ clenv with
- evd = evars_reset_evd gls.sigma clenv.evd;
- env = Global.env_of_context gls.it.evar_hyps }
+ evd = evd ;
+ env = Goal.V82.env evd (sig_it gls) }
(* [clenv_fchain mv clenv clenv']
*
diff --git a/pretyping/clenv.mli b/proofs/clenv.mli
index 4f7ac40920..2533fc5379 100644
--- a/pretyping/clenv.mli
+++ b/proofs/clenv.mli
@@ -50,16 +50,16 @@ val clenv_nf_meta : clausenv -> constr -> constr
(* type of a meta in clenv context *)
val clenv_meta_type : clausenv -> metavariable -> types
-val mk_clenv_from : evar_info sigma -> constr * types -> clausenv
+val mk_clenv_from : Goal.goal sigma -> constr * types -> clausenv
val mk_clenv_from_n :
- evar_info sigma -> int option -> constr * types -> clausenv
-val mk_clenv_type_of : evar_info sigma -> constr -> clausenv
+ Goal.goal sigma -> int option -> constr * types -> clausenv
+val mk_clenv_type_of : Goal.goal sigma -> constr -> clausenv
val mk_clenv_from_env : env -> evar_map -> int option -> constr * types -> clausenv
(***************************************************************)
(* linking of clenvs *)
-val connect_clenv : evar_info sigma -> clausenv -> clausenv
+val connect_clenv : Goal.goal sigma -> clausenv -> clausenv
val clenv_fchain :
?allow_K:bool -> ?flags:unify_flags -> metavariable -> clausenv -> clausenv -> clausenv
@@ -72,12 +72,12 @@ val clenv_unify :
(* unifies the concl of the goal with the type of the clenv *)
val clenv_unique_resolver :
- bool -> ?flags:unify_flags -> clausenv -> evar_info sigma -> clausenv
+ bool -> ?flags:unify_flags -> clausenv -> Goal.goal sigma -> clausenv
(* same as above ([allow_K=false]) but replaces remaining metas
with fresh evars if [evars_flag] is [true] *)
val evar_clenv_unique_resolver :
- bool -> ?flags:unify_flags -> clausenv -> evar_info sigma -> clausenv
+ bool -> ?flags:unify_flags -> clausenv -> Goal.goal sigma -> clausenv
val clenv_dependent : bool -> clausenv -> metavariable list
@@ -107,10 +107,10 @@ val clenv_unify_meta_types : ?flags:unify_flags -> clausenv -> clausenv
(* the optional int tells how many prods of the lemma have to be used *)
(* use all of them if None *)
val make_clenv_binding_apply :
- evar_info sigma -> int option -> constr * constr -> constr bindings ->
+ Goal.goal sigma -> int option -> constr * constr -> constr bindings ->
clausenv
val make_clenv_binding :
- evar_info sigma -> constr * constr -> constr bindings -> clausenv
+ Goal.goal sigma -> constr * constr -> constr bindings -> clausenv
(* [clenv_environments sigma n t] returns [sigma',lmeta,ccl] where
[lmetas] is a list of metas to be applied to a proof of [t] so that
diff --git a/proofs/clenvtac.ml b/proofs/clenvtac.ml
index 9bc818e861..f977768bd6 100644
--- a/proofs/clenvtac.ml
+++ b/proofs/clenvtac.ml
@@ -20,7 +20,6 @@ open Evd
open Evarutil
open Proof_type
open Refiner
-open Proof_trees
open Logic
open Reduction
open Reductionops
diff --git a/proofs/evar_refiner.ml b/proofs/evar_refiner.ml
index e4fab3f2f8..6dfbbdc121 100644
--- a/proofs/evar_refiner.ml
+++ b/proofs/evar_refiner.ml
@@ -14,7 +14,6 @@ open Term
open Evd
open Evarutil
open Sign
-open Proof_trees
open Refiner
(******************************************)
@@ -55,19 +54,10 @@ let w_refine (evk,evi) (ltac_var,rawc) sigma =
(* vernac command Existential *)
-let instantiate_pf_com n com pfts =
- let gls = top_goal_of_pftreestate pfts in
- let sigma = gls.sigma in
- let (evk,evi) =
- let evl = Evarutil.non_instantiated sigma in
- if (n <= 0) then
- error "incorrect existential variable index"
- else if List.length evl < n then
- error "not so many uninstantiated existential variables"
- else
- List.nth evl (n-1)
- in
+(* Main component of vernac command Existential *)
+let instantiate_pf_com evk com sigma =
+ let evi = Evd.find sigma evk in
let env = Evd.evar_env evi in
let rawc = Constrintern.intern_constr sigma env com in
let sigma' = w_refine (evk,evi) (([],[]),rawc) sigma in
- change_constraints_pftreestate sigma' pfts
+ sigma'
diff --git a/proofs/evar_refiner.mli b/proofs/evar_refiner.mli
index 0bd6168095..28c79d11e9 100644
--- a/proofs/evar_refiner.mli
+++ b/proofs/evar_refiner.mli
@@ -24,6 +24,6 @@ val w_refine : evar * evar_info ->
(var_map * unbound_ltac_var_map) * rawconstr -> evar_map -> evar_map
val instantiate_pf_com :
- int -> Topconstr.constr_expr -> pftreestate -> pftreestate
+ Evd.evar -> Topconstr.constr_expr -> Evd.evar_map -> Evd.evar_map
(* the instantiate tactic was moved to [tactics/evar_tactics.ml] *)
diff --git a/proofs/goal.ml b/proofs/goal.ml
new file mode 100644
index 0000000000..a9202318df
--- /dev/null
+++ b/proofs/goal.ml
@@ -0,0 +1,572 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesscer General Public License Version 2.1 *)
+(************************************************************************)
+
+(* $Id$ *)
+
+open Term
+
+(* This module implements the abstract interface to goals *)
+(* A general invariant of the module, is that a goal whose associated
+ evar is defined in the current evar_map, should not be accessed. *)
+
+(* type of the goals *)
+type goal = {
+ content : Evd.evar; (* Corresponding evar. Allows to retrieve
+ logical information once put together
+ with an evar_map. *)
+ tags : string list (* Heriditary? tags to be displayed *)
+}
+(* spiwack: I don't deal with the tags, yet. It is a worthy discussion
+ whether we do want some tags displayed besides the goal or not. *)
+
+(* access primitive *)
+(* invariant : [e] must exist in [em] *)
+let content evars { content = e } = Evd.find evars e
+
+
+(* Builds a new (empty) goal with evar [e] *)
+let build e =
+ { content = e ;
+ tags = []
+ }
+
+(* Builds a new goal with content evar [e] and
+ inheriting from goal [gl]*)
+let descendent gl e =
+ { gl with content = e}
+
+(* [advance sigma g] returns [Some g'] if [g'] is undefined and
+ is the current avatar of [g] (for instance [g] was changed by [clear]
+ into [g']). It returns [None] if [g] has been (partially) solved. *)
+open Store.Field
+let rec advance sigma g =
+ let evi = Evd.find sigma g.content in
+ if Option.default false (Evarutil.cleared.get evi.Evd.evar_extra) then
+ let v =
+ match evi.Evd.evar_body with
+ | Evd.Evar_defined c -> c
+ | _ -> Util.anomaly "Some goal is marked as 'cleared' but is uninstantiated"
+ in
+ let (e,_) = Term.destEvar v in
+ let g' = { g with content = e } in
+ advance sigma g'
+ else
+ match evi.Evd.evar_body with
+ | Evd.Evar_defined _ -> None
+ | _ -> Some g
+
+(*** Goal tactics ***)
+
+
+(* Goal tactics are [subgoal sensitive]-s *)
+type subgoals = { subgoals: goal list }
+
+(* type of the base elements of the goal API.*)
+(* it has an extra evar_info with respect to what would be expected,
+ it is supposed to be the evar_info of the goal in the evar_map.
+ The idea is that it is computed by the [run] function as an
+ optimisation, since it will generaly not change during the
+ evaluation. *)
+type 'a sensitive =
+ Environ.env -> Evd.evar_map ref -> goal -> Evd.evar_info -> 'a
+
+(* evaluates a goal sensitive value in a given goal (knowing the current evar_map). *)
+(* the evar_info corresponding to the goal is computed at once
+ as an optimisation (it shouldn't change during the evaluation). *)
+let eval t env defs gl =
+ let info = content defs gl in
+ let env = Environ.reset_with_named_context (Evd.evar_hyps info) env in
+ let rdefs = ref defs in
+ let r = t env rdefs gl info in
+ ( r , !rdefs )
+
+(* monadic bind on sensitive expressions *)
+let bind e f env rdefs goal info =
+ f (e env rdefs goal info) env rdefs goal info
+
+(* monadic return on sensitive expressions *)
+let return v _ _ _ _ = v
+
+(* interpretation of "open" constr *)
+(* spiwack: it is a wrapper around [Constrintern.interp_open_constr].
+ In an ideal world, this could/should be the other way round.
+ As of now, though, it seems at least quite useful to build tactics. *)
+let interp_constr cexpr env rdefs _ _ =
+ let (defs,c) = Constrintern.interp_open_constr !rdefs env cexpr in
+ rdefs := defs ;
+ c
+
+(* Type of constr with holes used by refine. *)
+(* The list of evars doesn't necessarily contain all the evars in the constr,
+ only those the constr has introduced. *)
+(* The variables in [myevars] are supposed to be stored
+ in decreasing order. Breaking this invariant might cause
+ many things to go wrong. *)
+type refinable = {
+ me: constr;
+ my_evars: Evd.evar list
+}
+
+module Refinable = struct
+ type t = refinable
+
+ type handle = Evd.evar list ref
+
+ let make t env rdefs gl info =
+ let r = ref [] in
+ let me = t r env rdefs gl info in
+ { me = me;
+ my_evars = !r }
+ let make_with t env rdefs gl info =
+ let r = ref [] in
+ let (me,side) = t r env rdefs gl info in
+ { me = me ; my_evars = !r } , side
+
+ let mkEvar handle env typ _ rdefs _ _ =
+ let ev = Evarutil.e_new_evar rdefs env typ in
+ let (e,_) = Term.destEvar ev in
+ handle := e::!handle;
+ ev
+
+ (* [with_type c typ] constrains term [c] to have type [typ]. *)
+ let with_type t typ env rdefs _ _ =
+ (* spiwack: this function assumes that no evars can be created during
+ this sort of coercion.
+ If it is not the case it could produce bugs. We would need to add a handle
+ and add the new evars to it. *)
+ let my_type = Retyping.get_type_of env !rdefs t in
+ let j = Environ.make_judge t my_type in
+ let tycon = Evarutil.mk_tycon_type typ in
+ let (new_defs,j') =
+ Coercion.Default.inh_conv_coerce_to (Util.dummy_loc) env !rdefs j tycon
+ in
+ rdefs := new_defs;
+ j'.Environ.uj_val
+
+ (* spiwack: it is not very fine grain since it solves all typeclasses holes,
+ not only those containing the current goal, or a given term. But it
+ seems to fit our needs so far. *)
+ let resolve_typeclasses ?onlyargs ?split ?(fail=false) () env rdefs _ _ =
+ rdefs:=Typeclasses.resolve_typeclasses ?onlyargs ?split ~fail env !rdefs;
+ ()
+
+
+
+ (* a pessimistic (i.e : there won't be many positive answers) filter
+ over evar_maps *)
+ let evar_map_filter f evm =
+ Evd.fold (fun ev evi r ->
+ if f ev evi then
+ Evd.add r ev evi
+ else
+ r
+ )
+ evm
+ Evd.empty
+
+ (* Union, sorted in decreasing order, of two lists of evars in decreasing order. *)
+ let rec fusion l1 l2 = match l1 , l2 with
+ | [] , _ -> l2
+ | _ , [] -> l1
+ | a::l1 , b::_ when a>b -> a::(fusion l1 l2)
+ | a::l1 , b::l2 when a=b -> a::(fusion l1 l2)
+ | _ , b::l2 -> b::(fusion l1 l2)
+
+ (* [constr_of_raw] is a pretyping function. The [check_type] argument
+ asks whether the term should have the same type as the conclusion.
+ [resolve_classes] is a flag on pretyping functions which, if set to true,
+ calls the typeclass resolver.
+ The principal argument is a [rawconstr] which is then pretyped in the
+ context of a term, the remaining evars are registered to the handle.
+ It is the main component of the toplevel refine tactic.*)
+ (* spiwack: it is not entirely satisfactory to have this function here. Plus it is
+ a bit hackish. However it does not seem possible to move it out until
+ pretyping is defined as some proof procedure. *)
+ let constr_of_raw handle check_type resolve_classes rawc env rdefs gl info =
+ (* We need to keep trace of what [rdefs] was originally*)
+ let init_defs = !rdefs in
+ (* if [check_type] is true, then creates a type constraint for the
+ proof-to-be *)
+ let tycon = Pretyping.OfType (Option.init check_type (Evd.evar_concl info)) in
+ (* call to [understand_tcc_evars] returns a constr with undefined evars
+ these evars will be our new goals *)
+ let open_constr =
+ Pretyping.Default.understand_tcc_evars ~resolve_classes rdefs env tycon rawc
+ in
+ (* [!rdefs] contains the evar_map outputed by [understand_...] *)
+ let post_defs = !rdefs in
+ (* [delta_evars] holds the evars that have been introduced by this
+ refinement (but not immediatly solved) *)
+ (* spiwack: this is the hackish part, don't know how to do any better though. *)
+ let delta_evars = evar_map_filter (fun ev evi ->
+ evi.Evd.evar_body = Evd.Evar_empty &&
+ not (Evd.mem init_defs ev)
+ )
+ post_defs
+ in
+ (* [delta_evars] in the shape of a list of [evar]-s*)
+ let delta_list = List.map fst (Evd.to_list delta_evars) in
+ (* The variables in [myevars] are supposed to be stored
+ in decreasing order. Breaking this invariant might cause
+ many things to go wrong. *)
+ handle := fusion delta_list !handle ;
+ open_constr
+
+end
+
+(* [refine t] takes a refinable term and use it as a partial proof for current
+ goal. *)
+let refine step env rdefs gl info =
+ (* subgoals to return *)
+ (* The evars in [my_evars] are stored in reverse order.
+ It is expectingly better however to display the goal
+ in increasing order. *)
+ rdefs := Evarconv.consider_remaining_unif_problems env !rdefs ;
+ let subgoals = List.map (descendent gl) (List.rev step.my_evars) in
+ (* creates the new [evar_map] by defining the evar of the current goal
+ as being [refine_step]. *)
+ let new_defs = Evd.define gl.content (step.me) !rdefs in
+ rdefs := new_defs;
+ (* Filtering the [subgoals] for uninstanciated (=unsolved) goals. *)
+ let subgoals =
+ Option.List.flatten (List.map (advance !rdefs) subgoals)
+ in
+ { subgoals = subgoals }
+
+
+(*** Cleaning goals ***)
+
+let clear ids env rdefs gl info =
+ let hyps = Evd.evar_hyps info in
+ let concl = Evd.evar_concl info in
+ let (hyps,concl) = Evarutil.clear_hyps_in_evi rdefs hyps concl ids in
+ let cleared_env = Environ.reset_with_named_context hyps env in
+ let cleared_concl = Evarutil.e_new_evar rdefs cleared_env concl in
+ let (cleared_evar,_) = Term.destEvar cleared_concl in
+ let cleared_goal = descendent gl cleared_evar in
+ rdefs := Evd.define gl.content cleared_concl !rdefs;
+ { subgoals = [cleared_goal] }
+
+let wrap_apply_to_hyp_and_dependent_on sign id f g =
+ try Environ.apply_to_hyp_and_dependent_on sign id f g
+ with Environ.Hyp_not_found ->
+ Util.error "No such assumption"
+
+let check_typability env sigma c =
+ let _ = Typing.type_of env sigma c in ()
+
+let recheck_typability (what,id) env sigma t =
+ try check_typability env sigma t
+ with _ ->
+ let s = match what with
+ | None -> "the conclusion"
+ | Some id -> "hypothesis "^(Names.string_of_id id) in
+ Util.error
+ ("The correctness of "^s^" relies on the body of "^(Names.string_of_id id))
+
+let remove_hyp_body env sigma id =
+ let sign =
+ wrap_apply_to_hyp_and_dependent_on (Environ.named_context_val env) id
+ (fun (_,c,t) _ ->
+ match c with
+ | None -> Util.error ((Names.string_of_id id)^" is not a local definition")
+ | Some c ->(id,None,t))
+ (fun (id',c,t as d) sign ->
+ (
+ begin
+ let env = Environ.reset_with_named_context sign env in
+ match c with
+ | None -> recheck_typability (Some id',id) env sigma t
+ | Some b ->
+ let b' = mkCast (b,DEFAULTcast, t) in
+ recheck_typability (Some id',id) env sigma b'
+ end;d))
+ in
+ Environ.reset_with_named_context sign env
+
+
+let clear_body idents env rdefs gl info =
+ let info = content !rdefs gl in
+ let full_env = Environ.reset_with_named_context (Evd.evar_hyps info) env in
+ let aux env id =
+ let env' = remove_hyp_body env !rdefs id in
+ recheck_typability (None,id) env' !rdefs (Evd.evar_concl info);
+ env'
+ in
+ let new_env =
+ List.fold_left aux full_env idents
+ in
+ let concl = Evd.evar_concl info in
+ let (defs',new_constr) = Evarutil.new_evar !rdefs new_env concl in
+ let (new_evar,_) = destEvar new_constr in
+ let new_goal = descendent gl new_evar in
+ rdefs := Evd.define gl.content new_constr defs';
+ { subgoals = [new_goal] }
+
+
+(*** Sensitive primitives ***)
+
+(* [concl] is the conclusion of the current goal *)
+let concl _ _ _ info =
+ Evd.evar_concl info
+
+(* [hyps] is the [named_context_val] representing the hypotheses
+ of the current goal *)
+let hyps _ _ _ info =
+ Evd.evar_hyps info
+
+(* [env] is the current [Environ.env] containing both the
+ environment in which the proof is ran, and the goal hypotheses *)
+let env env _ _ _ = env
+
+(* [defs] is the [Evd.evar_map] at the current evaluation point *)
+let defs _ rdefs _ _ =
+ !rdefs
+
+(* Cf mli for more detailed comment.
+ [null], [plus], [here] and [here_list] use internal exception [UndefinedHere]
+ to communicate whether or not the value is defined in the particular context. *)
+exception UndefinedHere
+let null _ _ _ _ = raise UndefinedHere
+
+let plus s1 s2 env rdefs goal info =
+ try s1 env rdefs goal info
+ with UndefinedHere -> s2 env rdefs goal info
+
+(* Equality of two goals *)
+let equal { content = e1 } { content = e2 } = e1 = e2
+
+let here goal value _ _ goal' _ =
+ if equal goal goal' then
+ value
+ else
+ raise UndefinedHere
+
+(* arnaud: voir à la passer dans Util ? *)
+let rec list_mem_with eq x = function
+ | y::_ when eq x y -> true
+ | _::l -> list_mem_with eq x l
+ | [] -> false
+
+let here_list goals value _ _ goal' _ =
+ if list_mem_with equal goal' goals then
+ value
+ else
+ raise UndefinedHere
+
+
+(*** Conversion in goals ***)
+
+let convert_hyp check (id,b,bt as d) env rdefs gl info =
+ let sigma = !rdefs in
+ (* This function substitutes the new type and body definitions
+ in the appropriate variable when used with {!Environ.apply_hyps}. *)
+ let replace_function =
+ (fun _ (_,c,ct) _ ->
+ if check && not (Reductionops.is_conv env sigma bt ct) then
+ Util.error ("Incorrect change of the type of "^(Names.string_of_id id));
+ if check && not (Option.Misc.compare (Reductionops.is_conv env sigma) b c) then
+ Util.error ("Incorrect change of the body of "^(Names.string_of_id id));
+ d)
+ in
+ (* Modified named context. *)
+ let new_hyps =
+ Environ.apply_to_hyp (hyps env rdefs gl info) id replace_function
+ in
+ let new_env = Environ.reset_with_named_context new_hyps env in
+ let new_constr =
+ Evarutil.e_new_evar rdefs new_env (concl env rdefs gl info)
+ in
+ let (new_evar,_) = Term.destEvar new_constr in
+ let new_goal = descendent gl new_evar in
+ rdefs := Evd.define gl.content new_constr !rdefs;
+ { subgoals = [new_goal] }
+
+let convert_concl check cl' env rdefs gl info =
+ let sigma = !rdefs in
+ let cl = concl env rdefs gl info in
+ check_typability env sigma cl';
+ if (not check) || Reductionops.is_conv_leq env sigma cl' cl then
+ let new_constr =
+ Evarutil.e_new_evar rdefs env cl'
+ in
+ let (new_evar,_) = Term.destEvar new_constr in
+ let new_goal = descendent gl new_evar in
+ rdefs := Evd.define gl.content new_constr !rdefs;
+ { subgoals = [new_goal] }
+ else
+ Util.error "convert-concl rule passed non-converting term"
+
+
+(*** Bureaucracy in hypotheses ***)
+
+(* Renames a hypothesis. *)
+let rename_hyp_sign id1 id2 sign =
+ Environ.apply_to_hyp_and_dependent_on sign id1
+ (fun (_,b,t) _ -> (id2,b,t))
+ (fun d _ -> map_named_declaration (replace_vars [id1,mkVar id2]) d)
+let rename_hyp id1 id2 env rdefs gl info =
+ let hyps = hyps env rdefs gl info in
+ if id1 <> id2 &&
+ List.mem id2 (Termops.ids_of_named_context (Environ.named_context_of_val hyps)) then
+ Util.error ((Names.string_of_id id2)^" is already used.");
+ let new_hyps = rename_hyp_sign id1 id2 hyps in
+ let new_env = Environ.reset_with_named_context new_hyps env in
+ let new_concl = Term.replace_vars [id1,mkVar id2] (concl env rdefs gl info) in
+ let new_subproof = Evarutil.e_new_evar rdefs new_env new_concl in
+ let new_subproof = Term.replace_vars [id2,mkVar id1] new_subproof in
+ let (new_evar,_) = Term.destEvar new_subproof in
+ let new_goal = descendent gl new_evar in
+ rdefs := Evd.define gl.content new_subproof !rdefs;
+ { subgoals = [new_goal] }
+
+(*** Additional functions ***)
+
+(* emulates List.map for functions of type
+ [Evd.evar_map -> 'a -> 'b * Evd.evar_map] on lists of type 'a, by propagating
+ new evar_map to next definition. *)
+(*This sort of construction actually works with any monad (here the State monade
+ in Haskell). There is a generic construction in Haskell called mapM.
+*)
+let rec list_map f l s =
+ match l with
+ | [] -> ([],s)
+ | a::l -> let (a,s) = f s a in
+ let (l,s) = list_map f l s in
+ (a::l,s)
+
+
+(* Layer to implement v8.2 tactic engine ontop of the new architecture.
+ Types are different from what they used to be due to a change of the
+ internal types. *)
+module V82 = struct
+
+ (* Old style env primitive *)
+ let env evars gl =
+ let evi = content evars gl in
+ Evd.evar_env evi
+
+ (* For printing *)
+ let unfiltered_env evars gl =
+ let evi = content evars gl in
+ Evd.evar_unfiltered_env evi
+
+ (* Old style hyps primitive *)
+ let hyps evars gl =
+ let evi = content evars gl in
+ evi.Evd.evar_hyps
+
+ (* Access to ".evar_concl" *)
+ let concl evars gl =
+ let evi = content evars gl in
+ evi.Evd.evar_concl
+
+ (* Access to ".evar_extra" *)
+ let extra evars gl =
+ let evi = content evars gl in
+ evi.Evd.evar_extra
+
+ (* Old style filtered_context primitive *)
+ let filtered_context evars gl =
+ let evi = content evars gl in
+ Evd.evar_filtered_context evi
+
+ (* Old style mk_goal primitive *)
+ let mk_goal evars hyps concl extra =
+ let evk = Evarutil.new_untyped_evar () in
+ let evi = { Evd.evar_hyps = hyps;
+ Evd.evar_concl = concl;
+ Evd.evar_filter = List.map (fun _ -> true)
+ (Environ.named_context_of_val hyps);
+ Evd.evar_body = Evd.Evar_empty;
+ Evd.evar_source = (Util.dummy_loc,Evd.GoalEvar);
+ Evd.evar_extra = extra }
+ in
+ let evi = Typeclasses.mark_unresolvable evi in
+ let evars = Evd.add evars evk evi in
+ let ids = List.map Util.pi1 (Environ.named_context_of_val hyps) in
+ let inst = Array.of_list (List.map mkVar ids) in
+ let ev = Term.mkEvar (evk,inst) in
+ (build evk, ev, evars)
+
+ (* Equality function on goals *)
+ let equal evars gl1 gl2 =
+ let evi1 = content evars gl1 in
+ let evi2 = content evars gl2 in
+ Evd.eq_evar_info evi1 evi2
+
+ (* Creates a dummy [goal sigma] for use in auto *)
+ let dummy_goal =
+ (* This goal seems to be marshalled somewhere. Therefore it cannot be
+ marked unresolvable for typeclasses, as non-empty Store.t-s happen
+ to have functional content. *)
+ let evi = Evd.make_evar Environ.empty_named_context_val Term.mkProp in
+ let evk = Evarutil.new_untyped_evar () in
+ let sigma = Evd.add Evd.empty evk evi in
+ { Evd.it = build evk ; Evd.sigma = sigma }
+
+ (* Makes a goal out of an evar *)
+ let build = build
+
+ (* Instantiates a goal with an open term *)
+ let partial_solution sigma { content=evk } c =
+ Evd.define evk c sigma
+
+ (* Parts of the progress tactical *)
+ let same_goal evars1 gl1 evars2 gl2 =
+ let evi1 = content evars1 gl1 in
+ let evi2 = content evars2 gl2 in
+ Term.eq_constr evi1.Evd.evar_concl evi2.Evd.evar_concl &&
+ Environ.eq_named_context_val evi1.Evd.evar_hyps evi2.Evd.evar_hyps
+
+ let weak_progress glss gls =
+ match glss.Evd.it with
+ | [ g ] -> not (same_goal glss.Evd.sigma g gls.Evd.sigma gls.Evd.it)
+ | _ -> true
+
+ let progress glss gls =
+ weak_progress glss gls
+ (* spiwack: progress normally goes like this:
+ (Evd.progress_evar_map gls.Evd.sigma glss.Evd.sigma) || (weak_progress glss gls)
+ This is immensly slow in the current implementation. Maybe we could
+ reimplement progress_evar_map with restricted folds like "fold_undefined",
+ with a good implementation of them.
+ *)
+
+ (* Used for congruence closure *)
+ let new_goal_with sigma gl new_hyps =
+ let evi = content sigma gl in
+ let new_evi = { evi with Evd.evar_hyps = new_hyps } in
+ let new_evi = Typeclasses.mark_unresolvable new_evi in
+ let evk = Evarutil.new_untyped_evar () in
+ let new_sigma = Evd.add Evd.empty evk new_evi in
+ { Evd.it = build evk ; sigma = new_sigma }
+
+ (* Used by the typeclasses *)
+ let nf_evar sigma gl =
+ let evi = content sigma gl in
+ let evi = Evarutil.nf_evar_info sigma evi in
+ let sigma = Evd.add sigma gl.content evi in
+ (gl,sigma)
+
+ (* Goal represented as a type, doesn't take into account section variables *)
+ let abstract_type sigma gl =
+ let (gl,sigma) = nf_evar sigma gl in
+ let env = env sigma gl in
+ let genv = Global.env () in
+ let is_proof_var decl =
+ try ignore (Environ.lookup_named (Util.pi1 decl) genv); false
+ with Not_found -> true in
+ Environ.fold_named_context_reverse (fun t decl ->
+ if is_proof_var decl then
+ mkNamedProd_or_LetIn decl t
+ else
+ t
+ ) ~init:(concl sigma gl) env
+
+end
diff --git a/proofs/goal.mli b/proofs/goal.mli
new file mode 100644
index 0000000000..49e3c9b1ae
--- /dev/null
+++ b/proofs/goal.mli
@@ -0,0 +1,228 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* $Id: goal.mli aspiwack $ *)
+
+(* This module implements the abstract interface to goals *)
+
+type goal
+
+(* spiwack: this primitive is not extremely brilliant. It may be a good
+ idea to define goals and proof views in the same file to avoid this
+ sort of communication pipes. But I find it heavy. *)
+val build : Evd.evar -> goal
+
+(* [advance sigma g] returns [Some g'] if [g'] is undefined and
+ is the current avatar of [g] (for instance [g] was changed by [clear]
+ into [g']). It returns [None] if [g] has been (partially) solved. *)
+open Store.Field
+val advance : Evd.evar_map -> goal -> goal option
+
+
+(*** Goal tactics ***)
+
+
+(* Goal tactics are [subgoal sensitive]-s *)
+type subgoals = private { subgoals: goal list }
+
+(* Goal sensitive values *)
+type +'a sensitive
+
+(* evaluates a goal sensitive value in a given goal (knowing the current evar_map). *)
+val eval : 'a sensitive -> Environ.env -> Evd.evar_map -> goal -> 'a * Evd.evar_map
+
+(* monadic bind on sensitive expressions *)
+val bind : 'a sensitive -> ('a -> 'b sensitive) -> 'b sensitive
+
+(* monadic return on sensitive expressions *)
+val return : 'a -> 'a sensitive
+
+
+(* interpretation of "open" constr *)
+(* spiwack: it is a wrapper around [Constrintern.interp_open_constr].
+ In an ideal world, this could/should be the other way round.
+ As of now, though, it seems at least quite useful to build tactics. *)
+val interp_constr : Topconstr.constr_expr -> Term.constr sensitive
+
+(* Type of constr with holes used by refine. *)
+type refinable
+
+module Refinable : sig
+ type t = refinable
+ type handle
+
+ val make : (handle -> Term.constr sensitive) -> refinable sensitive
+ val make_with : (handle -> (Term.constr*'a) sensitive) -> (refinable*'a) sensitive
+
+ val mkEvar : handle -> Environ.env -> Term.types -> Term.constr sensitive
+
+ (* [with_type c typ] constrains term [c] to have type [typ]. *)
+ val with_type : Term.constr -> Term.types -> Term.constr sensitive
+
+ val resolve_typeclasses : ?onlyargs:bool -> ?split:bool -> ?fail:bool -> unit -> unit sensitive
+
+
+ (* [constr_of_raw h check_type resolve_classes] is a pretyping function.
+ The [check_type] argument asks whether the term should have the same
+ type as the conclusion. [resolve_classes] is a flag on pretyping functions
+ which, if set to true, calls the typeclass resolver.
+ The principal argument is a [rawconstr] which is then pretyped in the
+ context of a term, the remaining evars are registered to the handle.
+ It is the main component of the toplevel refine tactic.*)
+ val constr_of_raw :
+ handle -> bool -> bool -> Rawterm.rawconstr -> Term.constr sensitive
+
+end
+
+(* [refine t] takes a refinable term and use it as a partial proof for current
+ goal. *)
+val refine : refinable -> subgoals sensitive
+
+
+(*** Cleaning goals ***)
+
+(* Implements the [clear] tactic *)
+val clear : Names.identifier list -> subgoals sensitive
+
+(* Implements the [clearbody] tactic *)
+val clear_body : Names.identifier list -> subgoals sensitive
+
+
+(*** Conversion in goals ***)
+
+(* Changes an hypothesis of the goal with a convertible type and body.
+ Checks convertibility if the boolean argument is true. *)
+val convert_hyp : bool -> Term.named_declaration -> subgoals sensitive
+
+(* Changes the conclusion of the goal with a convertible type and body.
+ Checks convertibility if the boolean argument is true. *)
+val convert_concl : bool -> Term.constr -> subgoals sensitive
+
+(*** Bureaucracy in hypotheses ***)
+
+(* Renames a hypothesis. *)
+val rename_hyp : Names.identifier -> Names.identifier -> subgoals sensitive
+
+(*** Sensitive primitives ***)
+
+(* [concl] is the conclusion of the current goal *)
+val concl : Term.constr sensitive
+
+(* [hyps] is the [named_context_val] representing the hypotheses
+ of the current goal *)
+val hyps : Environ.named_context_val sensitive
+
+(* [env] is the current [Environ.env] containing both the
+ environment in which the proof is ran, and the goal hypotheses *)
+val env : Environ.env sensitive
+
+(* [defs] is the [Evd.evar_map] at the current evaluation point *)
+val defs : Evd.evar_map sensitive
+
+(* These four functions serve as foundation for the goal sensitive part
+ of the tactic monad (see Proofview).
+ [here] is a special sort of [return]: [here g a] is the value [a], but
+ does not have any value (it raises an exception) if evaluated in
+ any other goal than [g].
+ [here_list] is the same, except with a list of goals rather than a single one.
+ [plus a b] is the same as [a] if [a] is defined in the current goal, otherwise
+ it is [b]. Effectively it's defined in the goals where [a] and [b] are defined.
+ [null] is defined in no goal. (it is a neutral element for [plus]). *)
+(* spiwack: these primitives are a bit hackish, but I couldn't find another way
+ to pass information between goals, like for an intro tactic which gives to
+ each goal the name of the variable it introduce.
+ In pratice, in my experience, the primitives given in Proofview (in terms of
+ [here] and [plus]) are sufficient to define any tactics, hence these might
+ be another example of communication primitives between Goal and Proofview.
+ Still, I can't see a way to prevent using the Proofview primitive to read
+ a goal sensitive value out of its valid context. *)
+val null : 'a sensitive
+
+val plus : 'a sensitive -> 'a sensitive -> 'a sensitive
+
+val here : goal -> 'a -> 'a sensitive
+
+val here_list : goal list -> 'a -> 'a sensitive
+
+(*** Additional functions ***)
+
+(* emulates List.map for functions of type
+ [Evd.evar_map -> 'a -> 'b * Evd.evar_map] on lists of type 'a, by propagating
+ new evar_map to next definition *)
+val list_map : (Evd.evar_map -> 'a -> 'b * Evd.evar_map) ->
+ 'a list ->
+ Evd.evar_map ->
+ 'b list *Evd.evar_map
+
+(* Layer to implement v8.2 tactic engine ontop of the new architecture.
+ Types are different from what they used to be due to a change of the
+ internal types. *)
+module V82 : sig
+
+ (* Old style env primitive *)
+ val env : Evd.evar_map -> goal -> Environ.env
+
+ (* For printing *)
+ val unfiltered_env : Evd.evar_map -> goal -> Environ.env
+
+ (* Old style hyps primitive *)
+ val hyps : Evd.evar_map -> goal -> Environ.named_context_val
+
+ (* Access to ".evar_concl" *)
+ val concl : Evd.evar_map -> goal -> Term.constr
+
+ (* Access to ".evar_extra" *)
+ val extra : Evd.evar_map -> goal -> Store.t
+
+ (* Old style filtered_context primitive *)
+ val filtered_context : Evd.evar_map -> goal -> Sign.named_context
+
+ (* Old style mk_goal primitive, returns a new goal with corresponding
+ hypotheses and conclusion, together with a term which is precisely
+ the evar corresponding to the goal, and an updated evar_map. *)
+ val mk_goal : Evd.evar_map ->
+ Environ.named_context_val ->
+ Term.constr ->
+ Store.t ->
+ goal * Term.constr * Evd.evar_map
+
+ (* Equality function on goals *)
+ val equal : Evd.evar_map -> goal -> goal -> bool
+
+ (* Creates a dummy [goal sigma] for use in auto *)
+ val dummy_goal : goal Evd.sigma
+
+ (* Makes a goal out of an evar *)
+ (* spiwack: used by [Proofview.init], not entirely clean probably, but it is
+ the only way I could think of to preserve compatibility with previous Coq
+ stuff. *)
+ val build : Evd.evar -> goal
+
+
+ (* Instantiates a goal with an open term *)
+ val partial_solution : Evd.evar_map -> goal -> Term.constr -> Evd.evar_map
+
+ (* Principal part of the weak-progress tactical *)
+ val weak_progress : goal list Evd.sigma -> goal Evd.sigma -> bool
+
+ (* Principal part of the progress tactical *)
+ val progress : goal list Evd.sigma -> goal Evd.sigma -> bool
+
+ (* Principal part of tclNOTSAMEGOAL *)
+ val same_goal : Evd.evar_map -> goal -> Evd.evar_map -> goal -> bool
+
+ (* Used for congruence closure *)
+ val new_goal_with : Evd.evar_map -> goal -> Environ.named_context_val -> goal Evd.sigma
+
+ (* Used by the typeclasses *)
+ val nf_evar : Evd.evar_map -> goal -> goal * Evd.evar_map
+
+ (* Goal represented as a type, doesn't take into account section variables *)
+ val abstract_type : Evd.evar_map -> goal -> Term.types
+
+end
diff --git a/proofs/logic.ml b/proofs/logic.ml
index 11155d3698..7c092bb66e 100644
--- a/proofs/logic.ml
+++ b/proofs/logic.ml
@@ -21,7 +21,6 @@ open Reductionops
open Inductive
open Inductiveops
open Typing
-open Proof_trees
open Proof_type
open Typeops
open Type_errors
@@ -326,22 +325,18 @@ let goal_type_of env sigma c =
(if !check then type_of else Retyping.get_type_of ~refresh:true) env sigma c
let rec mk_refgoals sigma goal goalacc conclty trm =
- let env = evar_env goal in
- let hyps = goal.evar_hyps in
- let mk_goal hyps concl = mk_goal hyps concl goal.evar_extra in
-(*
- if not (occur_meta trm) then
- let t'ty = (unsafe_machine env sigma trm).uj_type in
- let _ = conv_leq_goal env sigma trm t'ty conclty in
- (goalacc,t'ty)
- else
-*)
+ let env = Goal.V82.env sigma goal in
+ let hyps = Goal.V82.hyps sigma goal in
+ let mk_goal hyps concl =
+ Goal.V82.mk_goal sigma hyps concl (Goal.V82.extra sigma goal)
+ in
match kind_of_term trm with
| Meta _ ->
let conclty = nf_betaiota sigma conclty in
if !check && occur_meta conclty then
raise (RefinerError (MetaInType conclty));
- (mk_goal hyps conclty)::goalacc, conclty
+ let (gl,ev,sigma) = mk_goal hyps conclty in
+ gl::goalacc, conclty, sigma, ev
| Cast (t,_, ty) ->
check_typability env sigma ty;
@@ -349,30 +344,32 @@ let rec mk_refgoals sigma goal goalacc conclty trm =
mk_refgoals sigma goal goalacc ty t
| App (f,l) ->
- let (acc',hdty) =
+ let (acc',hdty,sigma,applicand) =
match kind_of_term f with
| Ind _ | Const _
when (isInd f or has_polymorphic_type (destConst f)) ->
(* Sort-polymorphism of definition and inductive types *)
goalacc,
- type_of_global_reference_knowing_conclusion env sigma f conclty
+ type_of_global_reference_knowing_conclusion env sigma f conclty,
+ sigma, f
| _ ->
mk_hdgoals sigma goal goalacc f
in
- let (acc'',conclty') =
+ let (acc'',conclty',sigma, args) =
mk_arggoals sigma goal acc' hdty (Array.to_list l) in
check_conv_leq_goal env sigma trm conclty' conclty;
- (acc'',conclty')
+ (acc'',conclty',sigma, Term.mkApp (applicand, Array.of_list args))
- | Case (_,p,c,lf) ->
- let (acc',lbrty,conclty') = mk_casegoals sigma goal goalacc p c in
+ | Case (ci,p,c,lf) ->
+ let (acc',lbrty,conclty',sigma,p',c') = mk_casegoals sigma goal goalacc p c in
check_conv_leq_goal env sigma trm conclty' conclty;
- let acc'' =
+ let (acc'',sigma, rbranches) =
array_fold_left2
- (fun lacc ty fi -> fst (mk_refgoals sigma goal lacc ty fi))
- acc' lbrty lf
+ (fun (lacc,sigma,bacc) ty fi ->
+ let (r,_,s,b') = mk_refgoals sigma goal lacc ty fi in r,s,(b'::bacc))
+ (acc',sigma,[]) lbrty lf
in
- (acc'',conclty')
+ (acc'',conclty',sigma, Term.mkCase (ci,p',c',Array.of_list (List.rev rbranches)))
| _ ->
if occur_meta trm then
@@ -380,70 +377,77 @@ let rec mk_refgoals sigma goal goalacc conclty trm =
let t'ty = goal_type_of env sigma trm in
check_conv_leq_goal env sigma trm t'ty conclty;
- (goalacc,t'ty)
+ (goalacc,t'ty,sigma, trm)
-(* Same as mkREFGOALS but without knowing te type of the term. Therefore,
+(* Same as mkREFGOALS but without knowing the type of the term. Therefore,
* Metas should be casted. *)
and mk_hdgoals sigma goal goalacc trm =
- let env = evar_env goal in
- let hyps = goal.evar_hyps in
- let mk_goal hyps concl = mk_goal hyps concl goal.evar_extra in
+ let env = Goal.V82.env sigma goal in
+ let hyps = Goal.V82.hyps sigma goal in
+ let mk_goal hyps concl =
+ Goal.V82.mk_goal sigma hyps concl (Goal.V82.extra sigma goal) in
match kind_of_term trm with
| Cast (c,_, ty) when isMeta c ->
check_typability env sigma ty;
- (mk_goal hyps (nf_betaiota sigma ty))::goalacc,ty
+ let (gl,ev,sigma) = mk_goal hyps (nf_betaiota sigma ty) in
+ gl::goalacc,ty,sigma,ev
| Cast (t,_, ty) ->
check_typability env sigma ty;
mk_refgoals sigma goal goalacc ty t
| App (f,l) ->
- let (acc',hdty) =
+ let (acc',hdty,sigma,applicand) =
if isInd f or isConst f
& not (array_exists occur_meta l) (* we could be finer *)
then
- (goalacc,type_of_global_reference_knowing_parameters env sigma f l)
+ (goalacc,type_of_global_reference_knowing_parameters env sigma f l,sigma,f)
else mk_hdgoals sigma goal goalacc f
in
- mk_arggoals sigma goal acc' hdty (Array.to_list l)
+ let (acc'',conclty',sigma, args) =
+ mk_arggoals sigma goal acc' hdty (Array.to_list l) in
+ (acc'',conclty',sigma, Term.mkApp (applicand, Array.of_list args))
- | Case (_,p,c,lf) ->
- let (acc',lbrty,conclty') = mk_casegoals sigma goal goalacc p c in
- let acc'' =
+ | Case (ci,p,c,lf) ->
+ let (acc',lbrty,conclty',sigma,p',c') = mk_casegoals sigma goal goalacc p c in
+ let (acc'',sigma,rbranches) =
array_fold_left2
- (fun lacc ty fi -> fst (mk_refgoals sigma goal lacc ty fi))
- acc' lbrty lf
+ (fun (lacc,sigma,bacc) ty fi ->
+ let (r,_,s,b') = mk_refgoals sigma goal lacc ty fi in r,s,(b'::bacc))
+ (acc',sigma,[]) lbrty lf
in
- (acc'',conclty')
+ (acc'',conclty',sigma, Term.mkCase (ci,p',c',Array.of_list (List.rev rbranches)))
| _ ->
if !check && occur_meta trm then
anomaly "refine called with a dependent meta";
- goalacc, goal_type_of env sigma trm
+ goalacc, goal_type_of env sigma trm, sigma, trm
and mk_arggoals sigma goal goalacc funty = function
- | [] -> goalacc,funty
+ | [] -> goalacc,funty,sigma, []
| harg::tlargs as allargs ->
- let t = whd_betadeltaiota (evar_env goal) sigma funty in
+ let t = whd_betadeltaiota (Goal.V82.env sigma goal) sigma funty in
match kind_of_term t with
| Prod (_,c1,b) ->
- let (acc',hargty) = mk_refgoals sigma goal goalacc c1 harg in
- mk_arggoals sigma goal acc' (subst1 harg b) tlargs
+ let (acc',hargty,sigma,arg') = mk_refgoals sigma goal goalacc c1 harg in
+ let (acc'',fty, sigma', args) =
+ mk_arggoals sigma goal acc' (subst1 harg b) tlargs in
+ (acc'',fty,sigma',arg'::args)
| LetIn (_,c1,_,b) ->
mk_arggoals sigma goal goalacc (subst1 c1 b) allargs
| _ -> raise (RefinerError (CannotApply (t,harg)))
and mk_casegoals sigma goal goalacc p c =
- let env = evar_env goal in
- let (acc',ct) = mk_hdgoals sigma goal goalacc c in
- let (acc'',pt) = mk_hdgoals sigma goal acc' p in
+ let env = Goal.V82.env sigma goal in
+ let (acc',ct,sigma,c') = mk_hdgoals sigma goal goalacc c in
+ let (acc'',pt,sigma,p') = mk_hdgoals sigma goal acc' p in
let indspec =
try find_mrectype env sigma ct
with Not_found -> anomaly "mk_casegoals" in
let (lbrty,conclty) =
type_case_branches_with_names env indspec p c in
- (acc'',lbrty,conclty)
+ (acc'',lbrty,conclty,sigma,p',c')
let convert_hyp sign sigma (id,b,bt as d) =
@@ -461,18 +465,6 @@ let convert_hyp sign sigma (id,b,bt as d) =
d) in
reorder_val_context env sign' !reorder
-(* Normalizing evars in a goal. Called by tactic Local_constraints
- (i.e. when the sigma of the proof tree changes). Detect if the
- goal is unchanged *)
-let norm_goal sigma gl =
- let red_fun = Evarutil.nf_evar sigma in
- let ncl = red_fun gl.evar_concl in
- let ngl =
- { gl with
- evar_concl = ncl;
- evar_hyps = map_named_val red_fun gl.evar_hyps } in
- if Evd.eq_evar_info ngl gl then None else Some ngl
-
(************************************************************************)
@@ -480,10 +472,12 @@ let norm_goal sigma gl =
(* Primitive tactics are handled here *)
let prim_refiner r sigma goal =
- let env = evar_env goal in
- let sign = goal.evar_hyps in
- let cl = goal.evar_concl in
- let mk_goal hyps concl = mk_goal hyps concl goal.evar_extra in
+ let env = Goal.V82.env sigma goal in
+ let sign = Goal.V82.hyps sigma goal in
+ let cl = Goal.V82.concl sigma goal in
+ let mk_goal hyps concl =
+ Goal.V82.mk_goal sigma hyps concl (Goal.V82.extra sigma goal)
+ in
match r with
(* Logical rules *)
| Intro id ->
@@ -491,19 +485,23 @@ let prim_refiner r sigma goal =
error "New variable is already declared";
(match kind_of_term (strip_outer_cast cl) with
| Prod (_,c1,b) ->
- let sg = mk_goal (push_named_context_val (id,None,c1) sign)
+ let (sg,ev,sigma) = mk_goal (push_named_context_val (id,None,c1) sign)
(subst1 (mkVar id) b) in
+ let sigma =
+ Goal.V82.partial_solution sigma goal (mkNamedLambda id c1 ev) in
([sg], sigma)
| LetIn (_,c1,t1,b) ->
- let sg =
+ let (sg,ev,sigma) =
mk_goal (push_named_context_val (id,Some c1,t1) sign)
(subst1 (mkVar id) b) in
+ let sigma =
+ Goal.V82.partial_solution sigma goal (mkNamedLetIn id c1 t1 ev) in
([sg], sigma)
| _ ->
raise (RefinerError IntroNeedsProduct))
| Cut (b,replace,id,t) ->
- let sg1 = mk_goal sign (nf_betaiota sigma t) in
+ let (sg1,ev1,sigma) = mk_goal sign (nf_betaiota sigma t) in
let sign,cl,sigma =
if replace then
let nexthyp = get_hyp_after id (named_context_of_val sign) in
@@ -515,7 +513,10 @@ let prim_refiner r sigma goal =
(if !check && mem_named_context id (named_context_of_val sign) then
error "New variable is already declared";
push_named_context_val (id,None,t) sign,cl,sigma) in
- let sg2 = mk_goal sign cl in
+ let (sg2,ev2,sigma) =
+ Goal.V82.mk_goal sigma sign cl (Goal.V82.extra sigma goal) in
+ let oterm = Term.mkApp (Term.mkNamedLambda id t ev2 , [| ev1 |]) in
+ let sigma = Goal.V82.partial_solution sigma goal oterm in
if b then ([sg1;sg2],sigma) else ([sg2;sg1],sigma)
| FixRule (f,n,rest,j) ->
@@ -545,9 +546,24 @@ let prim_refiner r sigma goal =
("Name "^string_of_id f^" already used in the environment");
mk_sign (push_named_context_val (f,None,ar) sign) oth
| [] ->
- List.map (fun (_,_,c) -> mk_goal sign c) all
+ Goal.list_map (fun sigma (_,_,c) ->
+ let (gl,ev,sig')=
+ Goal.V82.mk_goal sigma sign c
+ (Goal.V82.extra sigma goal)
+ in ((gl,ev),sig'))
+ all sigma
in
- (mk_sign sign all, sigma)
+ let (gls_evs,sigma) = mk_sign sign all in
+ let (gls,evs) = List.split gls_evs in
+ let ids = List.map pi1 all in
+ let evs = List.map (Term.subst_vars (List.rev ids)) evs in
+ let indxs = Array.of_list (List.map (fun n -> n-1) (List.map pi2 all)) in
+ let funnames = Array.of_list (List.map (fun i -> Name i) ids) in
+ let typarray = Array.of_list (List.map pi3 all) in
+ let bodies = Array.of_list evs in
+ let oterm = Term.mkFix ((indxs,0),(funnames,typarray,bodies)) in
+ let sigma = Goal.V82.partial_solution sigma goal oterm in
+ (gls,sigma)
| Cofix (f,others,j) ->
let rec check_is_coind env cl =
@@ -572,32 +588,55 @@ let prim_refiner r sigma goal =
with
| Not_found ->
mk_sign (push_named_context_val (f,None,ar) sign) oth)
- | [] -> List.map (fun (_,c) -> mk_goal sign c) all
+ | [] -> Goal.list_map (fun sigma(_,c) ->
+ let (gl,ev,sigma) =
+ Goal.V82.mk_goal sigma sign c
+ (Goal.V82.extra sigma goal)
+ in
+ ((gl,ev),sigma))
+ all sigma
in
- (mk_sign sign all, sigma)
+ let (gls_evs,sigma) = mk_sign sign all in
+ let (gls,evs) = List.split gls_evs in
+ let (ids,types) = List.split all in
+ let evs = List.map (Term.subst_vars (List.rev ids)) evs in
+ let funnames = Array.of_list (List.map (fun i -> Name i) ids) in
+ let typarray = Array.of_list types in
+ let bodies = Array.of_list evs in
+ let oterm = Term.mkCoFix (0,(funnames,typarray,bodies)) in
+ let sigma = Goal.V82.partial_solution sigma goal oterm in
+ (gls,sigma)
| Refine c ->
check_meta_variables c;
- let (sgl,cl') = mk_refgoals sigma goal [] cl c in
+ let (sgl,cl',sigma,oterm) = mk_refgoals sigma goal [] cl c in
let sgl = List.rev sgl in
+ let sigma = Goal.V82.partial_solution sigma goal oterm in
(sgl, sigma)
(* Conversion rules *)
| Convert_concl (cl',_) ->
check_typability env sigma cl';
if (not !check) || is_conv_leq env sigma cl' cl then
- let sg = mk_goal sign cl' in
+ let (sg,ev,sigma) = mk_goal sign cl' in
+ let sigma = Goal.V82.partial_solution sigma goal ev in
([sg], sigma)
else
error "convert-concl rule passed non-converting term"
| Convert_hyp (id,copt,ty) ->
- ([mk_goal (convert_hyp sign sigma (id,copt,ty)) cl], sigma)
+ let (gl,ev,sigma) = mk_goal (convert_hyp sign sigma (id,copt,ty)) cl in
+ let sigma = Goal.V82.partial_solution sigma goal ev in
+ ([gl], sigma)
(* And now the structural rules *)
| Thin ids ->
let (hyps,concl,nsigma) = clear_hyps sigma ids sign cl in
- ([mk_goal hyps concl], nsigma)
+ let (gl,ev,sigma) =
+ Goal.V82.mk_goal nsigma hyps concl (Goal.V82.extra nsigma goal)
+ in
+ let sigma = Goal.V82.partial_solution sigma goal ev in
+ ([gl], sigma)
| ThinBody ids ->
let clear_aux env id =
@@ -606,7 +645,8 @@ let prim_refiner r sigma goal =
env'
in
let sign' = named_context_val (List.fold_left clear_aux env ids) in
- let sg = mk_goal sign' cl in
+ let (sg,ev,sigma) = mk_goal sign' cl in
+ let sigma = Goal.V82.partial_solution sigma goal ev in
([sg], sigma)
| Move (withdep, hfrom, hto) ->
@@ -614,11 +654,15 @@ let prim_refiner r sigma goal =
split_sign hfrom hto (named_context_of_val sign) in
let hyps' =
move_hyp withdep toleft (left,declfrom,right) hto in
- ([mk_goal hyps' cl], sigma)
+ let (gl,ev,sigma) = mk_goal hyps' cl in
+ let sigma = Goal.V82.partial_solution sigma goal ev in
+ ([gl], sigma)
| Order ord ->
let hyps' = reorder_val_context env sign ord in
- ([mk_goal hyps' cl], sigma)
+ let (gl,ev,sigma) = mk_goal hyps' cl in
+ let sigma = Goal.V82.partial_solution sigma goal ev in
+ ([gl], sigma)
| Rename (id1,id2) ->
if !check & id1 <> id2 &&
@@ -626,12 +670,19 @@ let prim_refiner r sigma goal =
error ((string_of_id id2)^" is already used.");
let sign' = rename_hyp id1 id2 sign in
let cl' = replace_vars [id1,mkVar id2] cl in
- ([mk_goal sign' cl'], sigma)
+ let (gl,ev,sigma) = mk_goal sign' cl' in
+ let ev = Term.replace_vars [(id2,mkVar id1)] ev in
+ let sigma = Goal.V82.partial_solution sigma goal ev in
+ ([gl], sigma)
| Change_evars ->
- match norm_goal sigma goal with
- Some ngl -> ([ngl],sigma)
- | None -> ([goal], sigma)
+ (* spiwack: a priori [Change_evars] is now devoid of operational content.
+ The new proof engine keeping the evar_map up to date at all time.
+ As a compatibility mesure I leave the rule.
+ It is possible that my assumption is wrong and some uses of
+ [Change_evars] are not subsumed by the new engine. In which
+ case something has to be done here. (Feb. 2010) *)
+ ([goal],sigma)
(************************************************************************)
(************************************************************************)
@@ -671,77 +722,3 @@ let proof_variable_index x =
| [] -> raise Not_found
in
aux 1
-
-let prim_extractor subfun vl pft =
- let cl = pft.goal.evar_concl in
- match pft.ref with
- | Some (Prim (Intro id), [spf]) ->
- (match kind_of_term (strip_outer_cast cl) with
- | Prod (_,ty,_) ->
- let cty = subst_proof_vars vl ty in
- mkLambda (Name id, cty, subfun (add_proof_var id vl) spf)
- | LetIn (_,b,ty,_) ->
- let cb = subst_proof_vars vl b in
- let cty = subst_proof_vars vl ty in
- mkLetIn (Name id, cb, cty, subfun (add_proof_var id vl) spf)
- | _ -> error "Incomplete proof!")
-
- | Some (Prim (Cut (b,_,id,t)),[spf1;spf2]) ->
- let spf1, spf2 = if b then spf1, spf2 else spf2, spf1 in
- mkLetIn (Name id,subfun vl spf1,subst_proof_vars vl t,
- subfun (add_proof_var id vl) spf2)
-
- | Some (Prim (FixRule (f,n,others,j)),spfl) ->
- let firsts,lasts = list_chop j others in
- let all = Array.of_list (firsts@(f,n,cl)::lasts) in
- let lcty = Array.map (fun (_,_,ar) -> subst_proof_vars vl ar) all in
- let names = Array.map (fun (f,_,_) -> Name f) all in
- let vn = Array.map (fun (_,n,_) -> n-1) all in
- let newvl = List.fold_left (fun vl (id,_,_) -> add_proof_var id vl)
- (add_proof_var f vl) others in
- let lfix = Array.map (subfun newvl) (Array.of_list spfl) in
- mkFix ((vn,j),(names,lcty,lfix))
-
- | Some (Prim (Cofix (f,others,j)),spfl) ->
- let firsts,lasts = list_chop j others in
- let all = Array.of_list (firsts@(f,cl)::lasts) in
- let lcty = Array.map (fun (_,ar) -> subst_proof_vars vl ar) all in
- let names = Array.map (fun (f,_) -> Name f) all in
- let newvl = List.fold_left (fun vl (id,_)-> add_proof_var id vl)
- (add_proof_var f vl) others in
- let lfix = Array.map (subfun newvl) (Array.of_list spfl) in
- mkCoFix (j,(names,lcty,lfix))
-
- | Some (Prim (Refine c),spfl) ->
- let mvl = collect_meta_variables c in
- let metamap = List.combine mvl (List.map (subfun vl) spfl) in
- let cc = subst_proof_vars vl c in
- plain_instance metamap cc
-
- (* Structural and conversion rules do not produce any proof *)
- | Some (Prim (Convert_concl (t,k)),[pf]) ->
- if k = DEFAULTcast then subfun vl pf
- else mkCast (subfun vl pf,k,subst_proof_vars vl cl)
- | Some (Prim (Convert_hyp _),[pf]) ->
- subfun vl pf
-
- | Some (Prim (Thin _),[pf]) ->
- (* No need to make ids Anon in vl: subst_proof_vars take the most recent*)
- subfun vl pf
-
- | Some (Prim (ThinBody _),[pf]) ->
- subfun vl pf
-
- | Some (Prim (Move _|Order _),[pf]) ->
- subfun vl pf
-
- | Some (Prim (Rename (id1,id2)),[pf]) ->
- subfun (rebind id1 id2 vl) pf
-
- | Some (Prim Change_evars, [pf]) ->
- subfun vl pf
-
- | Some _ -> anomaly "prim_extractor"
-
- | None-> error "prim_extractor handed incomplete proof"
-
diff --git a/proofs/logic.mli b/proofs/logic.mli
index 0d56da382a..560e577366 100644
--- a/proofs/logic.mli
+++ b/proofs/logic.mli
@@ -38,9 +38,6 @@ val prim_refiner : prim_rule -> evar_map -> goal -> goal list * evar_map
type proof_variable
-val prim_extractor :
- (proof_variable list -> proof_tree -> constr)
- -> proof_variable list -> proof_tree -> constr
val proof_variable_index : identifier -> proof_variable list -> int
diff --git a/proofs/pfedit.ml b/proofs/pfedit.ml
index f3658ad4b6..6da73c2f67 100644
--- a/proofs/pfedit.ml
+++ b/proofs/pfedit.ml
@@ -20,322 +20,115 @@ open Environ
open Evd
open Typing
open Refiner
-open Proof_trees
open Tacexpr
open Proof_type
open Lib
open Safe_typing
-(*********************************************************************)
-(* Managing the proofs state *)
-(* Mainly contributed by C. Murthy *)
-(*********************************************************************)
+let refining = Proof_global.there_are_pending_proofs
+let check_no_pending_proofs = Proof_global.check_no_pending_proof
-type lemma_possible_guards = int list list
+let get_current_proof_name = Proof_global.get_current_proof_name
+let get_all_proof_names = Proof_global.get_all_proof_names
-type proof_topstate = {
- mutable top_end_tac : tactic option;
- top_init_tac : tactic option;
- top_compute_guard : lemma_possible_guards;
- top_goal : goal;
- top_strength : Decl_kinds.goal_kind;
- top_hook : declaration_hook }
+type lemma_possible_guards = Proof_global.lemma_possible_guards
-let proof_edits =
- (Edit.empty() : (identifier,pftreestate,proof_topstate) Edit.t)
+let delete_proof = Proof_global.discard
+let delete_current_proof = Proof_global.discard_current
+let delete_all_proofs = Proof_global.discard_all
-let get_all_proof_names () = Edit.dom proof_edits
+let undo n =
+ let p = Proof_global.give_me_the_proof () in
+ for i = 1 to n do
+ Proof.undo p
+ done
-let msg_proofs use_resume =
- match Edit.dom proof_edits with
- | [] -> (spc () ++ str"(No proof-editing in progress).")
- | l -> (str"." ++ fnl () ++ str"Proofs currently edited:" ++ spc () ++
- (prlist_with_sep pr_spc pr_id (get_all_proof_names ())) ++
- str"." ++
- (if use_resume then (fnl () ++ str"Use \"Resume\" first.")
- else (mt ()))
-)
-
-let undo_default = 50
-let undo_limit = ref undo_default
-
-(*********************************************************************)
-(* Functions for decomposing and modifying the proof state *)
-(*********************************************************************)
-
-let get_state () =
- match Edit.read proof_edits with
- | None -> errorlabstrm "Pfedit.get_state"
- (str"No focused proof" ++ msg_proofs true)
- | Some(_,pfs,ts) -> (pfs,ts)
-
-let get_topstate () = snd(get_state())
-let get_pftreestate () = fst(get_state())
-
-let get_end_tac () = let ts = get_topstate () in ts.top_end_tac
-
-let get_goal_context n =
- let pftree = get_pftreestate () in
- let goal = nth_goal_of_pftreestate n pftree in
- (project goal, pf_env goal)
-
-let get_current_goal_context () = get_goal_context 1
-
-let set_current_proof = Edit.focus proof_edits
-
-let resume_proof (loc,id) =
+let current_proof_depth () =
try
- Edit.focus proof_edits id
- with Invalid_argument "Edit.focus" ->
- user_err_loc(loc,"Pfedit.set_proof",str"No such proof" ++ msg_proofs false)
+ let p = Proof_global.give_me_the_proof () in
+ Proof.V82.depth p
+ with Proof_global.NoCurrentProof -> -1
-let suspend_proof () =
+(* [undo_todepth n] resets the proof to its nth step (does [undo (d-n)] where d
+ is the depth of the focus stack). *)
+let undo_todepth n =
try
- Edit.unfocus proof_edits
- with Invalid_argument "Edit.unfocus" ->
- errorlabstrm "Pfedit.suspend_current_proof"
- (str"No active proof" ++ (msg_proofs true))
-
-let resume_last_proof () =
- match (Edit.last_focused proof_edits) with
- | None ->
- errorlabstrm "resume_last" (str"No proof-editing in progress.")
- | Some p ->
- Edit.focus proof_edits p
+ undo ((current_proof_depth ()) - n )
+ with Proof_global.NoCurrentProof when n=0 -> ()
-let get_current_proof_name () =
- match Edit.read proof_edits with
- | None ->
- errorlabstrm "Pfedit.get_proof"
- (str"No focused proof" ++ msg_proofs true)
- | Some(na,_,_) -> na
+let set_undo _ = ()
+let get_undo _ = None
-let add_proof (na,pfs,ts) =
- Edit.create proof_edits (na,pfs,ts,!undo_limit+1)
-let delete_proof_gen = Edit.delete proof_edits
-
-let delete_proof (loc,id) =
- try
- delete_proof_gen id
- with (UserError ("Edit.delete",_)) ->
- user_err_loc
- (loc,"Pfedit.delete_proof",str"No such proof" ++ msg_proofs false)
-
-let mutate f =
- try
- Edit.mutate proof_edits (fun _ pfs -> f pfs)
- with Invalid_argument "Edit.mutate" ->
- errorlabstrm "Pfedit.mutate"
- (str"No focused proof" ++ msg_proofs true)
-
-let start (na,ts) =
- let pfs = mk_pftreestate ts.top_goal in
- let pfs = Option.fold_right solve_pftreestate ts.top_init_tac pfs in
- add_proof(na,pfs,ts)
+let start_proof id str hyps c ?init_tac ?compute_guard hook =
+ let goals = [ (Global.env_of_context hyps , c) ] in
+ let init_tac = Option.map Proofview.V82.tactic init_tac in
+ Proof_global.start_proof id str goals ?compute_guard hook;
+ Option.iter Proof_global.run_tactic init_tac
let restart_proof () =
- match Edit.read proof_edits with
- | None ->
- errorlabstrm "Pfedit.restart"
- (str"No focused proof to restart" ++ msg_proofs true)
- | Some(na,_,ts) ->
- delete_proof_gen na;
- start (na,ts);
- set_current_proof na
-
-let proof_term () =
- extract_pftreestate (get_pftreestate())
-
-(* Detect is one has completed a subtree of the initial goal *)
-let subtree_solved () =
- let pts = get_pftreestate () in
- is_complete_proof (proof_of_pftreestate pts) &
- not (is_top_pftreestate pts)
-
-let tree_solved () =
- let pts = get_pftreestate () in
- is_complete_proof (proof_of_pftreestate pts)
-
-let top_tree_solved () =
- let pts = get_pftreestate () in
- is_complete_proof (proof_of_pftreestate (top_of_tree pts))
-
-(*********************************************************************)
-(* Undo functions *)
-(*********************************************************************)
-
-let set_undo = function
- | None -> undo_limit := undo_default
- | Some n ->
- if n>=0 then
- undo_limit := n
- else
- error "Cannot set a negative undo limit"
-
-let get_undo () = Some !undo_limit
-
-let undo n =
- try
- Edit.undo proof_edits n;
- (* needed because the resolution of a subtree is done in 2 steps
- then a sequence of undo can lead to a focus on a completely solved
- subtree - this solution only works properly if undoing one step *)
- if subtree_solved() then Edit.undo proof_edits 1
- with (Invalid_argument "Edit.undo") ->
- errorlabstrm "Pfedit.undo" (str"No focused proof" ++ msg_proofs true)
-
-(* Undo current focused proof to reach depth [n]. This is used in
- [vernac_backtrack]. *)
-let undo_todepth n =
- try
- Edit.undo_todepth proof_edits n
- with (Invalid_argument "Edit.undo") ->
- errorlabstrm "Pfedit.undo" (str"No focused proof" ++ msg_proofs true)
-
-(* Return the depth of the current focused proof stack, this is used
- to put informations in coq prompt (in emacs mode). *)
-let current_proof_depth() =
- try
- Edit.depth proof_edits
- with (Invalid_argument "Edit.depth") -> -1
-
-(*********************************************************************)
-(* Proof cooking *)
-(*********************************************************************)
+ let p = Proof_global.give_me_the_proof () in
+ try while true do
+ Proof.undo p
+ done with Proof.EmptyUndoStack -> ()
+
+let resume_last_proof () = Proof_global.resume_last ()
+let resume_proof (_,id) = Proof_global.resume id
+let suspend_proof () = Proof_global.suspend ()
+
+let cook_proof hook =
+ let prf = Proof_global.give_me_the_proof () in
+ hook prf;
+ match Proof_global.close_proof () with
+ | (i,([e],cg,str,h)) -> (i,(e,cg,str,h))
+ | _ -> Util.anomaly "Pfedit.cook_proof: more than one proof term."
let xml_cook_proof = ref (fun _ -> ())
let set_xml_cook_proof f = xml_cook_proof := f
-let cook_proof k =
- let (pfs,ts) = get_state()
- and ident = get_current_proof_name () in
- let {evar_concl=concl} = ts.top_goal
- and strength = ts.top_strength in
- let pfterm = extract_pftreestate pfs in
- !xml_cook_proof (strength,pfs);
- k pfs;
- (ident,
- ({ const_entry_body = pfterm;
- const_entry_type = Some concl;
- const_entry_opaque = true;
- const_entry_boxed = false},
- ts.top_compute_guard, strength, ts.top_hook))
-
-let current_proof_statement () =
- let ts = get_topstate() in
- (get_current_proof_name (), ts.top_strength,
- ts.top_goal.evar_concl, ts.top_hook)
-
-(*********************************************************************)
-(* Abort functions *)
-(*********************************************************************)
-
-let refining () = [] <> (Edit.dom proof_edits)
-
-let check_no_pending_proofs () =
- if refining () then
- errorlabstrm "check_no_pending_proofs"
- (str"Proof editing in progress" ++ (msg_proofs false) ++ fnl() ++
- str"Use \"Abort All\" first or complete proof(s).")
+let get_pftreestate () =
+ Proof_global.give_me_the_proof ()
-let delete_current_proof () = delete_proof_gen (get_current_proof_name ())
-
-let delete_all_proofs () = Edit.clear proof_edits
-
-(*********************************************************************)
-(* Modifying the end tactic of the current profftree *)
-(*********************************************************************)
let set_end_tac tac =
- let top = get_topstate () in
- top.top_end_tac <- Some tac
-
-(*********************************************************************)
-(* Modifying the current prooftree *)
-(*********************************************************************)
-
-let start_proof na str sign concl ?init_tac ?(compute_guard=[]) hook =
- let top_goal = mk_goal sign concl None in
- let ts = {
- top_end_tac = None;
- top_init_tac = init_tac;
- top_compute_guard = compute_guard;
- top_goal = top_goal;
- top_strength = str;
- top_hook = hook}
- in
- start(na,ts);
- set_current_proof na
-
-let solve_nth k tac =
- let pft = get_pftreestate () in
- if not (List.mem (-1) (cursor_of_pftreestate pft)) then
- mutate (solve_nth_pftreestate k tac)
- else
- error "cannot apply a tactic when we are descended behind a tactic-node"
-
-let by tac = mutate (solve_pftreestate tac)
-
-let instantiate_nth_evar_com n c =
- mutate (Evar_refiner.instantiate_pf_com n c)
+ let tac = Proofview.V82.tactic tac in
+ Proof_global.set_endline_tactic tac
-let traverse n = mutate (traverse n)
-
-(* [traverse_to path]
-
- Traverses the current proof to get to the location specified by
- [path].
-
- ALGORITHM: The algorithm works on reversed paths. One just has to remove
- the common part on the reversed paths.
-
-*)
-
-let common_ancestor l1 l2 =
- let rec common_aux l1 l2 =
- match l1, l2 with
- | a1::l1', a2::l2' when a1 = a2 -> common_aux l1' l2'
- | _, _ -> List.rev l1, List.length l2
- in
- common_aux (List.rev l1) (List.rev l2)
-
-let rec traverse_up = function
- | 0 -> (function pf -> pf)
- | n -> (function pf -> Refiner.traverse 0 (traverse_up (n - 1) pf))
-
-let rec traverse_down = function
- | [] -> (function pf -> pf)
- | n::l -> (function pf -> Refiner.traverse n (traverse_down l pf))
-
-let traverse_to path =
- let up_and_down path pfs =
- let cursor = cursor_of_pftreestate pfs in
- let down_list, up_count = common_ancestor path cursor in
- traverse_down down_list (traverse_up up_count pfs)
- in
- mutate (up_and_down path)
-
-(* traverse the proof tree until it reach the nth subgoal *)
-let traverse_nth_goal n = mutate (nth_unproven n)
-
-let traverse_prev_unproven () = mutate prev_unproven
-
-let traverse_next_unproven () = mutate next_unproven
+let get_goal_context i =
+ try
+ let p = Proof_global.give_me_the_proof () in
+ let { it=goals ; sigma = sigma } = Proof.V82.subgoals p in
+ let goal = List.nth goals (i-1) in
+ (sigma, Refiner.pf_env { it=goal ; sigma=sigma })
+ with Proof_global.NoCurrentProof -> Util.error "No focused proof."
-(* The goal focused on *)
-let focus_n = ref 0
-let make_focus n = focus_n := n
-let focus () = !focus_n
-let focused_goal () = let n = !focus_n in if n=0 then 1 else n
+let get_current_goal_context () = get_goal_context 1
-let reset_top_of_tree () =
- mutate top_of_tree
+let current_proof_statement () =
+ match Proof_global.V82.get_current_initial_conclusions () with
+ | (id,([concl],strength,hook)) -> id,strength,concl,hook
+ | _ -> Util.anomaly "Pfedit.current_proof_statement: more than one statement"
+
+let solve_nth ?(with_end_tac=false) gi tac =
+ try
+ let tac = Proofview.V82.tactic tac in
+ let tac = if with_end_tac then
+ Proof_global.with_end_tac tac
+ else
+ tac
+ in
+ Proof_global.run_tactic (Proofview.tclFOCUS gi gi tac)
+ with
+ | Proof_global.NoCurrentProof -> Util.error "No focused proof"
+ | Proofview.IndexOutOfRange | Failure "list_chop" ->
+ let msg = str "No such goal: " ++ int gi ++ str "." in
+ Util.errorlabstrm "" msg
+
+let by = solve_nth 1
+
+let instantiate_nth_evar_com n com =
+ let pf = Proof_global.give_me_the_proof () in
+ Proof.V82.instantiate_evar n com pf
-let reset_top_of_script () =
- mutate (fun pts ->
- try
- up_until_matching_rule is_proof_instr pts
- with Not_found -> top_of_tree pts)
(**********************************************************************)
(* Shortcut to build a term using tactics *)
diff --git a/proofs/pfedit.mli b/proofs/pfedit.mli
index acb852471d..9e24061d32 100644
--- a/proofs/pfedit.mli
+++ b/proofs/pfedit.mli
@@ -20,6 +20,7 @@ open Tacmach
open Tacexpr
(*i*)
+
(*s Several proofs can be opened simultaneously but at most one is
focused at some time. The following functions work by side-effect
on current set of open proofs. In this module, ``proofs'' means an
@@ -57,17 +58,16 @@ val delete_all_proofs : unit -> unit
val undo : int -> unit
-(* Same as undo, but undoes the current proof stack to reach depth
- [n]. This is used in [vernac_backtrack]. *)
-val undo_todepth : int -> unit
+(* [undo_todepth n] resets the proof to its nth step (does [undo (d-n)] where d
+ is the depth of the undo stack). *)
+val undo_todepth : int -> unit
(* Returns the depth of the current focused proof stack, this is used
to put informations in coq prompt (in emacs mode). *)
val current_proof_depth: unit -> int
-(* [set_undo (Some n)] sets the size of the ``undo'' stack; [set_undo None]
- sets the size to the default value (12) *)
-
+(* [set_undo (Some n)] used to set the size of the ``undo'' stack.
+ These function now do nothing and will disapear. *)
val set_undo : int option -> unit
val get_undo : unit -> int option
@@ -78,7 +78,7 @@ val get_undo : unit -> int option
systematically apply at initialization time (e.g. to start the
proof of mutually dependent theorems) *)
-type lemma_possible_guards = int list list
+type lemma_possible_guards = Proof_global.lemma_possible_guards
val start_proof :
identifier -> goal_kind -> named_context_val -> constr ->
@@ -110,22 +110,18 @@ val suspend_proof : unit -> unit
it fails if there is no current proof of if it is not completed;
it also tells if the guardness condition has to be inferred. *)
-val cook_proof : (Refiner.pftreestate -> unit) ->
+val cook_proof : (Proof.proof -> unit) ->
identifier *
- (Entries.definition_entry * lemma_possible_guards * goal_kind *
- declaration_hook)
+ (Entries.definition_entry * lemma_possible_guards * goal_kind *
+ declaration_hook)
(* To export completed proofs to xml *)
-val set_xml_cook_proof : (goal_kind * pftreestate -> unit) -> unit
+val set_xml_cook_proof : (goal_kind * Proof.proof -> unit) -> unit
-(*s [get_pftreestate ()] returns the current focused pending proof or
+(*s [get_Proof.proof ()] returns the current focused pending proof or
raises [UserError "no focused proof"] *)
-val get_pftreestate : unit -> pftreestate
-
-(* [get_end_tac ()] returns the current tactic to apply to all new subgoal *)
-
-val get_end_tac : unit -> tactic option
+val get_pftreestate : unit -> Proof.proof
(* [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
@@ -160,7 +156,7 @@ val set_end_tac : tactic -> unit
current focused proof or raises a UserError if no proof is focused or
if there is no [n]th subgoal *)
-val solve_nth : int -> tactic -> unit
+val solve_nth : ?with_end_tac:bool -> int -> tactic -> unit
(* [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
@@ -175,31 +171,6 @@ val by : tactic -> unit
val instantiate_nth_evar_com : int -> Topconstr.constr_expr -> unit
-(*s To deal with subgoal focus *)
-
-val make_focus : int -> unit
-val focus : unit -> int
-val focused_goal : unit -> int
-val subtree_solved : unit -> bool
-val tree_solved : unit -> bool
-val top_tree_solved : unit -> bool
-
-val reset_top_of_tree : unit -> unit
-val reset_top_of_script : unit -> unit
-
-val traverse : int -> unit
-val traverse_nth_goal : int -> unit
-val traverse_next_unproven : unit -> unit
-val traverse_prev_unproven : unit -> unit
-
-
-(* These two functions make it possible to implement more elaborate
- proof and goal management, as it is done, for instance in pcoq *)
-
-val traverse_to : int list -> unit
-val mutate : (pftreestate -> pftreestate) -> unit
-
-
(* [build_by_tactic typ tac] returns a term of type [typ] by calling [tac] *)
val build_constant_by_tactic : named_context_val -> types -> tactic ->
diff --git a/proofs/proof.ml b/proofs/proof.ml
new file mode 100644
index 0000000000..0c298cc630
--- /dev/null
+++ b/proofs/proof.ml
@@ -0,0 +1,294 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* $Id$ *)
+
+(* Module defining the last essential tiles of interactive proofs.
+ The features of the Proof module are undoing and focusing.
+ A proof is a mutable object, it contains a proofview, and some information
+ to be able to undo actions, and to unfocus the current view. All three
+ of these being meant to evolve.
+ - Proofview: a proof is primarily the data of the current view.
+ That which is shown to the user (as a remainder, a proofview
+ is mainly the logical state of the proof, together with the
+ currently focused goals).
+ - Focus: a proof has a focus stack: the top of the stack contains
+ the context in which to unfocus the current view to a view focused
+ with the rest of the stack.
+ In addition, this contains, for each of the focus context, a
+ "focus kind" and a "focus condition" (in practice, and for modularity,
+ the focus kind is actually stored inside the condition). To unfocus, one
+ needs to know the focus kind, and the condition (for instance "no condition" or
+ the proof under focused must be complete) must be met.
+ - Undo: since proofviews and focus stacks are immutable objects,
+ it could suffice to hold the previous states, to allow to return to the past.
+ However, we also allow other modules to do actions that can be undone.
+ Therefore the undo stack stores action to be ran to undo.
+*)
+
+open Term
+
+type focus_kind = int
+type focus_condition = focus_kind -> Proofview.proofview -> bool
+
+let next_kind = ref 0
+let new_focus_kind () =
+ let r = !next_kind in
+ incr next_kind;
+ r
+
+(* To be authorized to unfocus one must meet the condition prescribed by
+ the action which focused.*)
+(* spiwack: we could consider having a list of authorized focus_kind instead
+ of just one, if anyone needs it *)
+(* [no_cond] only checks that the unfocusing command uses the right
+ [focus_kind]. *)
+let no_cond k0 k _ = k0 = k
+(* [done_cond] checks that the unfocusing command uses the right [focus_kind]
+ and that the focused proofview is complete. *)
+let done_cond k0 k p = k0 = k && Proofview.finished p
+
+
+(* Subpart of the type of proofs. It contains the parts of the proof which
+ are under control of the undo mechanism *)
+type proof_state = {
+ (* Current focused proofview *)
+ proofview: Proofview.proofview;
+ (* History of the focusings, provides information on how
+ to unfocus the proof.
+ The list is empty when the proof is fully unfocused. *)
+ focus_stack: (focus_condition*Proofview.focus_context) list;
+ (* Extra information which can be freely used to create new behaviours. *)
+ intel: Store.t
+}
+
+type proof_info = {
+ mutable endline_tactic : unit Proofview.tactic ;
+ initial_conclusions : Term.types list
+}
+
+type undo_action =
+ | State of proof_state
+ | Effect of (unit -> unit)
+
+type proof = { (* current proof_state *)
+ mutable state : proof_state;
+ (* The undo stack *)
+ mutable undo_stack : undo_action list;
+ info : proof_info
+ }
+
+
+(*** General proof functions ***)
+
+let start goals =
+ { state = { proofview = Proofview.init goals ;
+ focus_stack = [] ;
+ intel = Store.empty} ;
+ undo_stack = [] ;
+ info = { endline_tactic = Proofview.tclUNIT ();
+ initial_conclusions = List.map snd goals }
+ }
+
+let rec unroll_focus pv = function
+ | (_,ctx)::stk -> unroll_focus (Proofview.unfocus ctx pv) stk
+ | [] -> pv
+
+(* spiwack: a proof is considered completed even if its still focused, if the focus
+ doesn't hide any goal. *)
+let is_done p =
+ Proofview.finished p.state.proofview &&
+ Proofview.finished (unroll_focus p.state.proofview p.state.focus_stack)
+
+(* spiwack: for compatibility with <= 8.2 proof engine *)
+let has_unresolved_evar p =
+ Proofview.V82.has_unresolved_evar p.state.proofview
+
+(* Returns the list of partial proofs to initial goals *)
+let partial_proof p =
+ List.map fst (Proofview.return p.state.proofview)
+
+exception UnfinishedProof
+exception HasUnresolvedEvar
+let return p =
+ if not (is_done p) then
+ raise UnfinishedProof
+ else if has_unresolved_evar p then
+ (* spiwack: for compatibility with <= 8.2 proof engine *)
+ raise HasUnresolvedEvar
+ else
+ Proofview.return p.state.proofview
+
+(*** The following functions implement the basic internal mechanisms
+ of proofs, they are not meant to be exported in the .mli ***)
+
+(* An auxiliary function to push a {!focus_context} on the focus stack. *)
+let push_focus kind context pr =
+ pr.state <- { pr.state with focus_stack = (kind,context)::pr.state.focus_stack }
+
+exception FullyUnfocused
+
+(* An auxiliary function to read the kind of the next focusing step *)
+let cond_of_focus pr =
+ match pr.state.focus_stack with
+ | (cond,_)::_ -> cond
+ | _ -> raise FullyUnfocused
+
+(* An auxiliary function to pop and read the last {!Proofview.focus_context}
+ on the focus stack. *)
+let pop_focus pr =
+ match pr.state.focus_stack with
+ | focus::other_focuses ->
+ pr.state <- { pr.state with focus_stack = other_focuses }; focus
+ | _ ->
+ raise FullyUnfocused
+
+(* Auxiliary function to push a [proof_state] onto the undo stack. *)
+let push_undo save ({ undo_stack = stack } as pr) =
+ pr.undo_stack <- save::stack
+
+(* Auxiliary function to pop and read a [save_state] from the undo stack. *)
+exception EmptyUndoStack
+let pop_undo pr =
+ match pr.undo_stack with
+ | state::stack -> pr.undo_stack <- stack; state
+ | _ -> raise EmptyUndoStack
+
+
+(* This function focuses the proof [pr] between indices [i] and [j] *)
+let _focus cond i j pr =
+ let (focused,context) = Proofview.focus i j pr.state.proofview in
+ push_focus cond context pr;
+ pr.state <- { pr.state with proofview = focused }
+
+(* This function unfocuses the proof [pr], it raises [CannotUnfocus],
+ if the proof is already fully unfocused.
+ This function does not care about the condition of the current focus. *)
+let _unfocus pr =
+ let (_,fc) = pop_focus pr in
+ pr.state <- { pr.state with proofview = Proofview.unfocus fc pr.state.proofview }
+
+
+(*** Endline tactic ***)
+
+(* spiwack this is an information about the UI, it might be a good idea to move
+ it to the Proof_global module *)
+
+(* Sets the tactic to be used when a tactic line is closed with [...] *)
+let set_endline_tactic tac p =
+ p.info.endline_tactic <- tac
+
+let with_end_tac pr tac =
+ Proofview.tclTHEN tac pr.info.endline_tactic
+
+(*** The following functions define the safety mechanism of the
+ proof system, they may be unsafe if not used carefully. There is
+ currently no reason to export them in the .mli ***)
+
+(* This functions saves the current state into a [proof_state]. *)
+let save_state { state = ps } = State ps
+
+(* This function stores the current proof state in the undo stack. *)
+let save pr =
+ push_undo (save_state pr) pr
+
+(* This function restores a state, presumably from the top of the undo stack. *)
+let restore_state save pr =
+ match save with
+ | State save -> pr.state <- save
+ | Effect undo -> undo ()
+
+(* Interpretes the Undo command. *)
+let undo pr =
+ (* On a single line, since the effects commute *)
+ restore_state (pop_undo pr) pr
+
+(* Adds an undo effect to the undo stack. Use it with care, errors here might result
+ in inconsistent states. *)
+let add_undo effect pr =
+ push_undo (Effect effect) pr
+
+(* Focus command (focuses on the [i]th subgoal) *)
+(* spiwack: there could also, easily be a focus-on-a-range tactic, is there
+ a need for it? *)
+let focus cond i pr =
+ save pr;
+ _focus cond i i pr
+
+(* Unfocus command.
+ Fails if the proof is not focused. *)
+let unfocus kind pr =
+ let starting_point = save_state pr in
+ try
+ let cond = cond_of_focus pr in
+ if cond kind pr.state.proofview then
+ (_unfocus pr;
+ push_undo starting_point pr)
+ else
+ Util.error "This proof is focused, but cannot be unfocused this way"
+ with FullyUnfocused ->
+ Util.error "This proof is not focused"
+
+let no_focused_goal p =
+ Proofview.finished p.state.proofview
+
+(*** Function manipulation proof extra informations ***)
+
+let get_proof_info pr =
+ pr.state.intel
+
+let set_proof_info info pr =
+ save pr;
+ pr.state <- { pr.state with intel=info }
+
+
+(*** Tactics ***)
+
+let run_tactic env tac pr =
+ let starting_point = save_state pr in
+ let sp = pr.state.proofview in
+ try
+ let tacticced_proofview = Proofview.apply env tac sp in
+ pr.state <- { pr.state with proofview = tacticced_proofview };
+ push_undo starting_point pr
+ with e ->
+ restore_state starting_point pr;
+ raise e
+
+
+
+(*** Compatibility layer with <=v8.2 ***)
+module V82 = struct
+ let subgoals p =
+ Proofview.V82.goals p.state.proofview
+
+ let background_subgoals p =
+ Proofview.V82.goals (unroll_focus p.state.proofview p.state.focus_stack)
+
+ let get_initial_conclusions p =
+ p.info.initial_conclusions
+
+ let depth p = List.length p.undo_stack
+
+ let top_goal p =
+ let { Evd.it=gls ; sigma=sigma } =
+ Proofview.V82.top_goals p.state.proofview
+ in
+ { Evd.it=List.hd gls ; sigma=sigma }
+
+ let instantiate_evar n com pr =
+ let starting_point = save_state pr in
+ let sp = pr.state.proofview in
+ try
+ let new_proofview = Proofview.V82.instantiate_evar n com sp in
+ pr.state <- { pr.state with proofview = new_proofview };
+ push_undo starting_point pr
+ with e ->
+ restore_state starting_point pr;
+ raise e
+end
diff --git a/proofs/proof.mli b/proofs/proof.mli
new file mode 100644
index 0000000000..2b1c3f5c20
--- /dev/null
+++ b/proofs/proof.mli
@@ -0,0 +1,133 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* $Id: proof.mli aspiwack $ *)
+
+(* Module defining the last essential tiles of interactive proofs.
+ The features of the Proof module are undoing and focusing.
+ A proof is a mutable object, it contains a proofview, and some information
+ to be able to undo actions, and to unfocus the current view. All three
+ of these being meant to evolve.
+ - Proofview: a proof is primarily the data of the current view.
+ That which is shown to the user (as a remainder, a proofview
+ is mainly the logical state of the proof, together with the
+ currently focused goals).
+ - Focus: a proof has a focus stack: the top of the stack contains
+ the context in which to unfocus the current view to a view focused
+ with the rest of the stack.
+ In addition, this contains, for each of the focus context, a
+ "focus kind" and a "focus condition" (in practice, and for modularity,
+ the focus kind is actually stored inside the condition). To unfocus, one
+ needs to know the focus kind, and the condition (for instance "no condition" or
+ the proof under focused must be complete) must be met.
+ - Undo: since proofviews and focus stacks are immutable objects,
+ it could suffice to hold the previous states, to allow to return to the past.
+ However, we also allow other modules to do actions that can be undone.
+ Therefore the undo stack stores action to be ran to undo.
+*)
+
+open Term
+
+(* Type of a proof. *)
+type proof
+
+
+(*** General proof functions ***)
+
+val start : (Environ.env * Term.types) list -> proof
+
+(* Returns [true] is the considered proof is completed, that is if no goal remain
+ to be considered (this does not require that all evars have been solved). *)
+val is_done : proof -> bool
+
+(* Returns the list of partial proofs to initial goals. *)
+val partial_proof : proof -> Term.constr list
+
+(* Returns the proofs (with their type) of the initial goals.
+ Raises [UnfinishedProof] is some goals remain to be considered.
+ Raises [HasUnresolvedEvar] if some evars have been left undefined. *)
+exception UnfinishedProof
+exception HasUnresolvedEvar
+val return : proof -> (Term.constr * Term.types) list
+
+(* Interpretes the Undo command. Raises [EmptyUndoStack] if
+ the undo stack is empty. *)
+exception EmptyUndoStack
+val undo : proof -> unit
+
+(* Adds an undo effect to the undo stack. Use it with care, errors here might result
+ in inconsistent states. *)
+val add_undo : (unit -> unit) -> proof -> unit
+
+(*** Focusing actions ***)
+
+(* [focus_kind] is the type used by focusing and unfocusing
+ commands to synchronise. Focusing and unfocusing commands use
+ a particular focus_kind, and if they don't match, the unfocusing command
+ will fail. *)
+type focus_kind
+val new_focus_kind : unit -> focus_kind
+
+(* To be authorized to unfocus one must meet the condition prescribed by
+ the action which focused.*)
+type focus_condition
+(* [no_cond] only checks that the unfocusing command uses the right
+ [focus_kind]. *)
+val no_cond : focus_kind -> focus_condition
+(* [done_cond] checks that the unfocusing command uses the right [focus_kind]
+ and that the focused proofview is complete. *)
+val done_cond : focus_kind -> focus_condition
+
+(* focus command (focuses on the [i]th subgoal) *)
+(* there could also, easily be a focus-on-a-range tactic, is there
+ a need for it? *)
+val focus : focus_condition -> int -> proof -> unit
+
+exception FullyUnfocused
+(* Unfocusing command.
+ Raises [FullyUnfocused] if the proof is not focused. *)
+val unfocus : focus_kind -> proof -> unit
+
+(* returns [true] if there is no goal under focus. *)
+val no_focused_goal : proof -> bool
+
+(*** Function manipulation proof extra informations ***)
+
+val get_proof_info : proof -> Store.t
+
+val set_proof_info : Store.t -> proof -> unit
+
+
+(*** Endline tactic ***)
+
+(* Sets the tactic to be used when a tactic line is closed with [...] *)
+val set_endline_tactic : unit Proofview.tactic -> proof -> unit
+
+val with_end_tac : proof -> unit Proofview.tactic -> unit Proofview.tactic
+
+(*** Tactics ***)
+
+val run_tactic : Environ.env -> unit Proofview.tactic -> proof -> unit
+
+
+(*** Compatibility layer with <=v8.2 ***)
+module V82 : sig
+ val subgoals : proof -> Goal.goal list Evd.sigma
+
+ (* All the subgoals of the proof, including those which are not focused. *)
+ val background_subgoals : proof -> Goal.goal list Evd.sigma
+
+ val get_initial_conclusions : proof -> Term.types list
+
+ val depth : proof -> int
+
+ val top_goal : proof -> Goal.goal Evd.sigma
+
+ (* Implements the Existential command *)
+ val instantiate_evar : int -> Topconstr.constr_expr -> proof -> unit
+end
diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml
new file mode 100644
index 0000000000..420ff84325
--- /dev/null
+++ b/proofs/proof_global.ml
@@ -0,0 +1,295 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* $Id$ *)
+
+(***********************************************************************)
+(* *)
+(* This module defines the global proof environment *)
+(* In particular it keeps tracks of whether or not there is *)
+(* a proof which is currently being edited. *)
+(* *)
+(***********************************************************************)
+
+open Pp
+open Names
+
+(*** Proof Modes ***)
+
+(* Type of proof modes :
+ - A function [set] to set it *from standard mode*
+ - A function [reset] to reset the *standard mode* from it *)
+type proof_mode = {
+ name : string ;
+ set : unit -> unit ;
+ reset : unit -> unit
+}
+
+let proof_modes = Hashtbl.create 6
+let register_proof_mode ({ name = n } as m) = Hashtbl.add proof_modes n m
+(* initial mode: standard mode *)
+let standard = { name = "No" ; set = (fun ()->()) ; reset = (fun () -> ()) }
+let _ = register_proof_mode standard
+
+(* Default proof mode, to be set at the beginning of proofs. *)
+let default_proof_mode = ref standard
+
+let set_default_proof_mode =
+ Goptions.declare_string_option {Goptions.
+ optsync = true ;
+ optname = "default proof mode" ;
+ optkey = ["Default";"Proof";"Mode"] ;
+ optread = begin fun () ->
+ let { name = name } = !default_proof_mode in name
+ end;
+ optwrite = begin fun n ->
+ default_proof_mode := Hashtbl.find proof_modes n
+ end
+ }
+
+(*** Proof Global Environment ***)
+
+(* local shorthand *)
+type nproof = identifier*Proof.proof
+
+(* Extra info on proofs. *)
+type lemma_possible_guards = int list list
+type proof_info = {
+ strength : Decl_kinds.goal_kind ;
+ compute_guard : lemma_possible_guards;
+ hook :Tacexpr.declaration_hook ;
+ mode : proof_mode
+}
+
+(* Invariant: a proof is at most in one of current_proof and suspended. And the
+ domain of proof_info is the union of that of current_proof and suspended.*)
+(* The head of [!current_proof] is the actual current proof, the other ones are to
+ be resumed when the current proof is closed, aborted or suspended. *)
+let current_proof = ref ([]:nproof list)
+let suspended = ref ([] : nproof list)
+let proof_info = ref (Idmap.empty : proof_info Idmap.t)
+
+(* Current proof_mode, for bookkeeping *)
+let current_proof_mode = ref !default_proof_mode
+
+(* combinators for proof modes *)
+let update_proof_mode () =
+ match !current_proof with
+ | (id,_)::_ ->
+ let { mode = m } = Idmap.find id !proof_info in
+ !current_proof_mode.reset ();
+ current_proof_mode := m;
+ !current_proof_mode.set ()
+ | _ ->
+ !current_proof_mode.reset ();
+ current_proof_mode := standard
+
+(* combinators for the current_proof and suspended lists *)
+let push a l = l := a::!l;
+ update_proof_mode ()
+
+exception NoSuchProof
+let rec extract id l =
+ let rec aux = function
+ | ((id',_) as np)::l when id_ord id id' = 0 -> (np,l)
+ | np::l -> let (np', l) = aux l in (np' , np::l)
+ | [] -> raise NoSuchProof
+ in
+ let (np,l') = aux !l in
+ l := l';
+ update_proof_mode ();
+ np
+
+exception NoCurrentProof
+let extract_top l =
+ match !l with
+ | np::l' -> l := l' ; update_proof_mode (); np
+ | [] -> raise NoCurrentProof
+let find_top l =
+ match !l with
+ | np::_ -> np
+ | [] -> raise NoCurrentProof
+
+let rotate_top l1 l2 =
+ let np = extract_top l1 in
+ push np l2
+
+let rotate_find id l1 l2 =
+ let np = extract id l1 in
+ push np l2
+
+
+(* combinators for the proof_info map *)
+let add id info m =
+ m := Idmap.add id info !m
+let remove id m =
+ m := Idmap.remove id !m
+
+(*** Proof Global manipulation ***)
+
+let get_all_proof_names () =
+ List.map fst !current_proof @
+ List.map fst !suspended
+
+let give_me_the_proof () =
+ snd (find_top current_proof)
+
+let get_current_proof_name () =
+ fst (find_top current_proof)
+
+(* 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 use_resume =
+ match get_all_proof_names () with
+ | [] -> (spc () ++ str"(No proof-editing in progress).")
+ | l -> (str"." ++ fnl () ++ str"Proofs currently edited:" ++ spc () ++
+ (Util.prlist_with_sep Util.pr_spc Nameops.pr_id l) ++
+ str"." ++
+ (if use_resume then (fnl () ++ str"Use \"Resume\" first.")
+ else (mt ()))
+ )
+
+
+let there_is_a_proof () = !current_proof <> []
+let there_are_suspended_proofs () = !suspended <> []
+let there_are_pending_proofs () =
+ there_is_a_proof () ||
+ there_are_suspended_proofs ()
+let check_no_pending_proof () =
+ if not (there_are_pending_proofs ()) then
+ ()
+ else begin
+ pp_with Format.str_formatter
+ (str"Proof editing in progress" ++ (msg_proofs false) ++ fnl() ++
+ str"Use \"Abort All\" first or complete proof(s).") ;
+ Util.error (Format.flush_str_formatter ())
+ end
+
+
+let suspend () =
+ rotate_top current_proof suspended
+
+let resume_last () =
+ rotate_top suspended current_proof
+
+let resume id =
+ rotate_find id suspended current_proof
+
+let discard_gen id =
+ try
+ ignore (extract id current_proof);
+ remove id proof_info
+ with NoSuchProof -> ignore (extract id suspended)
+
+let discard (loc,id) =
+ try
+ discard_gen id
+ with NoSuchProof ->
+ Util.user_err_loc
+ (loc,"Pfedit.delete_proof",str"No such proof" ++ msg_proofs false)
+
+let discard_current () =
+ let (id,_) = extract_top current_proof in
+ remove id proof_info
+
+let discard_all () =
+ current_proof := [];
+ suspended := [];
+ proof_info := Idmap.empty
+
+(* [set_proof_mode] sets the proof mode to be used after it's called. It is
+ typically called by the Proof Mode command. *)
+(* Core component.
+ No undo handling.
+ Applies to proof [id], and proof mode [m]. *)
+let set_proof_mode m id =
+ let info = Idmap.find id !proof_info in
+ let info = { info with mode = m } in
+ proof_info := Idmap.add id info !proof_info;
+ update_proof_mode ()
+(* Complete function.
+ Handles undo.
+ Applies to current proof, and proof mode name [mn]. *)
+let set_proof_mode mn =
+ let m = Hashtbl.find proof_modes mn in
+ let id = get_current_proof_name () in
+ let pr = give_me_the_proof () in
+ Proof.add_undo begin let curr = !current_proof_mode in fun () ->
+ set_proof_mode curr id ; update_proof_mode ()
+ end pr ;
+ set_proof_mode m id
+
+(* [start_proof s str env t hook tac] starts a proof of name [s] and
+ conclusion [t]; [hook] is optionally a function to be applied at
+ proof end (e.g. to declare the built constructions as a coercion
+ or a setoid morphism); init_tac is possibly a tactic to
+ systematically apply at initialization time (e.g. to start the
+ proof of mutually dependent theorems).
+ It raises exception [ProofInProgress] if there is a proof being
+ currently edited. *)
+let start_proof id str goals ?(compute_guard=[]) hook =
+ (* arnaud: ajouter une vérification pour la présence de id dans le proof_global *)
+ let p = Proof.start goals in
+ add id { strength=str ;
+ compute_guard=compute_guard ;
+ hook=hook ;
+ mode = ! default_proof_mode } proof_info ;
+ push (id,p) current_proof
+
+(* arnaud: à enlever *)
+let run_tactic tac =
+ let p = give_me_the_proof () in
+ let env = Global.env () in
+ Proof.run_tactic env tac p
+
+(* Sets the tactic to be used when a tactic line is closed with [...] *)
+let set_endline_tactic tac =
+ let p = give_me_the_proof () in
+ Proof.set_endline_tactic tac p
+
+let with_end_tac tac =
+ let p = give_me_the_proof () in
+ Proof.with_end_tac p tac
+
+let close_proof () =
+ (* spiwack: for now close_proof doesn't actually discard the proof, it is done
+ by [Command.save]. *)
+ try
+ let id = get_current_proof_name () in
+ let p = give_me_the_proof () in
+ let proofs_and_types = Proof.return p in
+ let entries = List.map (fun (c,t) -> { Entries.const_entry_body = c ;
+ const_entry_type = Some t;
+ const_entry_opaque = true ;
+ const_entry_boxed = false } )
+ proofs_and_types
+ in
+ let { compute_guard=cg ; strength=str ; hook=hook } =
+ Idmap.find id !proof_info
+ in
+ (id, (entries,cg,str,hook))
+ with
+ | Proof.UnfinishedProof ->
+ Util.error "Attempt to save an incomplete proof"
+ | Proof.HasUnresolvedEvar ->
+ Util.error "Attempt to save a proof with existential variables still non-instantiated"
+
+module V82 = struct
+ let get_current_initial_conclusions () =
+ let p = give_me_the_proof () in
+ let id = get_current_proof_name () in
+ let { strength=str ; hook=hook } =
+ Idmap.find id !proof_info
+ in
+ (id,(Proof.V82.get_initial_conclusions p, str, hook))
+end
+
diff --git a/proofs/proof_global.mli b/proofs/proof_global.mli
new file mode 100644
index 0000000000..84a61c755a
--- /dev/null
+++ b/proofs/proof_global.mli
@@ -0,0 +1,89 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* $Id$ *)
+
+(***********************************************************************)
+(* *)
+(* This module defines the global proof environment *)
+(* Especially it keeps tracks of whether or not there is *)
+(* a proof which is currently being edited. *)
+(* *)
+(***********************************************************************)
+
+(* Type of proof modes :
+ - A name
+ - A function [set] to set it *from standard mode*
+ - A function [reset] to reset the *standard mode* from it *)
+type proof_mode = {
+ name : string ;
+ set : unit -> unit ;
+ reset : unit -> unit
+}
+(* Registers a new proof mode which can then be adressed by name
+ in [set_default_proof_mode].
+ One mode is already registered - the standard mode - named "No",
+ It corresponds to Coq default setting are they are set when coqtop starts. *)
+val register_proof_mode : proof_mode -> unit
+
+val there_is_a_proof : unit -> bool
+val there_are_pending_proofs : unit -> bool
+val check_no_pending_proof : unit -> unit
+
+val get_current_proof_name : unit -> Names.identifier
+val get_all_proof_names : unit -> Names.identifier list
+
+val discard : Names.identifier Util.located -> unit
+val discard_current : unit -> unit
+val discard_all : unit -> unit
+
+(* [set_proof_mode] sets the proof mode to be used after it's called. It is
+ typically called by the Proof Mode command. *)
+val set_proof_mode : string -> unit
+
+exception NoCurrentProof
+val give_me_the_proof : unit -> Proof.proof
+
+
+(* [start_proof s str goals ~init_tac ~compute_guard hook] starts
+ a proof of name [s] and
+ conclusion [t]; [hook] is optionally a function to be applied at
+ proof end (e.g. to declare the built constructions as a coercion
+ or a setoid morphism). *)
+type lemma_possible_guards = int list list
+val start_proof : Names.identifier ->
+ Decl_kinds.goal_kind ->
+ (Environ.env * Term.types) list ->
+ ?compute_guard:lemma_possible_guards ->
+ Tacexpr.declaration_hook ->
+ unit
+
+val close_proof : unit ->
+ Names.identifier *
+ (Entries.definition_entry list *
+ lemma_possible_guards *
+ Decl_kinds.goal_kind *
+ Tacexpr.declaration_hook)
+
+val suspend : unit -> unit
+val resume_last : unit -> unit
+val resume : Names.identifier -> unit
+
+(* Runs a tactic on the current proof. Raises [NoCurrentProof] is there is
+ no current proof. *)
+val run_tactic : unit Proofview.tactic -> unit
+
+(* Sets the tactic to be used when a tactic line is closed with [...] *)
+val set_endline_tactic : unit Proofview.tactic -> unit
+
+(* Appends the endline tactic of the current proof to a tactic. *)
+val with_end_tac : unit Proofview.tactic -> unit Proofview.tactic
+
+module V82 : sig
+ val get_current_initial_conclusions : unit -> Names.identifier *(Term.types list * Decl_kinds.goal_kind * Tacexpr.declaration_hook)
+end
diff --git a/proofs/proof_trees.ml b/proofs/proof_trees.ml
deleted file mode 100644
index a5bd073a30..0000000000
--- a/proofs/proof_trees.ml
+++ /dev/null
@@ -1,107 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(* $Id$ *)
-
-open Closure
-open Util
-open Names
-open Nameops
-open Term
-open Termops
-open Sign
-open Evd
-open Environ
-open Evarutil
-open Decl_expr
-open Proof_type
-open Tacred
-open Typing
-open Libnames
-open Nametab
-
-(*
-let is_bind = function
- | Tacexpr.Cbindings _ -> true
- | _ -> false
-*)
-
-(* Functions on goals *)
-
-let mk_goal hyps cl extra =
- { evar_hyps = hyps; evar_concl = cl;
- evar_filter = List.map (fun _ -> true) (named_context_of_val hyps);
- evar_body = Evar_empty; evar_source = (dummy_loc,GoalEvar);
- evar_extra = extra }
-
-(* Functions on proof trees *)
-
-let ref_of_proof pf =
- match pf.ref with
- | None -> failwith "rule_of_proof"
- | Some r -> r
-
-let rule_of_proof pf =
- let (r,_) = ref_of_proof pf in r
-
-let children_of_proof pf =
- let (_,cl) = ref_of_proof pf in cl
-
-let goal_of_proof pf = pf.goal
-
-let subproof_of_proof pf = match pf.ref with
- | Some (Nested (_,pf), _) -> pf
- | _ -> failwith "subproof_of_proof"
-
-let status_of_proof pf = pf.open_subgoals
-
-let is_complete_proof pf = pf.open_subgoals = 0
-
-let is_leaf_proof pf = (pf.ref = None)
-
-let is_tactic_proof pf = match pf.ref with
- | Some (Nested (Tactic _,_),_) -> true
- | _ -> false
-
-
-let pf_lookup_name_as_displayed env ccl s =
- Detyping.lookup_name_as_displayed env ccl s
-
-let pf_lookup_index_as_renamed env ccl n =
- Detyping.lookup_index_as_renamed env ccl n
-
-(* Functions on rules (Proof mode) *)
-
-let is_dem_rule = function
- Decl_proof _ -> true
- | _ -> false
-
-let is_proof_instr = function
- Nested(Proof_instr (_,_),_) -> true
- | _ -> false
-
-let is_focussing_command = function
- Decl_proof b -> b
- | Nested (Proof_instr (b,_),_) -> b
- | _ -> false
-
-
-(*********************************************************************)
-(* Pretty printing functions *)
-(*********************************************************************)
-
-open Pp
-
-let db_pr_goal g =
- let env = evar_env g in
- let penv = print_named_context env in
- let pc = print_constr_env env g.evar_concl in
- str" " ++ hv 0 (penv ++ fnl () ++
- str "============================" ++ fnl () ++
- str" " ++ pc) ++ fnl ()
-
diff --git a/proofs/proof_trees.mli b/proofs/proof_trees.mli
deleted file mode 100644
index 6d1fc143d3..0000000000
--- a/proofs/proof_trees.mli
+++ /dev/null
@@ -1,48 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(*i $Id$ i*)
-
-(*i*)
-open Util
-open Names
-open Term
-open Sign
-open Evd
-open Environ
-open Proof_type
-(*i*)
-
-(* This module declares readable constraints, and a few utilities on
- constraints and proof trees *)
-
-val mk_goal : named_context_val -> constr -> Dyn.t option -> goal
-
-val rule_of_proof : proof_tree -> rule
-val ref_of_proof : proof_tree -> (rule * proof_tree list)
-val children_of_proof : proof_tree -> proof_tree list
-val goal_of_proof : proof_tree -> goal
-val subproof_of_proof : proof_tree -> proof_tree
-val status_of_proof : proof_tree -> int
-val is_complete_proof : proof_tree -> bool
-val is_leaf_proof : proof_tree -> bool
-val is_tactic_proof : proof_tree -> bool
-
-val pf_lookup_name_as_displayed : env -> constr -> identifier -> int option
-val pf_lookup_index_as_renamed : env -> constr -> int -> int option
-
-val is_proof_instr : rule -> bool
-val is_focussing_command : rule -> bool
-
-(*s Pretty printing functions. *)
-
-(*i*)
-open Pp
-(*i*)
-
-val db_pr_goal : goal -> std_ppcmds
diff --git a/proofs/proof_type.ml b/proofs/proof_type.ml
index bc29534081..2fffa39527 100644
--- a/proofs/proof_type.ml
+++ b/proofs/proof_type.ml
@@ -16,7 +16,7 @@ open Libnames
open Term
open Util
open Tacexpr
-open Decl_expr
+(* open Decl_expr *)
open Rawterm
open Genarg
open Nametab
@@ -26,6 +26,10 @@ open Pattern
(* This module defines the structure of proof tree and the tactic type. So, it
is used by Proof_tree and Refiner *)
+type goal = Goal.goal
+
+type tactic = goal sigma -> goal list sigma
+
type prim_rule =
| Intro of identifier
| Cut of bool * bool * identifier * types
@@ -54,13 +58,6 @@ and rule =
and compound_rule=
| Tactic of tactic_expr * bool
- | Proof_instr of bool*proof_instr (* the boolean is for focus restrictions *)
-
-and goal = evar_info
-
-and tactic = goal sigma -> (goal list sigma * validation)
-
-and validation = (proof_tree list -> proof_tree)
and tactic_expr =
(constr,
diff --git a/proofs/proof_type.mli b/proofs/proof_type.mli
index b5c4a234d1..9692f19bc9 100644
--- a/proofs/proof_type.mli
+++ b/proofs/proof_type.mli
@@ -16,7 +16,6 @@ open Libnames
open Term
open Util
open Tacexpr
-open Decl_expr
open Rawterm
open Genarg
open Nametab
@@ -26,6 +25,10 @@ open Pattern
(* This module defines the structure of proof tree and the tactic type. So, it
is used by [Proof_tree] and [Refiner] *)
+type goal = Goal.goal
+
+type tactic = goal sigma -> goal list sigma
+
type prim_rule =
| Intro of identifier
| Cut of bool * bool * identifier * types
@@ -89,13 +92,6 @@ and rule =
and compound_rule=
(* the boolean of Tactic tells if the default tactic is used *)
| Tactic of tactic_expr * bool
- | Proof_instr of bool * proof_instr
-
-and goal = evar_info
-
-and tactic = goal sigma -> (goal list sigma * validation)
-
-and validation = (proof_tree list -> proof_tree)
and tactic_expr =
(constr,
diff --git a/proofs/proofs.mllib b/proofs/proofs.mllib
index 05b86b1a04..66001e77a5 100644
--- a/proofs/proofs.mllib
+++ b/proofs/proofs.mllib
@@ -1,12 +1,15 @@
+Goal
+Evar_refiner
+Proofview
+Proof
+Proof_global
Tacexpr
Proof_type
Redexpr
-Proof_trees
Logic
Refiner
-Evar_refiner
Tacmach
Pfedit
Tactic_debug
+Clenv
Clenvtac
-Decl_mode
diff --git a/proofs/proofview.ml b/proofs/proofview.ml
new file mode 100644
index 0000000000..a08941df03
--- /dev/null
+++ b/proofs/proofview.ml
@@ -0,0 +1,491 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* $Id$ *)
+
+(* The proofview datastructure is a pure datastructure underlying the notion
+ of proof (namely, a proof is a proofview which can evolve and has safety
+ mechanisms attached).
+ The general idea of the structure is that it is composed of a chemical
+ solution: an unstructured bag of stuff which has some relations with
+ one another, which represents the various subnodes of the proof, together
+ with a comb: a datastructure that gives order to some of these nodes,
+ namely the open goals.
+ The natural candidate for the solution is an {!Evd.evar_map}, that is
+ a calculus of evars. The comb is then a list of goals (evars wrapped
+ with some extra information, like possible name anotations).
+ There is also need of a list of the evars which initialised the proofview
+ to be able to return information about the proofview. *)
+
+(* Type of proofviews. *)
+type proofview = {
+ initial : (Term.constr * Term.types) list;
+ solution : Evd.evar_map;
+ comb : Goal.goal list
+ }
+
+(* Initialises a proofview, the argument is a list of environement,
+ conclusion types, and optional names, creating that many initial goals. *)
+let init =
+ let rec aux = function
+ | [] -> { initial = [] ;
+ solution = Evd.empty ;
+ comb = []
+ }
+ | (env,typ)::l -> let { initial = ret ; solution = sol ; comb = comb } =
+ aux l
+ in
+ let ( new_defs , econstr ) =
+ Evarutil.new_evar sol env typ
+ in
+ let (e,_) = Term.destEvar econstr in
+ let gl = Goal.build e in
+ { initial = (econstr,typ)::ret;
+ solution = new_defs ;
+ comb = gl::comb }
+ in
+ fun l -> let v = aux l in
+ (* Marks all the goal unresolvable for typeclasses. *)
+ { v with solution = Typeclasses.mark_unresolvables v.solution }
+
+(* Returns whether this proofview is finished or not. That is,
+ if it has empty subgoals in the comb. There could still be unsolved
+ subgoaled, but they would then be out of the view, focused out. *)
+let finished = function
+ | {comb = []} -> true
+ | _ -> false
+
+(* Returns the current value of the proofview partial proofs. *)
+let return { initial=init; solution=defs } =
+ List.map (fun (c,t) -> (Evarutil.nf_evar defs c , t)) init
+
+(* spiwack: this function should probably go in the Util section,
+ but I'd rather have Util (or a separate module for lists)
+ raise proper exceptions before *)
+(* [IndexOutOfRange] occurs in case of malformed indices
+ with respect to list lengths. *)
+exception IndexOutOfRange
+
+(* [list_goto i l] returns a pair of lists [c,t] where
+ [c] has length [i] and is the reversed of the [i] first
+ elements of [l], and [t] is the rest of the list.
+ The idea is to navigate through the list, [c] is then
+ seen as the context of the current position.
+ Raises [IndexOutOfRange] if [i > length l]*)
+let list_goto =
+ let rec aux acc index = function
+ | l when index = 0-> (acc,l)
+ | [] -> raise IndexOutOfRange
+ | a::q -> aux (a::acc) (index-1) q
+ in
+ fun i l ->
+ if i < 0 then
+ raise IndexOutOfRange
+ else
+ aux [] i l
+
+(* Type of the object which allow to unfocus a view.*)
+(* First component is a reverse list of what comes before
+ and second component is what goes after (in the expected
+ order) *)
+type focus_context = Goal.goal list * Goal.goal list
+
+(* This (internal) function extracts a sublist between two indices, and
+ returns this sublist together with its context:
+ if it returns [(a,(b,c))] then [a] is the sublist and (rev b)@a@c is the
+ original list.
+ The focused list has lenght [j-i-1] and contains the goals from
+ number [i] to number [j] (both included) the first goal of the list
+ being numbered [1].
+ [focus_sublist i j l] raises [IndexOutOfRange] if
+ [i > length l], or [j > length l] or [ j < i ]. *)
+let focus_sublist i j l =
+ let (left,sub_right) = list_goto (i-1) l in
+ let (sub, right) =
+ try
+ Util.list_chop (j-i+1) sub_right
+ with Failure "list_chop" ->
+ Util.errorlabstrm "nth_unproven" (Pp.str"Not such unproven subgoal")
+ in
+ (sub, (left,right))
+
+(* Inverse operation to the previous one. *)
+let unfocus_sublist (left,right) s =
+ List.rev_append left (s@right)
+
+
+(* [focus i j] focuses a proofview on the goals from index [i] to index [j]
+ (inclusive). (i.e. goals number [i] to [j] become the only goals of the
+ returned proofview).
+ It returns the focus proof, and a context for the focus trace. *)
+let focus i j sp =
+ let (new_comb, context) = focus_sublist i j sp.comb in
+ ( { sp with comb = new_comb } , context )
+
+(* Unfocuses a proofview with respect to a context. *)
+let undefined defs l =
+ Option.List.flatten (List.map (Goal.advance defs) l)
+let unfocus c sp =
+ { sp with comb = undefined sp.solution (unfocus_sublist c sp.comb) }
+
+
+(* The tactic monad:
+ - Tactics are objects which apply a transformation to all
+ the subgoals of the current view at the same time. By opposed
+ to the old vision of applying it to a single goal. It mostly
+ allows to consider tactic like [reorder] to reorder the goals
+ in the current view (which might be useful for the tactic designer)
+ (* spiwack: the ordering of goals, though, is perhaps a bit
+ brittle. It would be much more interesting to find a more
+ robust way to adress goals, I have no idea at this time
+ though*)
+ or global automation tactic for dependent subgoals (instantiating
+ an evar has influences on the other goals of the proof in progress,
+ not being able to take that into account causes the current eauto
+ tactic to fail on some instances where it could succeed).
+ - Tactics are a monad ['a tactic], in a sense a tactic can be
+ seens as a function (without argument) which returns a value
+ of type 'a and modifies the environement (in our case: the view).
+ Tactics of course have arguments, but these are given at the
+ meta-level as OCaml functions.
+ Most tactics, in the sense we are used to, return [ () ], that is
+ no really interesting values. But some might pass information
+ around; the [(>>--)] and [(>>==)] bind-like construction are the
+ main ingredients of this information passing.
+ (* spiwack: I don't know how much all this relates to F. Kirchner and
+ C. Muñoz. I wasn't able to understand how they used the monad
+ structure in there developpement.
+ *)
+ The tactics seen in Coq's Ltac are (for now at least) only
+ [unit tactic], the return values are kept for the OCaml toolkit.
+ The operation or the monad are [Proofview.tclUNIT] (which is the
+ "return" of the tactic monad) [Proofview.tclBIND] (which is
+ the "bind", also noted [(>=)]) and [Proofview.tclTHEN] (which is a
+ specialized bind on unit-returning tactics).
+*)
+
+(* type of tactics *)
+(* spiwack: double-continuation backtracking monads are reasonable
+ folklore for "search" implementations (including Tac interactive prover's
+ tactics). Yet it's quite hard to wrap your head around these.
+ I recommand reading a few times the "Backtracking, Interleaving, and Terminating
+ Monad Transformers" paper by O. Kiselyov, C. Chen, D. Fridman.
+ The peculiar shape of the monadic type is reminiscent of that of the continuation
+ monad transformer.
+ A good way to get a feel of what's happening is to look at what happens when
+ executing [apply (tclUNIT ())].
+ The disjunction function is unlike that of the LogicT paper, because we want and
+ need to backtrack over state as well as values. Therefore we cannot be
+ polymorphic over the inner monad. *)
+type proof_step = { goals : Goal.goal list ; defs : Evd.evar_map }
+type +'a result = { proof_step : proof_step ;
+ content : 'a }
+
+(* nb=non-backtracking *)
+type +'a nb_tactic = proof_step -> 'a result
+
+(* double-continutation backtracking *)
+(* "sk" stands for "success continuation", "fk" for "failure continuation" *)
+type 'r fk = exn -> 'r
+type (-'a,'r) sk = 'a -> 'r fk -> 'r
+type +'a tactic0 = { go : 'r. ('a, 'r nb_tactic) sk -> 'r nb_tactic fk -> 'r nb_tactic }
+
+(* We obtain a tactic by parametrizing with an environment *)
+(* spiwack: alternatively the environment could be part of the "nb_tactic" state
+ monad. As long as we do not intend to change the environment during a tactic,
+ it's probably better here. *)
+type +'a tactic = Environ.env -> 'a tactic0
+
+(* unit of [nb_tactic] *)
+let nb_tac_unit a step = { proof_step = step ; content = a }
+
+(* Applies a tactic to the current proofview. *)
+let apply env t sp =
+ let start = { goals = sp.comb ; defs = sp.solution } in
+ let res = (t env).go (fun a _ step -> nb_tac_unit a step) (fun e _ -> raise e) start in
+ let next = res.proof_step in
+ {sp with
+ solution = next.defs ;
+ comb = next.goals
+ }
+
+(*** tacticals ***)
+
+
+(* Unit of the tactic monad *)
+let tclUNIT a _ = { go = fun sk fk step -> sk a fk step }
+
+(* Bind operation of the tactic monad *)
+let tclBIND t k env = { go = fun sk fk step ->
+ (t env).go (fun a fk -> (k a env).go sk fk) fk step
+}
+
+(* Interpretes the ";" (semicolon) of Ltac.
+ As a monadic operation, it's a specialized "bind"
+ on unit-returning tactic (meaning "there is no value to bind") *)
+let tclTHEN t1 t2 env = { go = fun sk fk step ->
+ (t1 env).go (fun () fk -> (t2 env).go sk fk) fk step
+}
+
+(* [tclIGNORE t] has the same operational content as [t],
+ but drops the value at the end. *)
+let tclIGNORE tac env = { go = fun sk fk step ->
+ (tac env).go (fun _ fk -> sk () fk) fk step
+}
+
+(* [tclOR t1 t2 = t1] if t1 succeeds and [tclOR t1 t2 = t2] if t1 fails.
+ No interleaving for the moment. *)
+(* spiwack: compared to the LogicT paper, we backtrack at the same state
+ where [t1] has been called, not the state where [t1] failed. *)
+let tclOR t1 t2 env = { go = fun sk fk step ->
+ (t1 env).go sk (fun _ _ -> (t2 env).go sk fk step) step
+}
+
+(* [tclZERO e] always fails with error message [e]*)
+let tclZERO e env = { go = fun _ fk step -> fk e step }
+
+
+(* Focusing operation on proof_steps. *)
+let focus_proof_step i j ps =
+ let (new_subgoals, context) = focus_sublist i j ps.goals in
+ ( { ps with goals = new_subgoals } , context )
+
+(* Unfocusing operation of proof_steps. *)
+let unfocus_proof_step c ps =
+ { ps with
+ goals = undefined ps.defs (unfocus_sublist c ps.goals)
+ }
+
+(* Focuses a tactic at a range of subgoals, found by their indices. *)
+(* arnaud: bug if 0 goals ! *)
+let tclFOCUS i j t env = { go = fun sk fk step ->
+ let (focused,context) = focus_proof_step i j step in
+ (t env).go (fun a fk step -> sk a fk (unfocus_proof_step context step)) fk focused
+}
+
+(* Dispatch tacticals are used to apply a different tactic to each goal under
+ consideration. They come in two flavours:
+ [tclDISPATCH] takes a list of [unit tactic]-s and build a [unit tactic].
+ [tclDISPATCHS] takes a list of ['a sensitive tactic] and returns and returns
+ and ['a sensitive tactic] where the ['a sensitive] interpreted in a goal [g]
+ corresponds to that of the tactic which created [g].
+ It is to be noted that the return value of [tclDISPATCHS ts] makes only
+ sense in the goals immediatly built by it, and would cause an anomaly
+ is used otherwise. *)
+exception SizeMismatch
+(* spiwack: we use an parametrised function to generate the dispatch tacticals.
+ [tclDISPATCHGEN] takes a [null] argument to generate the return value
+ if there are no goal under focus, and a [join] argument to explain how
+ the return value at two given lists of subgoals are combined when
+ both lists are being concatenated.
+ [join] and [null] need be some sort of comutative monoid. *)
+let rec tclDISPATCHGEN null join tacs env = { go = fun sk fk step ->
+ match tacs,step.goals with
+ | [] , [] -> (tclUNIT null env).go sk fk step
+ | t::tacs , first::goals ->
+ (tclDISPATCHGEN null join tacs env).go
+ begin fun x fk step ->
+ match Goal.advance step.defs first with
+ | None -> sk x fk step
+ | Some first ->
+ (t env).go
+ begin fun y fk step' ->
+ sk (join x y) fk { step' with
+ goals = step'.goals@step.goals
+ }
+ end
+ fk
+ { step with goals = [first] }
+ end
+ fk
+ { step with goals = goals }
+ | _ -> raise SizeMismatch
+}
+
+(* takes a tactic which can raise exception and makes it pure by *failing*
+ on with these exceptions. Does not catch anomalies. *)
+let purify t =
+ let t' env = { go = fun sk fk step -> try (t env).go (fun x -> sk (Util.Inl x)) fk step
+ with Util.Anomaly _ as e -> raise e
+ | e -> sk (Util.Inr e) fk step
+ }
+ in
+ tclBIND t' begin function
+ | Util.Inl x -> tclUNIT x
+ | Util.Inr e -> tclZERO e
+ end
+let tclDISPATCHGEN null join tacs = purify (tclDISPATCHGEN null join tacs)
+
+let unitK () () = ()
+let tclDISPATCH = tclDISPATCHGEN () unitK
+
+let extend_to_list =
+ let rec copy n x l =
+ if n < 0 then raise SizeMismatch
+ else if n = 0 then l
+ else copy (n-1) x (x::l)
+ in
+ fun startxs rx endxs l ->
+ let ns = List.length startxs in
+ let ne = List.length endxs in
+ let n = List.length l in
+ startxs@(copy (n-ne-ns) rx endxs)
+let tclEXTEND tacs1 rtac tacs2 env = { go = fun sk fk step ->
+ let tacs = extend_to_list tacs1 rtac tacs2 step.goals in
+ (tclDISPATCH tacs env).go sk fk step
+}
+
+(* [tclGOALBIND] and [tclGOALBINDU] are sorts of bind which take a
+ [Goal.sensitive] as a first argument, the tactic then acts on each goal separately.
+ Allows backtracking between goals. *)
+let list_of_sensitive s k env step =
+ Goal.list_map begin fun defs g ->
+ let (a,defs) = Goal.eval s env defs g in
+ (k a) , defs
+ end step.goals step.defs
+(* In form of a tactic *)
+let list_of_sensitive s k env = { go = fun sk fk step ->
+ let (tacs,defs) = list_of_sensitive s k env step in
+ sk tacs fk { step with defs = defs }
+}
+
+(* This is a helper function for the dispatching tactics (like [tclGOALBIND] and
+ [tclDISPATCHS]). It takes an ['a sensitive] value, and returns a tactic
+ whose return value is, again, ['a sensitive] but only has value in the
+ (unmodified) goals under focus. *)
+let here_s b env = { go = fun sk fk step ->
+ sk (Goal.bind (Goal.here_list step.goals b) (fun b -> b)) fk step
+}
+
+let rec tclGOALBIND s k =
+ (* spiwack: the first line ensures that the value returned by the tactic [k] will
+ not "escape its scope". *)
+ let k a = tclBIND (k a) here_s in
+ purify begin
+ tclBIND (list_of_sensitive s k) begin fun tacs ->
+ tclDISPATCHGEN Goal.null Goal.plus tacs
+ end
+ end
+
+(* spiwack: this should probably be moved closer to the [tclDISPATCH] tactical. *)
+let tclDISPATCHS tacs =
+ let tacs =
+ List.map begin fun tac ->
+ tclBIND tac here_s
+ end tacs
+ in
+ purify begin
+ tclDISPATCHGEN Goal.null Goal.plus tacs
+ end
+
+let rec tclGOALBINDU s k =
+ purify begin
+ tclBIND (list_of_sensitive s k) begin fun tacs ->
+ tclDISPATCHGEN () unitK tacs
+ end
+ end
+
+(* spiwack: up to a few details, same errors are in the Logic module.
+ this should be maintained synchronized, probably. *)
+open Pretype_errors
+let rec catchable_exception = function
+ | Stdpp.Exc_located(_,e) -> catchable_exception e
+ | Util.UserError _ | Type_errors.TypeError _
+ | Indrec.RecursionSchemeError _
+ | Nametab.GlobalizationError _ | PretypeError (_,VarNotFound _)
+ (* unification errors *)
+ | PretypeError(_,(CannotUnify _|CannotUnifyLocal _|CannotGeneralize _
+ |NoOccurrenceFound _|CannotUnifyBindingType _|NotClean _
+ |CannotFindWellTypedAbstraction _
+ |UnsolvableImplicit _)) -> true
+ | Typeclasses_errors.TypeClassError
+ (_, Typeclasses_errors.UnsatisfiableConstraints _) -> true
+ | _ -> false
+
+(* No backtracking can happen here, hence, as opposed to the dispatch tacticals,
+ everything is done in one step. *)
+let sensitive_on_step s env step =
+ let wrap g ((defs, partial_list) as partial_res) =
+ match Goal.advance defs g with
+ | None ->partial_res
+ | Some g ->
+ let {Goal.subgoals = sg } , d' = Goal.eval s env defs g in
+ (d',sg::partial_list)
+ in
+ let ( new_defs , combed_subgoals ) =
+ List.fold_right wrap step.goals (step.defs,[])
+ in
+ { defs = new_defs;
+ goals = List.flatten combed_subgoals }
+let tclSENSITIVE s =
+ purify begin
+ fun env -> { go = fun sk fk step -> sk () fk (sensitive_on_step s env step) }
+ end
+
+module Notations = struct
+ let (>-) = Goal.bind
+ let (>>-) = tclGOALBINDU
+ let (>>--) = tclGOALBIND
+ let (>=) = tclBIND
+ let (>>=) t k = t >= fun s -> s >>- k
+ let (>>==) t k = t >= fun s -> s >>-- k
+ let (<*>) = tclTHEN
+ let (<+>) = tclOR
+end
+
+(*** Compatibility layer with <= 8.2 tactics ***)
+module V82 = struct
+ type tac = Goal.goal Evd.sigma -> Goal.goal list Evd.sigma
+
+ let tactic tac _ = { go = fun sk fk ps ->
+ (* spiwack: we ignore the dependencies between goals here, expectingly
+ preserving the semantics of <= 8.2 tactics *)
+ let tac evd gl =
+ let glsigma = tac { Evd.it = gl ; Evd.sigma = evd } in
+ let sigma = glsigma.Evd.sigma in
+ let g = glsigma.Evd.it in
+ ( g , sigma )
+ in
+ (* Old style tactics expect the goals normalized with respect to evars. *)
+ let (initgoals,initevd) =
+ Goal.list_map Goal.V82.nf_evar ps.goals ps.defs
+ in
+ let (goalss,evd) = Goal.list_map tac initgoals initevd in
+ let sgs = List.flatten goalss in
+ sk () fk { defs = evd ; goals = sgs }
+}
+
+ let has_unresolved_evar pv =
+ let evd = pv.solution in
+ (* arnaud: essayer une procédure moins coûteuse *)
+ not ((Evarutil.non_instantiated evd) = [])
+
+ (* Returns the open goals of the proofview together with the evar_map to
+ interprete them. *)
+ let goals { comb = comb ; solution = solution } =
+ { Evd.it = comb ; sigma = solution}
+
+ let top_goals { initial=initial ; solution=solution } =
+ let goals = List.map (fun (t,_) -> Goal.V82.build (fst (Term.destEvar t))) initial in
+ { Evd.it = goals ; sigma=solution }
+
+ let instantiate_evar n com pv =
+ let (evk,_) =
+ let evl = Evarutil.non_instantiated pv.solution in
+ if (n <= 0) then
+ Util.error "incorrect existential variable index"
+ else if List.length evl < n then
+ Util.error "not so many uninstantiated existential variables"
+ else
+ List.nth evl (n-1)
+ in
+ { pv with
+ solution = Evar_refiner.instantiate_pf_com evk com pv.solution }
+
+ let purify = purify
+end
diff --git a/proofs/proofview.mli b/proofs/proofview.mli
new file mode 100644
index 0000000000..cd5520d4e5
--- /dev/null
+++ b/proofs/proofview.mli
@@ -0,0 +1,203 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* $Id: proofview.mli aspiwack $ *)
+
+(* The proofview datastructure is a pure datastructure underlying the notion
+ of proof (namely, a proof is a proofview which can evolve and has safety
+ mechanisms attached).
+ The general idea of the structure is that it is composed of a chemical
+ solution: an unstructured bag of stuff which has some relations with
+ one another, which represents the various subnodes of the proof, together
+ with a comb: a datastructure that gives order to some of these nodes,
+ namely the open goals.
+ The natural candidate for the solution is an {!Evd.evar_map}, that is
+ a calculus of evars. The comb is then a list of goals (evars wrapped
+ with some extra information, like possible name anotations).
+ There is also need of a list of the evars which initialised the proofview
+ to be able to return information about the proofview. *)
+
+open Term
+
+type proofview
+
+(* Initialises a proofview, the argument is a list of environement,
+ conclusion types, creating that many initial goals. *)
+val init : (Environ.env * Term.types) list -> proofview
+
+(* Returns whether this proofview is finished or not.That is,
+ if it has empty subgoals in the comb. There could still be unsolved
+ subgoaled, but they would then be out of the view, focused out. *)
+val finished : proofview -> bool
+
+(* Returns the current value of the proofview partial proofs. *)
+val return : proofview -> (constr*types) list
+
+
+(*** Focusing operations ***)
+
+(* [IndexOutOfRange] occurs in case of malformed indices
+ with respect to list lengths. *)
+exception IndexOutOfRange
+
+(* Type of the object which allow to unfocus a view.*)
+type focus_context
+
+(* [focus i j] focuses a proofview on the goals from index [i] to index [j]
+ (inclusive). (i.e. goals number [i] to [j] become the only goals of the
+ returned proofview).
+ It returns the focus proof, and a context for the focus trace. *)
+val focus : int -> int -> proofview -> proofview * focus_context
+
+(* Unfocuses a proofview with respect to a context. *)
+val unfocus : focus_context -> proofview -> proofview
+
+(* The tactic monad:
+ - Tactics are objects which apply a transformation to all
+ the subgoals of the current view at the same time. By opposed
+ to the old vision of applying it to a single goal. It mostly
+ allows to consider tactic like [reorder] to reorder the goals
+ in the current view (which might be useful for the tactic designer)
+ (* spiwack: the ordering of goals, though, is actually rather
+ brittle. It would be much more interesting to find a more
+ robust way to adress goals, I have no idea at this time
+ though*)
+ or global automation tactic for dependent subgoals (instantiating
+ an evar has influences on the other goals of the proof in progress,
+ not being able to take that into account causes the current eauto
+ tactic to fail on some instances where it could succeed).
+ - Tactics are a monad ['a tactic], in a sense a tactic can be
+ seens as a function (without argument) which returns a value
+ of type 'a and modifies the environement (in our case: the view).
+ Tactics of course have arguments, but these are given at the
+ meta-level as OCaml functions.
+ Most tactics in the sense we are used to return [ () ], that is
+ no really interesting values. But some might, to pass information
+ around; for instance [Proofview.freeze] allows to store a certain
+ goal sensitive value "at the present time" (which means, considering the
+ structure of the dynamics of proofs, [Proofview.freeze s] will have,
+ for every current goal [gl], and for any of its descendent [g'] in
+ the future the same value in [g'] that in [gl]).
+ (* spiwack: I don't know how much all this relates to F. Kirchner and
+ C. Muñoz. I wasn't able to understand how they used the monad
+ structure in there developpement.
+ *)
+ The tactics seen in Coq's Ltac are (for now at least) only
+ [unit tactic], the return values are kept for the OCaml toolkit.
+ The operation or the monad are [Proofview.tclIDTAC] (which is the
+ "return" of the tactic monad) [Proofview.tclBIND] (which is
+ the "bind") and [Proofview.tclTHEN] (which is a specialized
+ bind on unit-returning tactics).
+*)
+
+
+type +'a tactic
+
+(* Applies a tactic to the current proofview. *)
+val apply : Environ.env -> 'a tactic -> proofview -> proofview
+
+(*** tacticals ***)
+
+(* Unit of the tactic monad *)
+val tclUNIT : 'a -> 'a tactic
+
+(* Bind operation of the tactic monad *)
+val tclBIND : 'a tactic -> ('a -> 'b tactic) -> 'b tactic
+
+(* Interpetes the ";" (semicolon) of Ltac.
+ As a monadic operation, it's a specialized "bind"
+ on unit-returning tactic (meaning "there is no value to bind") *)
+val tclTHEN : unit tactic -> 'a tactic -> 'a tactic
+
+(* [tclIGNORE t] has the same operational content as [t],
+ but drops the value at the end. *)
+val tclIGNORE : 'a tactic -> unit tactic
+
+(* [tclOR t1 t2 = t1] if t1 succeeds and [tclOR t1 t2 = t2] if t2 fails.
+ No interleaving at this point. *)
+val tclOR : 'a tactic -> 'a tactic -> 'a tactic
+
+(* [tclZERO] always fails *)
+val tclZERO : exn -> 'a tactic
+
+(* Focuses a tactic at a range of subgoals, found by their indices. *)
+val tclFOCUS : int -> int -> 'a tactic -> 'a tactic
+
+(* Dispatch tacticals are used to apply a different tactic to each goal under
+ consideration. They come in two flavours:
+ [tclDISPATCH] takes a list of [unit tactic]-s and build a [unit tactic].
+ [tclDISPATCHS] takes a list of ['a sensitive tactic] and returns and returns
+ and ['a sensitive tactic] where the ['a sensitive] interpreted in a goal [g]
+ corresponds to that of the tactic which created [g].
+ It is to be noted that the return value of [tclDISPATCHS ts] makes only
+ sense in the goals immediatly built by it, and would cause an anomaly
+ is used otherwise. *)
+val tclDISPATCH : unit tactic list -> unit tactic
+val tclDISPATCHS : 'a Goal.sensitive tactic list -> 'a Goal.sensitive tactic
+
+(* [tclEXTEND b r e] is a variant to [tclDISPATCH], where the [r] tactic
+ is "repeated" enough time such that every goal has a tactic assigned to it
+ ([b] is the list of tactics applied to the first goals, [e] to the last goals, and [r]
+ is applied to every goal in between. *)
+val tclEXTEND : unit tactic list -> unit tactic -> unit tactic list -> unit tactic
+
+(* A sort of bind which takes a [Goal.sensitive] as a first argument,
+ the tactic then acts on each goal separately.
+ Allows backtracking between goals. *)
+val tclGOALBIND : 'a Goal.sensitive -> ('a -> 'b Goal.sensitive tactic) -> 'b Goal.sensitive tactic
+val tclGOALBINDU : 'a Goal.sensitive -> ('a -> unit tactic) -> unit tactic
+
+(* [tclSENSITIVE] views goal-type tactics as a special kind of tactics.*)
+val tclSENSITIVE : Goal.subgoals Goal.sensitive -> unit tactic
+
+(* Notations for building tactics. *)
+module Notations : sig
+ (* Goal.bind *)
+ val (>-) : 'a Goal.sensitive -> ('a -> 'b Goal.sensitive) -> 'b Goal.sensitive
+ (* tclGOALBINDU *)
+ val (>>-) : 'a Goal.sensitive -> ('a -> unit tactic) -> unit tactic
+ (* tclGOALBIND *)
+ val (>>--) : 'a Goal.sensitive -> ('a -> 'b Goal.sensitive tactic) -> 'b Goal.sensitive tactic
+
+ (* tclBIND *)
+ val (>=) : 'a tactic -> ('a -> 'b tactic) -> 'b tactic
+
+ (* [(>>=)] (and its goal sensitive variant [(>>==)]) "binds" in one step the
+ tactic monad and the goal-sensitive monad.
+ It is strongly advised to use it everytieme an ['a Goal.sensitive tactic]
+ needs a bind, since it usually avoids to delay the interpretation of the
+ goal sensitive value to a location where it does not make sense anymore. *)
+ val (>>=) : 'a Goal.sensitive tactic -> ('a -> unit tactic) -> unit tactic
+ val (>>==) : 'a Goal.sensitive tactic -> ('a -> 'b Goal.sensitive tactic) -> 'b Goal.sensitive tactic
+
+ (* tclTHEN *)
+ val (<*>) : unit tactic -> 'a tactic -> 'a tactic
+ (* tclOR *)
+ val (<+>) : 'a tactic -> 'a tactic -> 'a tactic
+end
+
+(*** Compatibility layer with <= 8.2 tactics ***)
+module V82 : sig
+ type tac = Goal.goal Evd.sigma -> Goal.goal list Evd.sigma
+ val tactic : tac -> unit tactic
+
+ val has_unresolved_evar : proofview -> bool
+
+ (* Returns the open goals of the proofview together with the evar_map to
+ interprete them. *)
+ val goals : proofview -> Goal.goal list Evd.sigma
+
+ val top_goals : proofview -> Goal.goal list Evd.sigma
+
+ (* Implements the Existential command *)
+ val instantiate_evar : int -> Topconstr.constr_expr -> proofview -> proofview
+
+ (* spiwack: [purify] might be useful while writing tactics manipulating exception
+ explicitely or from the [V82] submodule (neither being advised, though *)
+ val purify : 'a tactic -> 'a tactic
+end
diff --git a/proofs/refiner.ml b/proofs/refiner.ml
index a320b67cda..ffb18f2655 100644
--- a/proofs/refiner.ml
+++ b/proofs/refiner.ml
@@ -18,53 +18,20 @@ open Sign
open Environ
open Reductionops
open Type_errors
-open Proof_trees
open Proof_type
open Logic
-type transformation_tactic = proof_tree -> (goal list * validation)
-
-let hypotheses gl = gl.evar_hyps
-let conclusion gl = gl.evar_concl
let sig_it x = x.it
let project x = x.sigma
-let pf_status pf = pf.open_subgoals
-
-let is_complete pf = (0 = (pf_status pf))
-
-let on_open_proofs f pf = if is_complete pf then pf else f pf
let and_status = List.fold_left (+) 0
(* Getting env *)
-let pf_env gls = Global.env_of_context (sig_it gls).evar_hyps
-let pf_hyps gls = named_context_of_val (sig_it gls).evar_hyps
-
-
-let descend n p =
- match p.ref with
- | None -> error "It is a leaf."
- | Some(r,pfl) ->
- if List.length pfl >= n then
- (match list_chop (n-1) pfl with
- | left,(wanted::right) ->
- (wanted,
- (fun pfl' ->
- if false (* debug *) then assert
- (List.length pfl'=1 & (List.hd pfl').goal = wanted.goal);
- let pf' = List.hd pfl' in
- let spfl = left@(pf'::right) in
- let newstatus = and_status (List.map pf_status spfl) in
- { p with
- open_subgoals = newstatus;
- ref = Some(r,spfl) }))
- | _ -> assert false)
- else
- error "Too few subproofs"
-
+let pf_env gls = Global.env_of_context (Goal.V82.hyps (project gls) (sig_it gls))
+let pf_hyps gls = named_context_of_val (Goal.V82.hyps (project gls) (sig_it gls))
(* [mapshape [ l1 ; ... ; lk ] [ v1 ; ... ; vk ] [ p_1 ; .... ; p_(l1+...+lk) ]]
gives
@@ -80,121 +47,9 @@ let rec mapshape nl (fl : (proof_tree list -> proof_tree) list)
let m,l = list_chop h l in
(List.hd fl m) :: (mapshape t (List.tl fl) l)
-(* [frontier : proof_tree -> goal list * validation]
- given a proof [p], [frontier p] gives [(l,v)] where [l] is the list of goals
- to be solved to complete the proof, and [v] is the corresponding
- validation *)
-
-let rec frontier p =
- match p.ref with
- | None ->
- ([p.goal],
- (fun lp' ->
- let p' = List.hd lp' in
- if Evd.eq_evar_info p'.goal p.goal then
- p'
- else
- errorlabstrm "Refiner.frontier"
- (str"frontier was handed back a ill-formed proof.")))
- | Some(r,pfl) ->
- let gll,vl = List.split(List.map frontier pfl) in
- (List.flatten gll,
- (fun retpfl ->
- let pfl' = mapshape (List.map List.length gll) vl retpfl in
- { p with
- open_subgoals = and_status (List.map pf_status pfl');
- ref = Some(r,pfl')}))
-
-(* TODO LEM: I might have to make sure that these hooks are called
- only when called from solve_nth_pftreestate; I can build the hook
- call into the "f", then.
- *)
-let solve_hook = ref ignore
-let set_solve_hook = (:=) solve_hook
-
-let rec frontier_map_rec f n p =
- if n < 1 || n > p.open_subgoals then p else
- match p.ref with
- | None ->
- let p' = f p in
- if Evd.eq_evar_info p'.goal p.goal then
- begin
- !solve_hook p';
- p'
- end
- else
- errorlabstrm "Refiner.frontier_map"
- (str"frontier_map was handed back a ill-formed proof.")
- | Some(r,pfl) ->
- let (_,rpfl') =
- List.fold_left
- (fun (n,acc) p -> (n-p.open_subgoals,frontier_map_rec f n p::acc))
- (n,[]) pfl in
- let pfl' = List.rev rpfl' in
- { p with
- open_subgoals = and_status (List.map pf_status pfl');
- ref = Some(r,pfl')}
-
-let frontier_map f n p =
- let nmax = p.open_subgoals in
- let n = if n < 0 then nmax + n + 1 else n in
- if n < 1 || n > nmax then
- errorlabstrm "Refiner.frontier_map" (str "No such subgoal");
- frontier_map_rec f n p
-
-let rec frontier_mapi_rec f i p =
- if p.open_subgoals = 0 then p else
- match p.ref with
- | None ->
- let p' = f i p in
- if Evd.eq_evar_info p'.goal p.goal then
- begin
- !solve_hook p';
- p'
- end
- else
- errorlabstrm "Refiner.frontier_mapi"
- (str"frontier_mapi was handed back a ill-formed proof.")
- | Some(r,pfl) ->
- let (_,rpfl') =
- List.fold_left
- (fun (n,acc) p -> (n+p.open_subgoals,frontier_mapi_rec f n p::acc))
- (i,[]) pfl in
- let pfl' = List.rev rpfl' in
- { p with
- open_subgoals = and_status (List.map pf_status pfl');
- ref = Some(r,pfl')}
-
-let frontier_mapi f p = frontier_mapi_rec f 1 p
-
-(* [list_pf p] is the lists of goals to be solved in order to complete the
- proof [p] *)
-
-let list_pf p = fst (frontier p)
-
-let rec nb_unsolved_goals pf = pf.open_subgoals
-
-(* leaf g is the canonical incomplete proof of a goal g *)
-
-let leaf g =
- { open_subgoals = 1;
- goal = g;
- ref = None }
-
-(* refiner r is a tactic applying the rule r *)
-
-let check_subproof_connection gl spfl =
- list_for_all2eq (fun g pf -> Evd.eq_evar_info g pf.goal) gl spfl
-
-let abstract_operation syntax semantics gls =
- let (sgl_sigma,validation) = semantics gls in
- let hidden_proof = validation (List.map leaf sgl_sigma.it) in
- (sgl_sigma,
- fun spfl ->
- assert (check_subproof_connection sgl_sigma.it spfl);
- { open_subgoals = and_status (List.map pf_status spfl);
- goal = gls.it;
- ref = Some(Nested(syntax,hidden_proof),spfl)})
+
+let abstract_operation syntax semantics =
+ semantics
let abstract_tactic_expr ?(dflt=false) te tacfun gls =
abstract_operation (Tactic(te,dflt)) tacfun gls
@@ -207,16 +62,11 @@ let abstract_extended_tactic ?(dflt=false) s args =
abstract_tactic ~dflt (Tacexpr.TacExtend (dummy_loc, s, args))
let refiner = function
- | Prim pr as r ->
+ | Prim pr ->
let prim_fun = prim_refiner pr in
(fun goal_sigma ->
let (sgl,sigma') = prim_fun goal_sigma.sigma goal_sigma.it in
- ({it=sgl; sigma = sigma'},
- (fun spfl ->
- assert (check_subproof_connection sgl spfl);
- { open_subgoals = and_status (List.map pf_status spfl);
- goal = goal_sigma.it;
- ref = Some(r,spfl) })))
+ {it=sgl; sigma = sigma'})
| Nested (_,_) | Decl_proof _ ->
@@ -226,83 +76,15 @@ let refiner = function
| Daimon ->
fun gls ->
- ({it=[];sigma=gls.sigma},
- fun spfl ->
- assert (spfl=[]);
- { open_subgoals = 0;
- goal = gls.it;
- ref = Some(Daimon,[])})
+ {it=[];sigma=gls.sigma}
let norm_evar_tac gl = refiner (Prim Change_evars) gl
-let norm_evar_proof sigma pf =
- let nf_subgoal i sgl =
- let (gll,v) = norm_evar_tac {it=sgl.goal;sigma=sigma} in
- v (List.map leaf gll.it) in
- frontier_mapi nf_subgoal pf
-
-(* [extract_open_proof : proof_tree -> constr * (int * constr) list]
- takes a (not necessarly complete) proof and gives a pair (pfterm,obl)
- where pfterm is the constr corresponding to the proof
- and [obl] is an [int*constr list] [ (m1,c1) ; ... ; (mn,cn)]
- where the mi are metavariables numbers, and ci are their types.
- Their proof should be completed in order to complete the initial proof *)
-
-let extract_open_proof sigma pf =
- let next_meta =
- let meta_cnt = ref 0 in
- let rec f () =
- incr meta_cnt;
- if Evd.mem sigma (existential_of_int !meta_cnt) then f ()
- else !meta_cnt
- in f
- in
- let open_obligations = ref [] in
- let rec proof_extractor vl = function
- | {ref=Some(Prim _,_)} as pf -> prim_extractor proof_extractor vl pf
-
- | {ref=Some(Nested(_,hidden_proof),spfl)} ->
- let sgl,v = frontier hidden_proof in
- let flat_proof = v spfl in
- proof_extractor vl flat_proof
-
- | {ref=Some(Decl_proof _,[pf])} -> (proof_extractor vl) pf
-
- | {ref=(None|Some(Daimon,[]));goal=goal} ->
- let visible_rels =
- map_succeed
- (fun id ->
- try let n = proof_variable_index id vl in (n,id)
- with Not_found -> failwith "caught")
- (ids_of_named_context (named_context_of_val goal.evar_hyps)) in
- let sorted_rels =
- Sort.list (fun (n1,_) (n2,_) -> n1 > n2 ) visible_rels in
- let sorted_env =
- List.map (fun (n,id) -> (n,lookup_named_val id goal.evar_hyps))
- sorted_rels in
- let abs_concl =
- List.fold_right (fun (_,decl) c -> mkNamedProd_or_LetIn decl c)
- sorted_env goal.evar_concl in
- let inst = List.filter (fun (_,(_,b,_)) -> b = None) sorted_env in
- let meta = next_meta () in
- open_obligations := (meta,abs_concl):: !open_obligations;
- applist (mkMeta meta, List.map (fun (n,_) -> mkRel n) inst)
-
- | _ -> anomaly "Bug: a case has been forgotten in proof_extractor"
- in
- let pfterm = proof_extractor [] pf in
- (pfterm, List.rev !open_obligations)
-
(*********************)
(* Tacticals *)
(*********************)
-(* unTAC : tactic -> goal sigma -> proof sigma *)
-
-let unTAC tac g =
- let (gl_sigma,v) = tac g in
- { it = v (List.map leaf gl_sigma.it); sigma = gl_sigma.sigma }
let unpackage glsig = (ref (glsig.sigma)),glsig.it
@@ -310,13 +92,9 @@ let repackage r v = {it=v;sigma = !r}
let apply_sig_tac r tac g =
check_for_interrupt (); (* Breakpoint *)
- let glsigma,v = tac (repackage r g) in
+ let glsigma = tac (repackage r g) in
r := glsigma.sigma;
- (glsigma.it,v)
-
-let idtac_valid = function
- [pf] -> pf
- | _ -> anomaly "Refiner.idtac_valid"
+ glsigma.it
(* [goal_goal_list : goal sigma -> goal list sigma] *)
let goal_goal_list gls = {it=[gls.it];sigma=gls.sigma}
@@ -325,7 +103,7 @@ let goal_goal_list gls = {it=[gls.it];sigma=gls.sigma}
let tclNORMEVAR = norm_evar_tac
(* identity tactic without any message *)
-let tclIDTAC gls = (goal_goal_list gls, idtac_valid)
+let tclIDTAC gls = goal_goal_list gls
(* the message printing identity tactic *)
let tclIDTAC_MESSAGE s gls =
@@ -344,23 +122,21 @@ let tclFAIL_lazy lvl s g = raise (FailError (lvl,s))
let start_tac gls =
let (sigr,g) = unpackage gls in
- (sigr,[g],idtac_valid)
+ (sigr,[g])
-let finish_tac (sigr,gl,p) = (repackage sigr gl, p)
+let finish_tac (sigr,gl) = repackage sigr gl
(* Apply [taci.(i)] on the first n subgoals and [tac] on the others *)
-let thens3parts_tac tacfi tac tacli (sigr,gs,p) =
+let thens3parts_tac tacfi tac tacli (sigr,gs) =
let nf = Array.length tacfi in
let nl = Array.length tacli in
let ng = List.length gs in
if ng<nf+nl then errorlabstrm "Refiner.thensn_tac" (str "Not enough subgoals.");
- let gll,pl =
- List.split
+ let gll =
(list_map_i (fun i ->
apply_sig_tac sigr (if i<nf then tacfi.(i) else if i>=ng-nl then tacli.(nl-ng+i) else tac))
0 gs) in
- (sigr, List.flatten gll,
- compose p (mapshape (List.map List.length gll) pl))
+ (sigr,List.flatten gll)
(* Apply [taci.(i)] on the first n subgoals and [tac] on the others *)
let thensf_tac taci tac = thens3parts_tac taci tac [||]
@@ -369,10 +145,10 @@ let thensf_tac taci tac = thens3parts_tac taci tac [||]
let thensl_tac tac taci = thens3parts_tac [||] tac taci
(* Apply [tac i] on the ith subgoal (no subgoals number check) *)
-let thensi_tac tac (sigr,gs,p) =
- let gll,pl =
- List.split (list_map_i (fun i -> apply_sig_tac sigr (tac i)) 1 gs) in
- (sigr, List.flatten gll, compose p (mapshape (List.map List.length gll) pl))
+let thensi_tac tac (sigr,gs) =
+ let gll =
+ list_map_i (fun i -> apply_sig_tac sigr (tac i)) 1 gs in
+ (sigr, List.flatten gll)
let then_tac tac = thensf_tac [||] tac
@@ -382,7 +158,7 @@ let non_existent_goal n =
(* Apply tac on the i-th goal (if i>0). If i<0, then start counting from
the last goal (i=-1). *)
-let theni_tac i tac ((_,gl,_) as subgoals) =
+let theni_tac i tac ((_,gl) as subgoals) =
let nsg = List.length gl in
let k = if i < 0 then nsg + i + 1 else i in
if nsg < 1 then errorlabstrm "theni_tac" (str"No more subgoals.")
@@ -451,42 +227,29 @@ let rec tclTHENLIST = function
let tclMAP tacfun l =
List.fold_right (fun x -> (tclTHEN (tacfun x))) l tclIDTAC
-(* various progress criterions *)
-let same_goal gl subgoal =
- eq_constr (conclusion subgoal) (conclusion gl) &&
- eq_named_context_val (hypotheses subgoal) (hypotheses gl)
-
-
-let weak_progress gls ptree =
- (List.length gls.it <> 1) ||
- (not (same_goal (List.hd gls.it) ptree.it))
-
-let progress gls ptree =
- (progress_evar_map ptree.sigma gls.sigma) ||
- (weak_progress gls ptree)
-
+(* PROGRESS tac ptree applies tac to the goal ptree and fails if tac leaves
+the goal unchanged *)
+let tclWEAK_PROGRESS tac ptree =
+ let rslt = tac ptree in
+ if Goal.V82.weak_progress rslt ptree then rslt
+ else errorlabstrm "Refiner.WEAK_PROGRESS" (str"Failed to progress.")
(* PROGRESS tac ptree applies tac to the goal ptree and fails if tac leaves
the goal unchanged *)
let tclPROGRESS tac ptree =
let rslt = tac ptree in
- if progress (fst rslt) ptree then rslt
+ if Goal.V82.progress rslt ptree then rslt
else errorlabstrm "Refiner.PROGRESS" (str"Failed to progress.")
-(* weak_PROGRESS tac ptree applies tac to the goal ptree and fails
- if tac leaves the goal unchanged, possibly modifying sigma *)
-let tclWEAK_PROGRESS tac ptree =
- let rslt = tac ptree in
- if weak_progress (fst rslt) ptree then rslt
- else errorlabstrm "Refiner.tclWEAK_PROGRESS" (str"Failed to progress.")
-
-
(* Same as tclWEAK_PROGRESS but fails also if tactics generates several goals,
one of them being identical to the original goal *)
let tclNOTSAMEGOAL (tac : tactic) goal =
+ let same_goal gls1 evd2 gl2 =
+ Goal.V82.same_goal gls1.sigma gls1.it evd2 gl2
+ in
let rslt = tac goal in
- let gls = (fst rslt).it in
- if List.exists (same_goal goal.it) gls
+ let {it=gls;sigma=sigma} = rslt in
+ if List.exists (same_goal goal sigma) gls
then errorlabstrm "Refiner.tclNOTSAMEGOAL"
(str"Tactic generated a subgoal identical to the original goal.")
else rslt
@@ -525,9 +288,9 @@ let tclORELSE_THEN t1 t2then t2else gls =
with e -> catch_failerror e; None
with
| None -> t2else gls
- | Some (sgl,v) ->
+ | Some sgl ->
let (sigr,gl) = unpackage sgl in
- finish_tac (then_tac t2then (sigr,gl,v))
+ finish_tac (then_tac t2then (sigr,gl))
(* TRY f tries to apply f, and if it fails, leave the goal unchanged *)
let tclTRY f = (tclORELSE0 f tclIDTAC)
@@ -601,14 +364,12 @@ let rec tclREPEAT_MAIN t g =
(*s Tactics handling a list of goals. *)
-type validation_list = proof_tree list -> proof_tree list
-
-type tactic_list = (goal list sigma) -> (goal list sigma) * validation_list
+type tactic_list = (goal list sigma) -> (goal list sigma)
(* Functions working on goal list for correct backtracking in Prolog *)
let tclFIRSTLIST = tclFIRST
-let tclIDTAC_list gls = (gls, fun x -> x)
+let tclIDTAC_list gls = gls
(* first_goal : goal list sigma -> goal sigma *)
@@ -628,286 +389,20 @@ let apply_tac_list tac glls =
let (sigr,lg) = unpackage glls in
match lg with
| (g1::rest) ->
- let (gl,p) = apply_sig_tac sigr tac g1 in
- let n = List.length gl in
- (repackage sigr (gl@rest),
- fun pfl -> let (pfg,pfrest) = list_chop n pfl in (p pfg)::pfrest)
+ let gl = apply_sig_tac sigr tac g1 in
+ repackage sigr (gl@rest)
| _ -> error "apply_tac_list"
let then_tactic_list tacl1 tacl2 glls =
- let (glls1,pl1) = tacl1 glls in
- let (glls2,pl2) = tacl2 glls1 in
- (glls2, compose pl1 pl2)
+ let glls1 = tacl1 glls in
+ let glls2 = tacl2 glls1 in
+ glls2
(* Transform a tactic_list into a tactic *)
let tactic_list_tactic tac gls =
- let (glres,vl) = tac (goal_goal_list gls) in
- (glres, compose idtac_valid vl)
-
-
-
-(* The type of proof-trees state and a few utilities
- A proof-tree state is built from a proof-tree, a set of global
- constraints, and a stack which allows to navigate inside the
- proof-tree remembering how to rebuild the global proof-tree
- possibly after modification of one of the focused children proof-tree.
- The number in the stack corresponds to
- either the selected subtree and the validation is a function from a
- proof-tree list consisting only of one proof-tree to the global
- proof-tree
- or -1 when the move is done behind a registered tactic in which
- case the validation corresponds to a constant function giving back
- the original proof-tree. *)
-
-type pftreestate = {
- tpf : proof_tree ;
- tpfsigma : evar_map;
- tstack : (int * validation) list }
-
-let proof_of_pftreestate pts = pts.tpf
-let is_top_pftreestate pts = pts.tstack = []
-let cursor_of_pftreestate pts = List.map fst pts.tstack
-let evc_of_pftreestate pts = pts.tpfsigma
-
-let top_goal_of_pftreestate pts =
- { it = goal_of_proof pts.tpf; sigma = pts.tpfsigma }
-
-let nth_goal_of_pftreestate n pts =
- let goals = fst (frontier pts.tpf) in
- try {it = List.nth goals (n-1); sigma = pts.tpfsigma }
- with Invalid_argument _ | Failure _ -> non_existent_goal n
-
-let traverse n pts = match n with
- | 0 -> (* go to the parent *)
- (match pts.tstack with
- | [] -> error "traverse: no ancestors"
- | (_,v)::tl ->
- let pf = v [pts.tpf] in
- let pf = norm_evar_proof pts.tpfsigma pf in
- { tpf = pf;
- tstack = tl;
- tpfsigma = pts.tpfsigma })
- | -1 -> (* go to the hidden tactic-proof, if any, otherwise fail *)
- (match pts.tpf.ref with
- | Some (Nested (_,spf),_) ->
- let v = (fun pfl -> pts.tpf) in
- { tpf = spf;
- tstack = (-1,v)::pts.tstack;
- tpfsigma = pts.tpfsigma }
- | _ -> error "traverse: not a tactic-node")
- | n -> (* when n>0, go to the nth child *)
- let (npf,v) = descend n pts.tpf in
- { tpf = npf;
- tpfsigma = pts.tpfsigma;
- tstack = (n,v):: pts.tstack }
-
-let change_constraints_pftreestate newgc pts = { pts with tpfsigma = newgc }
-
-let app_tac sigr tac p =
- let (gll,v) = tac {it=p.goal;sigma= !sigr} in
- sigr := gll.sigma;
- v (List.map leaf gll.it)
-
-(* modify proof state at current position *)
-
-let map_pftreestate f pts =
- let sigr = ref pts.tpfsigma in
- let tpf' = f sigr pts.tpf in
- let tpf'' =
- if !sigr == pts.tpfsigma then tpf' else norm_evar_proof !sigr tpf' in
- { tpf = tpf'';
- tpfsigma = !sigr;
- tstack = pts.tstack }
-
-(* solve the nth subgoal with tactic tac *)
-
-let solve_nth_pftreestate n tac =
- map_pftreestate
- (fun sigr pt -> frontier_map (app_tac sigr tac) n pt)
-
-let solve_pftreestate = solve_nth_pftreestate 1
-
-(* This function implements a poor man's undo at the current goal.
- This is a gross approximation as it does not attempt to clean correctly
- the global constraints given in tpfsigma. *)
-
-let weak_undo_pftreestate pts =
- let pf = leaf pts.tpf.goal in
- { tpf = pf;
- tpfsigma = pts.tpfsigma;
- tstack = pts.tstack }
-
-(* Gives a new proof (a leaf) of a goal gl *)
-let mk_pftreestate g =
- { tpf = leaf g;
- tstack = [];
- tpfsigma = Evd.empty }
-
-(* Extracts a constr from a proof-tree state ; raises an error if the
- proof is not complete or the state does not correspond to the head
- of the proof-tree *)
-
-let extract_open_pftreestate pts =
- extract_open_proof pts.tpfsigma pts.tpf
-
-let extract_pftreestate pts =
- if pts.tstack <> [] then
- errorlabstrm "extract_pftreestate" (str"Proof blocks need to be closed");
- let pfterm,subgoals = extract_open_pftreestate pts in
- let exl = Evarutil.non_instantiated pts.tpfsigma in
- if subgoals <> [] or exl <> [] then
- errorlabstrm "extract_proof"
- (if subgoals <> [] then
- str "Attempt to save an incomplete proof"
- else
- str "Attempt to save a proof with existential variables still non-instantiated");
- let env = Global.env_of_context pts.tpf.goal.evar_hyps in
- nf_betaiota_preserving_vm_cast env pts.tpfsigma pfterm
- (* strong whd_betaiotaevar env pts.tpfsigma pfterm *)
- (***
- local_strong (Evarutil.whd_ise (ts_it pts.tpfsigma)) pfterm
- ***)
-(* Focus on the first leaf proof in a proof-tree state *)
-
-let rec first_unproven pts =
- let pf = (proof_of_pftreestate pts) in
- if is_complete_proof pf then
- errorlabstrm "first_unproven" (str"No unproven subgoals");
- if is_leaf_proof pf then
- pts
- else
- let childnum =
- list_try_find_i
- (fun n pf ->
- if not(is_complete_proof pf) then n else failwith "caught")
- 1 (children_of_proof pf)
- in
- first_unproven (traverse childnum pts)
-
-(* Focus on the last leaf proof in a proof-tree state *)
-
-let rec last_unproven pts =
- let pf = proof_of_pftreestate pts in
- if is_complete_proof pf then
- errorlabstrm "last_unproven" (str"No unproven subgoals");
- if is_leaf_proof pf then
- pts
- else
- let children = (children_of_proof pf) in
- let nchilds = List.length children in
- let childnum =
- list_try_find_i
- (fun n pf ->
- if not(is_complete_proof pf) then n else failwith "caught")
- 1 (List.rev children)
- in
- last_unproven (traverse (nchilds-childnum+1) pts)
-
-let rec nth_unproven n pts =
- let pf = proof_of_pftreestate pts in
- if is_complete_proof pf then
- errorlabstrm "nth_unproven" (str"No unproven subgoals");
- if is_leaf_proof pf then
- if n = 1 then
- pts
- else
- errorlabstrm "nth_unproven" (str"Not enough unproven subgoals")
- else
- let children = children_of_proof pf in
- let rec process i k = function
- | [] ->
- errorlabstrm "nth_unproven" (str"Not enough unproven subgoals")
- | pf1::rest ->
- let k1 = nb_unsolved_goals pf1 in
- if k1 < k then
- process (i+1) (k-k1) rest
- else
- nth_unproven k (traverse i pts)
- in
- process 1 n children
-
-let rec node_prev_unproven loc pts =
- let pf = proof_of_pftreestate pts in
- match cursor_of_pftreestate pts with
- | [] -> last_unproven pts
- | n::l ->
- if is_complete_proof pf or loc = 1 then
- node_prev_unproven n (traverse 0 pts)
- else
- let child = List.nth (children_of_proof pf) (loc - 2) in
- if is_complete_proof child then
- node_prev_unproven (loc - 1) pts
- else
- first_unproven (traverse (loc - 1) pts)
-
-let rec node_next_unproven loc pts =
- let pf = proof_of_pftreestate pts in
- match cursor_of_pftreestate pts with
- | [] -> first_unproven pts
- | n::l ->
- if is_complete_proof pf ||
- loc = (List.length (children_of_proof pf)) then
- node_next_unproven n (traverse 0 pts)
- else if is_complete_proof (List.nth (children_of_proof pf) loc) then
- node_next_unproven (loc + 1) pts
- else
- last_unproven(traverse (loc + 1) pts)
-
-let next_unproven pts =
- let pf = proof_of_pftreestate pts in
- if is_leaf_proof pf then
- match cursor_of_pftreestate pts with
- | [] -> error "next_unproven"
- | n::_ -> node_next_unproven n (traverse 0 pts)
- else
- node_next_unproven (List.length (children_of_proof pf)) pts
-
-let prev_unproven pts =
- let pf = proof_of_pftreestate pts in
- if is_leaf_proof pf then
- match cursor_of_pftreestate pts with
- | [] -> error "prev_unproven"
- | n::_ -> node_prev_unproven n (traverse 0 pts)
- else
- node_prev_unproven 1 pts
-
-let rec top_of_tree pts =
- if is_top_pftreestate pts then pts else top_of_tree(traverse 0 pts)
-
-(* FIXME: cette fonction n'est (as of October 2007) appelée nulle part *)
-let change_rule f pts =
- let mark_top _ pt =
- match pt.ref with
- Some (oldrule,l) ->
- {pt with ref=Some (f oldrule,l)}
- | _ -> invalid_arg "change_rule" in
- map_pftreestate mark_top pts
-
-let match_rule p pts =
- match (proof_of_pftreestate pts).ref with
- Some (r,_) -> p r
- | None -> false
-
-let rec up_until_matching_rule p pts =
- if is_top_pftreestate pts then
- raise Not_found
- else
- let one_up = traverse 0 pts in
- if match_rule p one_up then
- pts
- else
- up_until_matching_rule p one_up
-
-let rec up_to_matching_rule p pts =
- if match_rule p pts then
- pts
- else
- if is_top_pftreestate pts then
- raise Not_found
- else
- let one_up = traverse 0 pts in
- up_to_matching_rule p one_up
+ let glres = tac (goal_goal_list gls) in
+ glres
(* Change evars *)
let tclEVARS sigma gls = tclIDTAC {gls with sigma=sigma}
@@ -918,28 +413,8 @@ let pp_info = ref (fun _ _ _ -> assert false)
let set_info_printer f = pp_info := f
let tclINFO (tac : tactic) gls =
- let (sgl,v) as res = tac gls in
- begin try
- let pf = v (List.map leaf (sig_it sgl)) in
- let sign = named_context_of_val (sig_it gls).evar_hyps in
- msgnl (hov 0 (str" == " ++
- !pp_info (project gls) sign pf))
- with e when catchable_exception e ->
- msgnl (hov 0 (str "Info failed to apply validation"))
- end;
- res
-
-let pp_proof = ref (fun _ _ _ -> assert false)
-let set_proof_printer f = pp_proof := f
-
-let print_pftreestate {tpf = pf; tpfsigma = sigma; tstack = stack } =
- (if stack = []
- then str "Rooted proof tree is:"
- else (str "Proof tree at occurrence [" ++
- prlist_with_sep (fun () -> str ";") (fun (n,_) -> int n)
- (List.rev stack) ++ str "] is:")) ++ fnl() ++
- !pp_proof sigma (Global.named_context()) pf ++
- Evd.pr_evar_map sigma
+ msgnl (hov 0 (str "Warning: info is currently not working"));
+ tac gls
(* Check that holes in arguments have been resolved *)
@@ -962,5 +437,5 @@ let tclWITHHOLES accept_unresolved_holes tac sigma c gl =
else
let res = tclTHEN (tclEVARS sigma) (tac c) gl in
if not accept_unresolved_holes then
- check_evars (pf_env gl) (fst res).sigma sigma gl;
+ check_evars (pf_env gl) (res).sigma sigma gl;
res
diff --git a/proofs/refiner.mli b/proofs/refiner.mli
index e853c12b7c..77f2e48a7d 100644
--- a/proofs/refiner.mli
+++ b/proofs/refiner.mli
@@ -12,9 +12,9 @@
open Term
open Sign
open Evd
-open Proof_trees
open Proof_type
open Tacexpr
+open Logic
(*i*)
(* The refiner (handles primitive rules and high-level tactics). *)
@@ -28,14 +28,14 @@ val pf_hyps : goal sigma -> named_context
val unpackage : 'a sigma -> evar_map ref * 'a
val repackage : evar_map ref -> 'a -> 'a sigma
val apply_sig_tac :
- evar_map ref -> (goal sigma -> (goal list) sigma * validation) -> goal -> (goal list) * validation
-
-type transformation_tactic = proof_tree -> (goal list * validation)
+ evar_map ref -> (goal sigma -> goal list sigma) -> goal -> goal list
(*s Hiding the implementation of tactics. *)
(* [abstract_tactic tac] hides the (partial) proof produced by [tac] under
a single proof node. The boolean tells if the default tactic is used. *)
+(* spiwack: currently here for compatibility, abstract_operation
+ is a second projection *)
val abstract_operation : compound_rule -> tactic -> tactic
val abstract_tactic : ?dflt:bool -> atomic_tactic_expr -> tactic -> tactic
val abstract_tactic_expr : ?dflt:bool -> tactic_expr -> tactic -> tactic
@@ -43,22 +43,6 @@ val abstract_extended_tactic :
?dflt:bool -> string -> typed_generic_argument list -> tactic -> tactic
val refiner : rule -> tactic
-val frontier : transformation_tactic
-val list_pf : proof_tree -> goal list
-val unTAC : tactic -> goal sigma -> proof_tree sigma
-
-
-(* Install a hook frontier_map and frontier_mapi call on the new node they create *)
-val set_solve_hook : (Proof_type.proof_tree -> unit) -> unit
-(* [frontier_map f n p] applies f on the n-th open subgoal of p and
- rebuilds proof-tree.
- n=1 for first goal, n negative counts from the right *)
-val frontier_map :
- (proof_tree -> proof_tree) -> int -> proof_tree -> proof_tree
-
-(* [frontier_mapi f p] applies (f i) on the i-th open subgoal of p. *)
-val frontier_mapi :
- (int -> proof_tree -> proof_tree) -> proof_tree -> proof_tree
(*s Tacticals. *)
@@ -153,8 +137,8 @@ val tclAT_LEAST_ONCE : tactic -> tactic
val tclFAIL : int -> Pp.std_ppcmds -> tactic
val tclFAIL_lazy : int -> Pp.std_ppcmds Lazy.t -> tactic
val tclDO : int -> tactic -> tactic
-val tclPROGRESS : tactic -> tactic
val tclWEAK_PROGRESS : tactic -> tactic
+val tclPROGRESS : tactic -> tactic
val tclNOTSAMEGOAL : tactic -> tactic
val tclINFO : tactic -> tactic
@@ -173,9 +157,7 @@ val tclIFTHENTRYELSEMUST : tactic -> tactic -> tactic
(*s Tactics handling a list of goals. *)
-type validation_list = proof_tree list -> proof_tree list
-
-type tactic_list = (goal list sigma) -> (goal list sigma) * validation_list
+type tactic_list = goal list sigma -> goal list sigma
val tclFIRSTLIST : tactic_list list -> tactic_list
val tclIDTAC_list : tactic_list
@@ -191,57 +173,4 @@ val goal_goal_list : 'a sigma -> 'a list sigma
extension of the sigma of the goal *)
val tclWITHHOLES : bool -> ('a -> tactic) -> evar_map -> 'a -> tactic
-(*s Functions for handling the state of the proof editor. *)
-
-type pftreestate
-
-val proof_of_pftreestate : pftreestate -> proof_tree
-val cursor_of_pftreestate : pftreestate -> int list
-val is_top_pftreestate : pftreestate -> bool
-val match_rule : (rule -> bool) -> pftreestate -> bool
-val evc_of_pftreestate : pftreestate -> evar_map
-val top_goal_of_pftreestate : pftreestate -> goal sigma
-val nth_goal_of_pftreestate : int -> pftreestate -> goal sigma
-
-val traverse : int -> pftreestate -> pftreestate
-val map_pftreestate :
- (evar_map ref -> proof_tree -> proof_tree) -> pftreestate -> pftreestate
-val solve_nth_pftreestate : int -> tactic -> pftreestate -> pftreestate
-val solve_pftreestate : tactic -> pftreestate -> pftreestate
-
-(* a weak version of logical undoing, that is really correct only *)
-(* if there are no existential variables. *)
-val weak_undo_pftreestate : pftreestate -> pftreestate
-
-val mk_pftreestate : goal -> pftreestate
-val extract_open_proof : evar_map -> proof_tree -> constr * (int * types) list
-val extract_open_pftreestate : pftreestate -> constr * Termops.meta_type_map
-val extract_pftreestate : pftreestate -> constr
-val first_unproven : pftreestate -> pftreestate
-val last_unproven : pftreestate -> pftreestate
-val nth_unproven : int -> pftreestate -> pftreestate
-val node_prev_unproven : int -> pftreestate -> pftreestate
-val node_next_unproven : int -> pftreestate -> pftreestate
-val next_unproven : pftreestate -> pftreestate
-val prev_unproven : pftreestate -> pftreestate
-val top_of_tree : pftreestate -> pftreestate
-val match_rule : (rule -> bool) -> pftreestate -> bool
-val up_until_matching_rule : (rule -> bool) ->
- pftreestate -> pftreestate
-val up_to_matching_rule : (rule -> bool) ->
- pftreestate -> pftreestate
-val change_rule : (rule -> rule) -> pftreestate -> pftreestate
-val change_constraints_pftreestate
- : evar_map -> pftreestate -> pftreestate
-
-
(*s Pretty-printers. *)
-
-(*i*)
-open Pp
-(*i*)
-val set_info_printer :
- (evar_map -> named_context -> proof_tree -> Pp.std_ppcmds) -> unit
-val set_proof_printer :
- (evar_map -> named_context -> proof_tree -> Pp.std_ppcmds) -> unit
-val print_pftreestate : pftreestate -> Pp.std_ppcmds
diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml
index 9e35abfc89..6dbdf17cb8 100644
--- a/proofs/tacmach.ml
+++ b/proofs/tacmach.ml
@@ -21,7 +21,6 @@ open Evd
open Typing
open Redexpr
open Tacred
-open Proof_trees
open Proof_type
open Logic
open Refiner
@@ -34,7 +33,6 @@ let re_sig it gc = { it = it; sigma = gc }
(**************************************************************)
type 'a sigma = 'a Evd.sigma;;
-type validation = Proof_type.validation;;
type tactic = Proof_type.tactic;;
let unpackage = Refiner.unpackage
@@ -46,7 +44,7 @@ let project = Refiner.project
let pf_env = Refiner.pf_env
let pf_hyps = Refiner.pf_hyps
-let pf_concl gls = (sig_it gls).evar_concl
+let pf_concl gls = Goal.V82.concl (project gls) (sig_it gls)
let pf_hyps_types gls =
let sign = Environ.named_context (pf_env gls) in
List.map (fun (id,_,x) -> (id, x)) sign
@@ -123,11 +121,11 @@ let pf_matches = pf_apply Matching.matches_conv
(* Tactics handling a list of goals *)
(************************************)
-type transformation_tactic = proof_tree -> (goal list * validation)
+type transformation_tactic = proof_tree -> goal list
type validation_list = proof_tree list -> proof_tree list
-type tactic_list = (goal list sigma) -> (goal list sigma) * validation_list
+type tactic_list = Refiner.tactic_list
let first_goal = first_goal
let goal_goal_list = goal_goal_list
@@ -138,37 +136,6 @@ let tclFIRSTLIST = tclFIRSTLIST
let tclIDTAC_list = tclIDTAC_list
-(********************************************************)
-(* Functions for handling the state of the proof editor *)
-(********************************************************)
-
-type pftreestate = Refiner.pftreestate
-
-let proof_of_pftreestate = proof_of_pftreestate
-let cursor_of_pftreestate = cursor_of_pftreestate
-let is_top_pftreestate = is_top_pftreestate
-let evc_of_pftreestate = evc_of_pftreestate
-let top_goal_of_pftreestate = top_goal_of_pftreestate
-let nth_goal_of_pftreestate = nth_goal_of_pftreestate
-let traverse = traverse
-let solve_nth_pftreestate = solve_nth_pftreestate
-let solve_pftreestate = solve_pftreestate
-let weak_undo_pftreestate = weak_undo_pftreestate
-let mk_pftreestate = mk_pftreestate
-let extract_pftreestate = extract_pftreestate
-let extract_open_pftreestate = extract_open_pftreestate
-let first_unproven = first_unproven
-let last_unproven = last_unproven
-let nth_unproven = nth_unproven
-let node_prev_unproven = node_prev_unproven
-let node_next_unproven = node_next_unproven
-let next_unproven = next_unproven
-let prev_unproven = prev_unproven
-let top_of_tree = top_of_tree
-let frontier = frontier
-let change_constraints_pftreestate = change_constraints_pftreestate
-
-
(********************************************)
(* Definition of the most primitive tactics *)
(********************************************)
@@ -243,10 +210,18 @@ let rec pr_list f = function
| [] -> mt ()
| a::l1 -> (f a) ++ pr_list f l1
+let db_pr_goal sigma g =
+ let env = Goal.V82.env sigma g in
+ let penv = print_named_context env in
+ let pc = print_constr_env env (Goal.V82.concl sigma g) in
+ str" " ++ hv 0 (penv ++ fnl () ++
+ str "============================" ++ fnl () ++
+ str" " ++ pc) ++ fnl ()
+
let pr_gls gls =
- hov 0 (pr_evar_map (sig_sig gls) ++ fnl () ++ db_pr_goal (sig_it gls))
+ hov 0 (pr_evar_map (sig_sig gls) ++ fnl () ++ db_pr_goal (project gls) (sig_it gls))
let pr_glls glls =
hov 0 (pr_evar_map (sig_sig glls) ++ fnl () ++
- prlist_with_sep pr_fnl db_pr_goal (sig_it glls))
+ prlist_with_sep pr_fnl (db_pr_goal (project glls)) (sig_it glls))
diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli
index a808ca4190..f4bb1d9220 100644
--- a/proofs/tacmach.mli
+++ b/proofs/tacmach.mli
@@ -15,7 +15,6 @@ open Sign
open Environ
open Evd
open Reduction
-open Proof_trees
open Proof_type
open Refiner
open Redexpr
@@ -27,7 +26,6 @@ open Pattern
(* Operations for handling terms under a local typing context. *)
type 'a sigma = 'a Evd.sigma;;
-type validation = Proof_type.validation;;
type tactic = Proof_type.tactic;;
val sig_it : 'a sigma -> 'a
@@ -38,7 +36,7 @@ val re_sig : 'a -> evar_map -> 'a sigma
val unpackage : 'a sigma -> evar_map ref * 'a
val repackage : evar_map ref -> 'a -> 'a sigma
val apply_sig_tac :
- evar_map ref -> (goal sigma -> (goal list) sigma * validation) -> goal -> (goal list) * validation
+ evar_map ref -> (goal sigma -> (goal list) sigma) -> goal -> (goal list)
val pf_concl : goal sigma -> types
val pf_env : goal sigma -> env
@@ -90,38 +88,6 @@ val pf_conv_x_leq : goal sigma -> constr -> constr -> bool
val pf_matches : goal sigma -> constr_pattern -> constr -> patvar_map
val pf_is_matching : goal sigma -> constr_pattern -> constr -> bool
-type transformation_tactic = proof_tree -> (goal list * validation)
-
-val frontier : transformation_tactic
-
-
-(*s Functions for handling the state of the proof editor. *)
-
-type pftreestate = Refiner.pftreestate
-
-val proof_of_pftreestate : pftreestate -> proof_tree
-val cursor_of_pftreestate : pftreestate -> int list
-val is_top_pftreestate : pftreestate -> bool
-val evc_of_pftreestate : pftreestate -> evar_map
-val top_goal_of_pftreestate : pftreestate -> goal sigma
-val nth_goal_of_pftreestate : int -> pftreestate -> goal sigma
-val traverse : int -> pftreestate -> pftreestate
-val weak_undo_pftreestate : pftreestate -> pftreestate
-val solve_nth_pftreestate : int -> tactic -> pftreestate -> pftreestate
-val solve_pftreestate : tactic -> pftreestate -> pftreestate
-val mk_pftreestate : goal -> pftreestate
-val extract_open_pftreestate : pftreestate -> constr * Termops.meta_type_map
-val extract_pftreestate : pftreestate -> constr
-val first_unproven : pftreestate -> pftreestate
-val last_unproven : pftreestate -> pftreestate
-val nth_unproven : int -> pftreestate -> pftreestate
-val node_prev_unproven : int -> pftreestate -> pftreestate
-val node_next_unproven : int -> pftreestate -> pftreestate
-val next_unproven : pftreestate -> pftreestate
-val prev_unproven : pftreestate -> pftreestate
-val top_of_tree : pftreestate -> pftreestate
-val change_constraints_pftreestate :
- evar_map -> pftreestate -> pftreestate
(*s The most primitive tactics. *)
@@ -159,7 +125,7 @@ val rename_hyp : (identifier*identifier) list -> tactic
type validation_list = proof_tree list -> proof_tree list
-type tactic_list = (goal list sigma) -> (goal list sigma) * validation_list
+type tactic_list = Refiner.tactic_list
val first_goal : 'a list sigma -> 'a sigma
val goal_goal_list : 'a sigma -> 'a list sigma
diff --git a/proofs/tactic_debug.ml b/proofs/tactic_debug.ml
index ea8ab5b625..f6b2ce4525 100644
--- a/proofs/tactic_debug.ml
+++ b/proofs/tactic_debug.ml
@@ -36,8 +36,18 @@ let explain_logic_error = ref (fun e -> mt())
let explain_logic_error_no_anomaly = ref (fun e -> mt())
(* Prints the goal *)
+
+let db_pr_goal g =
+ let env = Refiner.pf_env g in
+ let penv = print_named_context env in
+ let pc = print_constr_env env (Goal.V82.concl (Refiner.project g) (Refiner.sig_it g)) in
+ str" " ++ hv 0 (penv ++ fnl () ++
+ str "============================" ++ fnl () ++
+ str" " ++ pc) ++ fnl ()
+
let db_pr_goal g =
- msgnl (str "Goal:" ++ fnl () ++ Proof_trees.db_pr_goal (Refiner.sig_it g))
+ msgnl (str "Goal:" ++ fnl () ++ db_pr_goal g)
+
(* Prints the commands *)
let help () =
diff --git a/tactics/auto.ml b/tactics/auto.ml
index 06c4fab6e1..d2bb1a06fc 100644
--- a/tactics/auto.ml
+++ b/tactics/auto.ml
@@ -280,9 +280,7 @@ let try_head_pattern c =
try head_pattern_bound c
with BoundPattern -> error "Bound head variable."
-let dummy_goal =
- {it = make_evar empty_named_context_val mkProp;
- sigma = empty}
+let dummy_goal = Goal.V82.dummy_goal
let make_exact_entry sigma pri (c,cty) =
let cty = strip_outer_cast cty in
@@ -700,7 +698,8 @@ let print_hint_term cl = ppnl (pr_hint_term cl)
(* print all hints that apply to the concl of the current goal *)
let print_applicable_hint () =
let pts = get_pftreestate () in
- let gl = nth_goal_of_pftreestate 1 pts in
+ let glss = Proof.V82.subgoals pts in
+ let gl = { Evd.it = List.hd glss.Evd.it; sigma = glss.Evd.sigma } in
print_hint_term (pf_concl gl)
(* displays the whole hint database db *)
diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml
index b0645744b4..a0dea0292e 100644
--- a/tactics/autorewrite.ml
+++ b/tactics/autorewrite.ml
@@ -132,16 +132,16 @@ let autorewrite_multi_in ?(conds=Naive) idl tac_main lbas : tactic =
let to_be_cleared = ref false in
fun dir cstr tac gl ->
let last_hyp_id =
- match (Environ.named_context_of_val gl.Evd.it.Evd.evar_hyps) with
+ match Tacmach.pf_hyps gl with
(last_hyp_id,_,_)::_ -> last_hyp_id
| _ -> (* even the hypothesis id is missing *)
error ("No such hypothesis: " ^ (string_of_id !id) ^".")
in
let gl' = general_rewrite_in dir all_occurrences ~tac:(tac, conds) false !id cstr false gl in
- let gls = (fst gl').Evd.it in
+ let gls = gl'.Evd.it in
match gls with
g::_ ->
- (match Environ.named_context_of_val g.Evd.evar_hyps with
+ (match Environ.named_context_of_val (Goal.V82.hyps gl'.Evd.sigma g) with
(lastid,_,_)::_ ->
if last_hyp_id <> lastid then
begin
diff --git a/tactics/class_tactics.ml4 b/tactics/class_tactics.ml4
index 4c58edf595..e0e7aae2fe 100644
--- a/tactics/class_tactics.ml4
+++ b/tactics/class_tactics.ml4
@@ -20,7 +20,6 @@ open Termops
open Sign
open Reduction
open Proof_type
-open Proof_trees
open Declarations
open Tacticals
open Tacmach
@@ -54,17 +53,6 @@ let is_dependent ev evm =
else dep || occur_evar ev evi.evar_concl)
evm false
-let valid goals p res_sigma l =
- let evm =
- List.fold_left2
- (fun sigma (ev, evi) prf ->
- let cstr, obls = Refiner.extract_open_proof !res_sigma prf in
- if not (Evd.is_defined sigma ev) then
- Evd.define ev cstr sigma
- else sigma)
- !res_sigma goals l
- in raise (Found evm)
-
let evar_filter evi =
let hyps' = evar_filtered_context evi in
{ evi with
@@ -78,7 +66,7 @@ let evars_to_goals p evm =
if evi.evar_body = Evar_empty then
let evi', goal = p evm ev evi in
if goal then
- ((ev, evi') :: gls, Evd.add evm' ev evi')
+ ((ev,Goal.V82.build ev) :: gls, Evd.add evm' ev evi')
else (gls, Evd.add evm' ev evi')
else (gls, Evd.add evm' ev evi))
evm ([], Evd.empty)
@@ -223,29 +211,17 @@ let rec catchable = function
| Stdpp.Exc_located (_, e) -> catchable e
| e -> Logic.catchable_exception e
-let is_dep gl gls =
- let evs = Evarutil.evars_of_term gl.evar_concl in
- if evs = Intset.empty then false
- else
- List.fold_left
- (fun b gl ->
- if b then b
- else
- let evs' = Evarutil.evars_of_term gl.evar_concl in
- intersects evs evs')
- false gls
-
let is_ground gl =
Evarutil.is_ground_term (project gl) (pf_concl gl)
let nb_empty_evars s =
Evd.fold (fun ev evi acc -> if evi.evar_body = Evar_empty then succ acc else acc) s 0
-let pr_ev evs ev = Printer.pr_constr_env (Evd.evar_env ev) (Evarutil.nf_evar evs ev.Evd.evar_concl)
+let pr_ev evs ev = Printer.pr_constr_env (Goal.V82.env evs ev) (Evarutil.nf_evar evs (Goal.V82.concl evs ev))
-let typeclasses_debug = ref false
+let pr_depth l = prlist_with_sep (fun () -> str ".") pr_int (List.rev l)
-type validation = evar_map -> proof_tree list -> proof_tree
+let typeclasses_debug = ref false
let pr_depth l = prlist_with_sep (fun () -> str ".") pr_int (List.rev l)
@@ -256,7 +232,7 @@ type 'ans fk = unit -> 'ans
type ('a,'ans) sk = 'a -> 'ans fk -> 'ans
type 'a tac = { skft : 'ans. ('a,'ans) sk -> 'ans fk -> autogoal sigma -> 'ans }
-type auto_result = autogoal list sigma * validation
+type auto_result = autogoal list sigma
type atac = auto_result tac
@@ -281,7 +257,7 @@ let make_resolve_hyp env sigma st flags only_classes pri (id, _, cty) =
else []
let pf_filtered_hyps gls =
- evar_filtered_context (sig_it gls)
+ Goal.V82.filtered_context gls.Evd.sigma (sig_it gls)
let make_autogoal_hints only_classes ?(st=full_transparent_state) g =
let sign = pf_filtered_hyps g in
@@ -292,7 +268,7 @@ let lift_tactic tac (f : goal list sigma -> autoinfo -> autogoal list sigma) : '
{ skft = fun sk fk {it = gl,hints; sigma=s} ->
let res = try Some (tac {it=gl; sigma=s}) with e when catchable e -> None in
match res with
- | Some (gls,v) -> sk (f gls hints, fun _ -> v) fk
+ | Some gls -> sk (f gls hints) fk
| None -> fk () }
let intro_tac : atac =
@@ -300,28 +276,33 @@ let intro_tac : atac =
(fun {it = gls; sigma = s} info ->
let gls' =
List.map (fun g' ->
- let env = evar_env g' in
+ let env = Goal.V82.env s g' in
+ let context = Environ.named_context_of_val (Goal.V82.hyps s g') in
let hint = make_resolve_hyp env s (Hint_db.transparent_state info.hints)
- (true,false,false) info.only_classes None (List.hd (evar_context g')) in
+ (true,false,false) info.only_classes None (List.hd context) in
let ldb = Hint_db.add_list hint info.hints in
(g', { info with is_evar = None; hints = ldb; auto_last_tac = str"intro" })) gls
in {it = gls'; sigma = s})
let normevars_tac : atac =
- lift_tactic tclNORMEVAR
- (fun {it = gls; sigma = s} info ->
- let gls' =
- List.map (fun g' ->
- (g', { info with auto_last_tac = str"NORMEVAR" })) gls
- in {it = gls'; sigma = s})
+ { skft = fun sk fk {it = gl; sigma = s} ->
+ let gl', sigma' = Goal.V82.nf_evar s (fst gl) in
+ sk {it = [gl', snd gl]; sigma = sigma'} fk }
+
+ (* lift_tactic tclNORMEVAR *)
+ (* (fun {it = gls; sigma = s} info -> *)
+ (* let gls' = *)
+ (* List.map (fun g' -> *)
+ (* (g', { info with auto_last_tac = str"NORMEVAR" })) gls *)
+ (* in {it = gls'; sigma = s}) *)
let id_tac : atac =
{ skft = fun sk fk {it = gl; sigma = s} ->
- sk ({it = [gl]; sigma = s}, fun _ pfs -> List.hd pfs) fk }
+ sk {it = [gl]; sigma = s} fk }
(* Ordering of states is lexicographic on the number of remaining goals. *)
-let compare (pri, _, _, (res, _)) (pri', _, _, (res', _)) =
+let compare (pri, _, _, res) (pri', _, _, res') =
let nbgoals s =
List.length (sig_it s) + nb_empty_evars (sig_sig s)
in
@@ -344,11 +325,11 @@ let solve_unif_tac : atac =
let hints_tac hints =
{ skft = fun sk fk {it = gl,info; sigma = s} ->
- let possible_resolve ((lgls,v) as res, pri, b, pp) =
+ let possible_resolve (lgls as res, pri, b, pp) =
(pri, pp, b, res)
in
let tacs =
- let concl = gl.evar_concl in
+ let concl = Goal.V82.concl s gl in
let poss = e_possible_resolve hints info.hints concl in
let l =
let tacgl = {it = gl; sigma = s} in
@@ -358,25 +339,26 @@ let hints_tac hints =
in
if l = [] && !typeclasses_debug then
msgnl (pr_depth info.auto_depth ++ str": no match for " ++
- Printer.pr_constr_env (Evd.evar_env gl) concl ++
+ Printer.pr_constr_env (Goal.V82.env s gl) concl ++
spc () ++ int (List.length poss) ++ str" possibilities");
List.map possible_resolve l
in
let tacs = List.sort compare tacs in
let rec aux i = function
- | (_, pp, b, ({it = gls; sigma = s}, v)) :: tl ->
+ | (_, pp, b, {it = gls; sigma = s}) :: tl ->
if !typeclasses_debug then msgnl (pr_depth (i :: info.auto_depth) ++ str": " ++ pp
++ str" on" ++ spc () ++ pr_ev s gl);
let fk =
(fun () -> (* if !typeclasses_debug then msgnl (str"backtracked after " ++ pp); *)
aux (succ i) tl)
in
- let sgls = evars_to_goals (fun evm ev evi ->
- if Typeclasses.is_resolvable evi &&
- (not info.only_classes || Typeclasses.is_class_evar evm evi) then
- Typeclasses.mark_unresolvable evi, true
- else evi, false) s
- in
+ let sgls = None in
+ (* evars_to_goals (fun evm ev evi -> *)
+ (* if Typeclasses.is_resolvable evi && *)
+ (* (not info.only_classes || Typeclasses.is_class_evar evm evi) then *)
+ (* Typeclasses.mark_unresolvable evi, true *)
+ (* else evi, false) s *)
+ (* in *)
let nbgls, newgls, s' =
let gls' = List.map (fun g -> (None, g)) gls in
match sgls with
@@ -389,12 +371,12 @@ let hints_tac hints =
{ info with auto_depth = j :: i :: info.auto_depth; auto_last_tac = pp;
is_evar = evar;
hints =
- if b && g.evar_hyps <> gl.evar_hyps
+ if b && Goal.V82.hyps s g <> Goal.V82.hyps s gl
then make_autogoal_hints info.only_classes
~st:(Hint_db.transparent_state info.hints) {it = g; sigma = s'}
else info.hints }
in g, info) 1 newgls in
- let glsv = {it = gls'; sigma = s'}, (fun _ pfl -> v (list_firstn nbgls pfl)) in
+ let glsv = {it = gls'; sigma = s'} in
sk glsv fk
| [] -> fk ()
in aux 1 tacs }
@@ -423,9 +405,9 @@ let dependent only_classes evd oev concl =
in not (Intset.is_empty concl_evars)
let then_list (second : atac) (sk : (auto_result, 'a) sk) : (auto_result, 'a) sk =
- let rec aux s (acc : (autogoal list * validation) list) fk = function
+ let rec aux s (acc : autogoal list list) fk = function
| (gl,info) :: gls ->
- second.skft (fun ({it=gls';sigma=s'},v') fk' ->
+ second.skft (fun {it=gls';sigma=s'} fk' ->
let s', needs_backtrack =
if gls' = [] then
match info.is_evar with
@@ -433,30 +415,23 @@ let then_list (second : atac) (sk : (auto_result, 'a) sk) : (auto_result, 'a) sk
let s' =
if Evd.is_defined s' ev then s'
else
- let prf = v' s' [] in
- let term, _ = Refiner.extract_open_proof s' prf in
- Evd.define ev term s'
- in s', dependent info.only_classes s' (Some ev) gl.evar_concl
- | None -> s', dependent info.only_classes s' None gl.evar_concl
+ s'
+ in s', dependent info.only_classes s' (Some ev) (Goal.V82.concl s' gl)
+ | None ->
+ s', dependent info.only_classes s' None (Goal.V82.concl s' gl)
else s', true
in
let fk'' = if not needs_backtrack then
(if !typeclasses_debug then msgnl (str"no backtrack on " ++ pr_ev s gl); fk) else fk'
- in aux s' ((gls',v')::acc) fk'' gls)
+ in aux s' (gls'::acc) fk'' gls)
fk {it = (gl,info); sigma = s}
| [] -> Some (List.rev acc, s, fk)
- in fun ({it = gls; sigma = s},v) fk ->
+ in fun {it = gls; sigma = s} fk ->
let rec aux' = function
| None -> fk ()
| Some (res, s', fk') ->
- let goals' = List.concat (List.map (fun (gls,v) -> gls) res) in
- let v' s' pfs' : proof_tree =
- let (newpfs, rest) = List.fold_left (fun (newpfs,pfs') (gls,v) ->
- let before, after = list_chop (List.length gls) pfs' in
- (v s' before :: newpfs, after))
- ([], pfs') res
- in assert(rest = []); v s' (List.rev newpfs)
- in sk ({it = goals'; sigma = s'}, v') (fun () -> aux' (fk' ()))
+ let goals' = List.concat res in
+ sk {it = goals'; sigma = s'} (fun () -> aux' (fk' ()))
in aux' (aux s [] (fun () -> None) gls)
let then_tac (first : atac) (second : atac) : atac =
@@ -469,7 +444,7 @@ type run_list_res = (auto_result * run_list_res fk) option
let run_list_tac (t : 'a tac) p goals (gl : autogoal list sigma) : run_list_res =
(then_list t (fun x fk -> Some (x, fk)))
- (gl, fun s pfs -> valid goals p (ref s) pfs)
+ gl
(fun _ -> None)
let rec fix (t : 'a tac) : 'a tac =
@@ -488,10 +463,8 @@ let make_autogoals ?(only_classes=true) ?(st=full_transparent_state) gs evm' =
let get_result r =
match r with
| None -> None
- | Some ((gls, v), fk) ->
- try ignore(v (sig_sig gls) []); assert(false)
- with Found evm' -> Some (evm', fk)
-
+ | Some (gls, fk) -> Some (gls.sigma,fk)
+
let run_on_evars ?(only_classes=true) ?(st=full_transparent_state) p evm tac =
match evars_to_goals p evm with
| None -> None (* This happens only because there's no evar having p *)
@@ -508,8 +481,8 @@ let eauto ?(only_classes=true) ?st hints g =
let gl = { it = make_autogoal ~only_classes ?st None g; sigma = project g } in
match run_tac (eauto_tac hints) gl with
| None -> raise Not_found
- | Some ({it = goals; sigma = s}, valid) ->
- {it = List.map fst goals; sigma = s}, valid s
+ | Some {it = goals; sigma = s} ->
+ {it = List.map fst goals; sigma = s}
let real_eauto st hints p evd =
let rec aux evd fails =
@@ -531,17 +504,15 @@ let resolve_all_evars_once debug (mode, depth) p evd =
let db = searchtable_map typeclasses_db in
real_eauto (Hint_db.transparent_state db) [db] p evd
-exception FoundTerm of constr
-
let resolve_one_typeclass env ?(sigma=Evd.empty) gl =
- let gls = { it = Evd.make_evar (Environ.named_context_val env) gl; sigma = sigma } in
+ let (gl,t,sigma) =
+ Goal.V82.mk_goal sigma (Environ.named_context_val env) gl Store.empty in
+ let gls = { it = gl ; sigma = sigma } in
let hints = searchtable_map typeclasses_db in
- let gls', v = eauto ~st:(Hint_db.transparent_state hints) [hints] gls in
- let term = v [] in
+ let gls' = eauto ~st:(Hint_db.transparent_state hints) [hints] gls in
let evd = sig_sig gls' in
- let term = fst (Refiner.extract_open_proof evd term) in
- let term = Evarutil.nf_evar evd term in
- evd, term
+ let term = Evarutil.nf_evar evd t in
+ evd, term
let _ =
Typeclasses.solve_instanciation_problem := (fun x y z -> resolve_one_typeclass x ~sigma:y z)
diff --git a/tactics/eauto.ml4 b/tactics/eauto.ml4
index b0fef2b71b..49af8b40ea 100644
--- a/tactics/eauto.ml4
+++ b/tactics/eauto.ml4
@@ -19,7 +19,6 @@ open Termops
open Sign
open Reduction
open Proof_type
-open Proof_trees
open Declarations
open Tacticals
open Tacmach
@@ -170,7 +169,7 @@ let find_first_goal gls =
type search_state = {
depth : int; (*r depth of search before failing *)
- tacres : goal list sigma * validation;
+ tacres : goal list sigma;
last_tactic : std_ppcmds;
dblist : Auto.hint_db list;
localdb : Auto.hint_db list }
@@ -179,7 +178,7 @@ module SearchProblem = struct
type state = search_state
- let success s = (sig_it (fst s.tacres)) = []
+ let success s = (sig_it s.tacres) = []
let pr_ev evs ev = Printer.pr_constr_env (Evd.evar_env ev) (Evarutil.nf_evar evs ev.Evd.evar_concl)
@@ -187,7 +186,7 @@ module SearchProblem = struct
let evars = Evarutil.nf_evars (Refiner.project gls) in
prlist (pr_ev evars) (sig_it gls)
- let filter_tactics (glls,v) l =
+ let filter_tactics glls l =
(* let _ = Proof_trees.db_pr_goal (List.hd (sig_it glls)) in *)
(* let evars = Evarutil.nf_evars (Refiner.project glls) in *)
(* msg (str"Goal:" ++ pr_ev evars (List.hd (sig_it glls)) ++ str"\n"); *)
@@ -195,11 +194,10 @@ module SearchProblem = struct
| [] -> []
| (tac,pptac) :: tacl ->
try
- let (lgls,ptl) = apply_tac_list tac glls in
- let v' p = v (ptl p) in
+ let lgls = apply_tac_list tac glls in
(* let gl = Proof_trees.db_pr_goal (List.hd (sig_it glls)) in *)
(* msg (hov 1 (pptac ++ str" gives: \n" ++ pr_goals lgls ++ str"\n")); *)
- ((lgls,v'),pptac) :: aux tacl
+ (lgls,pptac) :: aux tacl
with e -> Refiner.catch_failerror e; aux tacl
in aux l
@@ -207,14 +205,14 @@ module SearchProblem = struct
number of remaining goals. *)
let compare s s' =
let d = s'.depth - s.depth in
- let nbgoals s = List.length (sig_it (fst s.tacres)) in
+ let nbgoals s = List.length (sig_it s.tacres) in
if d <> 0 then d else nbgoals s - nbgoals s'
let branching s =
if s.depth = 0 then
[]
else
- let lg = fst s.tacres in
+ let lg = s.tacres in
let nbgl = List.length (sig_it lg) in
assert (nbgl > 0);
let g = find_first_goal lg in
@@ -232,7 +230,7 @@ module SearchProblem = struct
in
let intro_tac =
List.map
- (fun ((lgls,_) as res,pp) ->
+ (fun (lgls as res,pp) ->
let g' = first_goal lgls in
let hintl =
make_resolve_hyp (pf_env g') (project g') (pf_last_hyp g')
@@ -248,7 +246,7 @@ module SearchProblem = struct
filter_tactics s.tacres (e_possible_resolve s.dblist (List.hd s.localdb) (pf_concl g))
in
List.map
- (fun ((lgls,_) as res, pp) ->
+ (fun (lgls as res, pp) ->
let nbgl' = List.length (sig_it lgls) in
if nbgl' < nbgl then
{ depth = s.depth; tacres = res; last_tactic = pp;
diff --git a/tactics/eqdecide.ml4 b/tactics/eqdecide.ml4
index 0d1699b1cf..d5e4ca17d7 100644
--- a/tactics/eqdecide.ml4
+++ b/tactics/eqdecide.ml4
@@ -29,7 +29,6 @@ open Auto
open Pattern
open Matching
open Hipattern
-open Proof_trees
open Proof_type
open Tacmach
open Coqlib
diff --git a/tactics/evar_tactics.ml b/tactics/evar_tactics.ml
index ad392c7d84..906c32c57a 100644
--- a/tactics/evar_tactics.ml
+++ b/tactics/evar_tactics.ml
@@ -33,9 +33,9 @@ let instantiate n (ist,rawc) ido gl =
let sigma = gl.sigma in
let evl =
match ido with
- ConclLocation () -> evar_list sigma gl.it.evar_concl
+ ConclLocation () -> evar_list sigma (pf_concl gl)
| HypLocation (id,hloc) ->
- let decl = Environ.lookup_named_val id gl.it.evar_hyps in
+ let decl = Environ.lookup_named_val id (Goal.V82.hyps sigma (sig_it gl)) in
match hloc with
InHyp ->
(match decl with
diff --git a/tactics/hipattern.ml4 b/tactics/hipattern.ml4
index 9aec0e0914..b6112c34fb 100644
--- a/tactics/hipattern.ml4
+++ b/tactics/hipattern.ml4
@@ -21,7 +21,6 @@ open Reductionops
open Inductiveops
open Evd
open Environ
-open Proof_trees
open Clenv
open Pattern
open Matching
diff --git a/tactics/hipattern.mli b/tactics/hipattern.mli
index d98d2a2b36..9b04a2cd27 100644
--- a/tactics/hipattern.mli
+++ b/tactics/hipattern.mli
@@ -15,7 +15,6 @@ open Term
open Sign
open Evd
open Pattern
-open Proof_trees
open Coqlib
(*i*)
diff --git a/tactics/leminv.ml b/tactics/leminv.ml
index d0f6e82263..395a7c206e 100644
--- a/tactics/leminv.ml
+++ b/tactics/leminv.ml
@@ -24,7 +24,6 @@ open Entries
open Inductiveops
open Environ
open Tacmach
-open Proof_trees
open Proof_type
open Pfedit
open Evar_refiner
@@ -217,29 +216,31 @@ let inversion_scheme env sigma t sort dep_option inv_op =
errorlabstrm "lemma_inversion"
(str"Computed inversion goal was not closed in initial signature.");
*)
- let invSign = named_context_val invEnv in
- let pfs = mk_pftreestate (mk_goal invSign invGoal None) in
- let pfs = solve_pftreestate (tclTHEN intro (onLastHypId inv_op)) pfs in
- let (pfterm,meta_types) = extract_open_pftreestate pfs in
+ let pf = Proof.start [invEnv,invGoal] in
+ Proof.run_tactic env (Proofview.V82.tactic (tclTHEN intro (onLastHypId inv_op))) pf;
+ let pfterm = List.hd (Proof.partial_proof pf) in
let global_named_context = Global.named_context () in
- let ownSign =
+ let ownSign = ref begin
fold_named_context
(fun env (id,_,_ as d) sign ->
if mem_named_context id global_named_context then sign
else add_named_decl d sign)
invEnv ~init:empty_named_context
- in
- let (_,ownSign,mvb) =
- List.fold_left
- (fun (avoid,sign,mvb) (mv,mvty) ->
- let h = next_ident_away (id_of_string "H") avoid in
- (h::avoid, add_named_decl (h,None,mvty) sign, (mv,mkVar h)::mvb))
- (ids_of_context invEnv, ownSign, [])
- meta_types
+ end in
+ let avoid = ref [] in
+ let { sigma=sigma } = Proof.V82.subgoals pf in
+ let rec fill_holes c =
+ match kind_of_term c with
+ | Evar (e,_) ->
+ let h = next_ident_away (id_of_string "H") !avoid in
+ let ty = (Evd.find sigma e).evar_concl in
+ avoid := h::!avoid;
+ ownSign := add_named_decl (h,None,ty) !ownSign;
+ mkVar h
+ | _ -> map_constr fill_holes c
in
let invProof =
- it_mkNamedLambda_or_LetIn
- (local_strong (fun _ -> whd_meta mvb) Evd.empty pfterm) ownSign
+ it_mkNamedLambda_or_LetIn (fill_holes pfterm) !ownSign
in
invProof
@@ -255,26 +256,17 @@ let add_inversion_lemma name env sigma t sort dep inv_op =
IsProof Lemma)
in ()
-(* open Pfedit *)
-
(* inv_op = Inv (derives de complete inv. lemma)
* inv_op = InvNoThining (derives de semi inversion lemma) *)
let inversion_lemma_from_goal n na (loc,id) sort dep_option inv_op =
let pts = get_pftreestate() in
- let gl = nth_goal_of_pftreestate n pts in
+ let { it=gls ; sigma=sigma } = Proof.V82.subgoals pts in
+ let gl = { it = List.nth gls (n-1) ; sigma=sigma } in
let t =
try pf_get_hyp_typ gl id
with Not_found -> Pretype_errors.error_var_not_found_loc loc id in
let env = pf_env gl and sigma = project gl in
-(* Pourquoi ???
- let fv = global_vars env t in
- let thin_ids = thin_ids (hyps,fv) in
- if not(list_subset thin_ids fv) then
- errorlabstrm "lemma_inversion"
- (str"Cannot compute lemma inversion when there are" ++ spc () ++
- str"free variables in the types of an inductive" ++ spc () ++
- str"which are not free in its instance."); *)
add_inversion_lemma na env sigma t sort dep_option inv_op
let add_inversion_lemma_exn na com comsort bool tac =
diff --git a/tactics/refine.ml b/tactics/refine.ml
index 67a73b9be5..cb6cb961f4 100644
--- a/tactics/refine.ml
+++ b/tactics/refine.ml
@@ -382,5 +382,3 @@ let refine (evd,c) gl =
complicated to update meta types when passing through a binder *)
let th = compute_metamap (pf_env gl) evd c in
tclTHEN (Refiner.tclEVARS evd) (tcc_aux [] th) gl
-
-let _ = Decl_proof_instr.set_refine refine (* dirty trick to solve circular dependency *)
diff --git a/tactics/rewrite.ml4 b/tactics/rewrite.ml4
index e08e5e9ede..1af2d33988 100644
--- a/tactics/rewrite.ml4
+++ b/tactics/rewrite.ml4
@@ -21,7 +21,6 @@ open Termops
open Sign
open Reduction
open Proof_type
-open Proof_trees
open Declarations
open Tacticals
open Tacmach
@@ -569,7 +568,7 @@ let apply_rule hypinfo loccs : strategy =
if eq_constr t c2 then Some None
else
let goalevars = Evd.evar_merge (fst evars)
- (Evd.undefined_evars (Evarutil.nf_evar_map env'.evd))
+ (Evd.undefined_evars env'.evd)
in
let res = { rew_car = ty; rew_from = c1;
rew_to = c2; rew_prf = RewPrf (rel, prf); rew_evars = goalevars, snd evars }
@@ -908,8 +907,8 @@ module Strategies =
let hints (db : string) : strategy =
fun env sigma t ty cstr evars ->
- let rules = Autorewrite.find_matches db t in
- lemmas (List.map (fun hint -> (inj_open hint.Autorewrite.rew_lemma, hint.Autorewrite.rew_l2r)) rules)
+ let rules = Autorewrite.find_matches db t in
+ lemmas (List.map (fun hint -> (inj_open hint.Autorewrite.rew_lemma, hint.Autorewrite.rew_l2r)) rules)
env sigma t ty cstr evars
let reduce (r : Redexpr.red_expr) : strategy =
@@ -1051,8 +1050,8 @@ let cl_rewrite_clause_aux ?(abs=None) strat goal_meta clause gl =
change_in_concl None newt)
in
let evartac =
- if not (undef = Evd.empty) then
- Refiner.tclEVARS undef
+ if not (Evd.is_empty undef) then
+ Refiner.tclEVARS evars
else tclIDTAC
in tclTHENLIST [evartac; rewtac] gl
with
@@ -1575,9 +1574,10 @@ let general_s_rewrite cl l2r occs (c,l) ~new_goals gl =
let meta = Evarutil.new_meta() in
let hypinfo, strat = apply_lemma gl c cl l2r occs in
try
- tclTHEN
- (Refiner.tclEVARS hypinfo.cl.evd)
- (cl_rewrite_clause_aux ~abs:hypinfo.abs strat meta cl) gl
+ tclWEAK_PROGRESS
+ (tclTHEN
+ (Refiner.tclEVARS hypinfo.cl.evd)
+ (cl_rewrite_clause_aux ~abs:hypinfo.abs strat meta cl)) gl
with RewriteFailure ->
let {l2r=l2r; c1=x; c2=y} = hypinfo in
raise (Pretype_errors.PretypeError
diff --git a/tactics/tacinterp.ml b/tactics/tacinterp.ml
index 0e352110a0..6e3957ac05 100644
--- a/tactics/tacinterp.ml
+++ b/tactics/tacinterp.ml
@@ -73,7 +73,7 @@ type ltac_type =
(* Values for interpretation *)
type value =
- | VRTactic of (goal list sigma * validation) (* For Match results *)
+ | VRTactic of (goal list sigma) (* For Match results *)
(* Not a true value *)
| VFun of ltac_trace * (identifier*value) list *
identifier option list * glob_tactic_expr
@@ -347,11 +347,6 @@ let vars_of_ist (lfun,_,_,env) =
List.fold_left (fun s id -> Idset.add id s)
(vars_of_env env) lfun
-let get_current_context () =
- try Pfedit.get_current_goal_context ()
- with e when Logic.catchable_exception e ->
- (Evd.empty, Global.env())
-
let strict_check = ref false
let adjust_loc loc = if !strict_check then dloc else loc
@@ -1794,11 +1789,11 @@ let mk_int_or_var_value ist c = VInteger (interp_int_or_var ist c)
let pack_sigma (sigma,c) = {it=c;sigma=sigma}
-let extend_gl_hyps gl sign =
- { gl with
- it = { gl.it with
- evar_hyps =
- List.fold_right Environ.push_named_context_val sign gl.it.evar_hyps } }
+let extend_gl_hyps { it=gl ; sigma=sigma } sign =
+ let hyps = Goal.V82.hyps sigma gl in
+ let new_hyps = List.fold_right Environ.push_named_context_val sign hyps in
+ (* spiwack: (2010/01/13) if a bug was reintroduced in [change] in is probably here *)
+ Goal.V82.new_goal_with sigma gl new_hyps
(* Interprets an l-tac expression into a value *)
let rec val_interp ist gl (tac:glob_tactic_expr) =
@@ -1988,6 +1983,8 @@ and interp_letin ist gl llc u =
(* Interprets the Match Context expressions *)
and interp_match_goal ist goal lz lr lmr =
+ let (gl,sigma) = Goal.V82.nf_evar (project goal) (sig_it goal) in
+ let goal = { it = gl ; sigma = sigma } in
let hyps = pf_hyps goal in
let hyps = if lr then List.rev hyps else hyps in
let concl = pf_concl goal in
diff --git a/tactics/tacinterp.mli b/tactics/tacinterp.mli
index fe4e9f6aab..71ee29f8cc 100644
--- a/tactics/tacinterp.mli
+++ b/tactics/tacinterp.mli
@@ -25,7 +25,7 @@ open Redexpr
(* Values for interpretation *)
type value =
- | VRTactic of (goal list sigma * validation)
+ | VRTactic of (goal list sigma)
| VFun of ltac_trace * (identifier*value) list *
identifier option list * glob_tactic_expr
| VVoid
diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml
index 332855052d..045f70c61b 100644
--- a/tactics/tacticals.ml
+++ b/tactics/tacticals.ml
@@ -61,8 +61,8 @@ let tclAT_LEAST_ONCE = Refiner.tclAT_LEAST_ONCE
let tclFAIL = Refiner.tclFAIL
let tclFAIL_lazy = Refiner.tclFAIL_lazy
let tclDO = Refiner.tclDO
-let tclPROGRESS = Refiner.tclPROGRESS
let tclWEAK_PROGRESS = Refiner.tclWEAK_PROGRESS
+let tclPROGRESS = Refiner.tclPROGRESS
let tclNOTSAMEGOAL = Refiner.tclNOTSAMEGOAL
let tclTHENTRY = Refiner.tclTHENTRY
let tclIFTHENELSE = Refiner.tclIFTHENELSE
diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli
index b9c8ab928b..3dd73c92c3 100644
--- a/tactics/tacticals.mli
+++ b/tactics/tacticals.mli
@@ -55,8 +55,8 @@ val tclAT_LEAST_ONCE : tactic -> tactic
val tclFAIL : int -> std_ppcmds -> tactic
val tclFAIL_lazy : int -> std_ppcmds Lazy.t -> tactic
val tclDO : int -> tactic -> tactic
-val tclPROGRESS : tactic -> tactic
val tclWEAK_PROGRESS : tactic -> tactic
+val tclPROGRESS : tactic -> tactic
val tclNOTSAMEGOAL : tactic -> tactic
val tclTHENTRY : tactic -> tactic -> tactic
val tclMAP : ('a -> tactic) -> 'a list -> tactic
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index 41fab4e716..e6201aad9d 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -27,7 +27,6 @@ open Pfedit
open Tacred
open Rawterm
open Tacmach
-open Proof_trees
open Proof_type
open Logic
open Evar_refiner
@@ -516,8 +515,8 @@ let intro_move idopt hto = match idopt with
| Some id -> intro_gen dloc (IntroMustBe id) hto true false
let pf_lookup_hypothesis_as_renamed env ccl = function
- | AnonHyp n -> pf_lookup_index_as_renamed env ccl n
- | NamedHyp id -> pf_lookup_name_as_displayed env ccl id
+ | AnonHyp n -> Detyping.lookup_index_as_renamed env ccl n
+ | NamedHyp id -> Detyping.lookup_name_as_displayed env ccl id
let pf_lookup_hypothesis_as_renamed_gen red h gl =
let env = pf_env gl in
@@ -614,7 +613,7 @@ let bring_hyps hyps =
let resolve_classes gl =
let env = pf_env gl and evd = project gl in
- if evd = Evd.empty then tclIDTAC gl
+ if Evd.is_empty evd then tclIDTAC gl
else
let evd' = Typeclasses.resolve_typeclasses env (Evd.create_evar_defs evd) in
(tclTHEN (tclEVARS evd') tclNORMEVAR) gl
diff --git a/tactics/tactics.mllib b/tactics/tactics.mllib
index b885b15243..333d6a3a21 100644
--- a/tactics/tactics.mllib
+++ b/tactics/tactics.mllib
@@ -19,6 +19,4 @@ Leminv
Tacinterp
Evar_tactics
Autorewrite
-Decl_interp
-Decl_proof_instr
Tactic_option
diff --git a/theories/Classes/RelationClasses.v b/theories/Classes/RelationClasses.v
index 0d9c9ae9d5..a8193ef3c8 100644
--- a/theories/Classes/RelationClasses.v
+++ b/theories/Classes/RelationClasses.v
@@ -393,7 +393,7 @@ Program Instance subrelation_partial_order :
Next Obligation.
Proof.
- unfold relation_equivalence in *. firstorder.
+ unfold relation_equivalence in *. compute; firstorder.
Qed.
Typeclasses Opaque arrows predicate_implication predicate_equivalence
diff --git a/theories/Init/Prelude.v b/theories/Init/Prelude.v
index 685c724703..45b949716c 100644
--- a/theories/Init/Prelude.v
+++ b/theories/Init/Prelude.v
@@ -18,9 +18,11 @@ Require Export Coq.Init.Tactics.
(* Initially available plugins
(+ nat_syntax_plugin loaded in Datatypes) *)
Declare ML Module "extraction_plugin".
+Declare ML Module "decl_mode_plugin".
Declare ML Module "cc_plugin".
Declare ML Module "ground_plugin".
Declare ML Module "dp_plugin".
Declare ML Module "recdef_plugin".
Declare ML Module "subtac_plugin".
Declare ML Module "xml_plugin".
+Global Set Default Proof Mode "Classic".
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v
index ee12c6a8db..e2c6232013 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v
@@ -1097,7 +1097,7 @@ intros x; case x; simpl ww_is_even.
rewrite wwB_wBwB; rewrite Zpower_2.
apply Zmult_le_compat_r; auto with zarith.
case (spec_to_Z w4);auto with zarith.
- Qed.
+Qed.
Lemma spec_ww_is_zero: forall x,
if ww_is_zero x then [[x]] = 0 else 0 < [[x]].
diff --git a/theories/Numbers/Cyclic/Int31/Cyclic31.v b/theories/Numbers/Cyclic/Int31/Cyclic31.v
index 4fd2a8f0a2..e2e5b563f2 100644
--- a/theories/Numbers/Cyclic/Int31/Cyclic31.v
+++ b/theories/Numbers/Cyclic/Int31/Cyclic31.v
@@ -2387,7 +2387,7 @@ Section Int31_Specs.
replace [|il|] with (([|il|] - [|il1|]) + [|il1|]); try ring.
rewrite <-Hil2.
change (-1 * 2 ^ Z_of_nat size) with (-base); ring.
- Qed.
+Qed.
(** [iszero] *)
diff --git a/theories/Program/Tactics.v b/theories/Program/Tactics.v
index 6f7645841c..d187941d20 100644
--- a/theories/Program/Tactics.v
+++ b/theories/Program/Tactics.v
@@ -304,8 +304,10 @@ Ltac refine_hyp c :=
possibly using [program_simplify] to use standard goal-cleaning tactics. *)
Ltac program_simplify :=
- simpl in |- *; intros ; destruct_all_rec_calls ; repeat (destruct_conjs; simpl proj1_sig in *);
- subst*; autoinjections ; try discriminates ;
+(* arnaud: enlever la première ligne en faveur de la suite. *)
+simpl ; intros ; destruct_conjs ; simpl proj1_sig in * ; subst* ; autoinjections ; try discriminates ;
+(* arnaud: restore: simpl in |- *; intros ; destruct_all_rec_calls ; repeat (destruct_conjs; simpl proj1_sig in * );
+ subst*; autoinjections ; try discriminates ;*)
try (solve [ red ; intros ; destruct_conjs ; autoinjections ; discriminates ]).
Ltac program_solve_wf :=
diff --git a/theories/Reals/RiemannInt.v b/theories/Reals/RiemannInt.v
index ae2c3d77f8..9fd6e00882 100644
--- a/theories/Reals/RiemannInt.v
+++ b/theories/Reals/RiemannInt.v
@@ -2242,7 +2242,7 @@ Proof.
unfold RiemannInt_SF in |- *; case (Rle_dec a b); intro.
eapply StepFun_P17.
apply StepFun_P1.
- simpl in |- *; apply StepFun_P1.
+ simpl in |- *; apply StepFun_P1.
apply Ropp_eq_compat; eapply StepFun_P17.
apply StepFun_P1.
simpl in |- *; apply StepFun_P1.
diff --git a/toplevel/auto_ind_decl.ml b/toplevel/auto_ind_decl.ml
index 0e66c43c3b..d8c7326a9e 100644
--- a/toplevel/auto_ind_decl.ml
+++ b/toplevel/auto_ind_decl.ml
@@ -551,7 +551,7 @@ repeat ( apply andb_prop in z;let z1:= fresh "Z" in destruct z as [z1 z]).
Ci a1 ... an = Ci b1 ... bn
replace bi with ai; auto || replace bi with ai by apply typeofbi_prod ; auto
*)
- fun gls-> let gl = (gls.Evd.it).Evd.evar_concl in
+ fun gls-> let gl = pf_concl gls in
match (kind_of_term gl) with
| App (c,ca) -> (
match (kind_of_term c) with
@@ -674,7 +674,7 @@ let compute_lb_tact lb_scheme_key ind lnamesparrec nparrec gsig =
tclTHENSEQ [apply (andb_true_intro());
simplest_split ;Auto.default_auto ]
);
- fun gls -> let gl = (gls.Evd.it).Evd.evar_concl in
+ fun gls -> let gl = pf_concl gls in
(* assume the goal to be eq (eq_type ...) = true *)
match (kind_of_term gl) with
| App(c,ca) -> (match (kind_of_term ca.(1)) with
diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml
index 01e83b1ad0..0fb9a33b2b 100644
--- a/toplevel/himsg.ml
+++ b/toplevel/himsg.ml
@@ -512,7 +512,14 @@ let explain_no_instance env (_,id) l =
str "applied to arguments" ++ spc () ++
prlist_with_sep pr_spc (pr_lconstr_env env) l
+let undefined_evars evm =
+ Evd.fold (fun ev evi undef ->
+ if evi.evar_body = Evar_empty then
+ Evd.add undef ev (Evarutil.nf_evar_info evm evi)
+ else undef) evm Evd.empty
+
let pr_constraints printenv env evm =
+ let evm = undefined_evars evm in
let l = Evd.to_list evm in
let (ev, evi) = List.hd l in
if List.for_all (fun (ev', evi') ->
diff --git a/toplevel/lemmas.mli b/toplevel/lemmas.mli
index 5d4f014a36..f4e9f6f587 100644
--- a/toplevel/lemmas.mli
+++ b/toplevel/lemmas.mli
@@ -36,7 +36,7 @@ val start_proof_with_initialization :
declaration_hook -> unit
(* A hook the next three functions pass to cook_proof *)
-val set_save_hook : (Refiner.pftreestate -> unit) -> unit
+val set_save_hook : (Proof.proof -> unit) -> unit
(*s [save_named b] saves the current completed proof under the name it
was started; boolean [b] tells if the theorem is declared opaque; it
diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml
index 886f4033d8..36f1e96e94 100644
--- a/toplevel/vernacentries.ml
+++ b/toplevel/vernacentries.ml
@@ -19,8 +19,6 @@ open Nameops
open Term
open Pfedit
open Tacmach
-open Proof_trees
-open Decl_mode
open Constrintern
open Prettyp
open Printer
@@ -66,52 +64,34 @@ let cl_of_qualid = function
(* "Show" commands *)
let show_proof () =
- let pts = get_pftreestate () in
- let cursor = cursor_of_pftreestate pts in
- let evc = evc_of_pftreestate pts in
- let (pfterm,meta_types) = extract_open_pftreestate pts in
- msgnl (str"LOC: " ++
- prlist_with_sep pr_spc pr_int (List.rev cursor) ++ fnl () ++
- str"Subgoals" ++ fnl () ++
- prlist (fun (mv,ty) -> Nameops.pr_meta mv ++ str" -> " ++
- pr_ltype ty ++ fnl ())
- meta_types
- ++ str"Proof: " ++ pr_lconstr (Evarutil.nf_evar evc pfterm))
+ (* spiwack: this would probably be cooler with a bit of polishing. *)
+ let p = Proof_global.give_me_the_proof () in
+ let pprf = Proof.partial_proof p in
+ msgnl (Util.prlist_with_sep Pp.fnl Printer.pr_constr pprf)
let show_node () =
- let pts = get_pftreestate () in
- let pf = proof_of_pftreestate pts
- and cursor = cursor_of_pftreestate pts in
- msgnl (prlist_with_sep pr_spc pr_int (List.rev cursor) ++ fnl () ++
- pr_goal (goal_of_proof pf) ++ fnl () ++
- (match pf.Proof_type.ref with
- | None -> (str"BY <rule>")
- | Some(r,spfl) ->
- (str"BY " ++ pr_rule r ++ fnl () ++
- str" " ++
- hov 0 (prlist_with_sep pr_fnl pr_goal
- (List.map goal_of_proof spfl)))))
+ (* spiwack: I'm have little clue what this function used to do. I deactivated it,
+ could, possibly, be cleaned away. (Feb. 2010) *)
+ ()
let show_script () =
- let pts = get_pftreestate () in
- let pf = proof_of_pftreestate pts
- and evc = evc_of_pftreestate pts in
- msgnl_with !Pp_control.deep_ft (print_treescript evc pf)
+ (* spiwack: show_script is currently not working *)
+ ()
let show_thesis () =
msgnl (anomaly "TODO" )
let show_top_evars () =
+ (* spiwack: new as of Feb. 2010: shows goal evars in addition to non-goal evars. *)
let pfts = get_pftreestate () in
- let gls = top_goal_of_pftreestate pfts in
- let sigma = project gls in
+ let gls = Proof.V82.subgoals pfts in
+ let sigma = gls.Evd.sigma in
msg (pr_evars_int 1 (Evarutil.non_instantiated sigma))
+
let show_prooftree () =
- let pts = get_pftreestate () in
- let pf = proof_of_pftreestate pts
- and evc = evc_of_pftreestate pts in
- msg (print_proof evc (Global.named_context()) pf)
+ (* Spiwack: proof tree is currently not working *)
+ ()
let print_subgoals () = if_verbose (fun () -> msg (pr_open_subgoals ())) ()
@@ -119,7 +99,8 @@ let print_subgoals () = if_verbose (fun () -> msg (pr_open_subgoals ())) ()
let show_intro all =
let pf = get_pftreestate() in
- let gl = nth_goal_of_pftreestate 1 pf in
+ let {Evd.it=gls ; sigma=sigma} = Proof.V82.subgoals pf in
+ let gl = {Evd.it=List.hd gls ; sigma = sigma} in
let l,_= decompose_prod_assum (strip_outer_cast (pf_concl gl)) in
if all
then
@@ -349,14 +330,10 @@ let vernac_end_proof = function
the theories [??] *)
let vernac_exact_proof c =
- let pfs = top_of_tree (get_pftreestate()) in
- let pf = proof_of_pftreestate pfs in
- if (is_leaf_proof pf) then begin
- by (Tactics.exact_proof c);
- save_named true end
- else
- errorlabstrm "Vernacentries.ExactProof"
- (strbrk "Command 'Proof ...' can only be used at the beginning of the proof.")
+ (* spiwack: for simplicity I do not enforce that "Proof proof_term" is
+ called only at the begining of a proof. *)
+ by (Tactics.exact_proof c);
+ save_named true
let vernac_assumption kind l nl=
if Pfedit.refining () then
@@ -626,24 +603,61 @@ let vernac_declare_class id =
(***********)
(* Solving *)
-let vernac_solve n tcom b =
+
+let command_focus = Proof.new_focus_kind ()
+let focus_command_cond = Proof.no_cond command_focus
+
+(* Gestion of bullets. *)
+open Store.Field
+(* spiwack: we need only one focus kind as we keep a stack of (distinct!) bullets *)
+let bullet_kind = Proof.new_focus_kind ()
+let bullet_cond = Proof.done_cond bullet_kind
+let (get_bullets,set_bullets) =
+ let bullets = Store.field () in
+ ( begin fun pr -> Option.default [] (bullets.get (Proof.get_proof_info pr)) end ,
+ begin fun bs pr -> Proof.set_proof_info (bullets.set bs (Proof.get_proof_info pr)) pr end )
+let has_bullet bul pr =
+ let rec has_bullet = function
+ | b'::_ when bul=b' -> true
+ | _::l -> has_bullet l
+ | [] -> false
+ in
+ has_bullet (get_bullets pr)
+(* precondition: the stack is not empty *)
+let pop_bullet pr =
+ match get_bullets pr with
+ | b::stk -> Proof.unfocus bullet_kind pr ;
+ set_bullets stk pr ;
+ b
+ | [] -> Util.anomaly "Tried to pop bullet from an empty stack"
+let push_bullet b pr =
+ Proof.focus bullet_cond 1 pr ;
+ set_bullets (b::get_bullets pr) pr
+
+let put_bullet p bul =
+ if has_bullet bul p then
+ begin
+ while bul <> pop_bullet p do () done;
+ push_bullet bul p
+ end
+ else
+ push_bullet bul p
+
+let vernac_solve n bullet tcom b =
if not (refining ()) then
error "Unknown command of the non proof-editing mode.";
- Decl_mode.check_not_proof_mode "Unknown proof instruction";
- begin
- if b then
- solve_nth n (Tacinterp.hide_interp tcom (get_end_tac ()))
- else solve_nth n (Tacinterp.hide_interp tcom None)
- end;
+ let p = Proof_global.give_me_the_proof () in
+ Option.iter (put_bullet p) bullet ;
+ solve_nth n (Tacinterp.hide_interp tcom None) ~with_end_tac:b;
(* in case a strict subtree was completed,
go back to the top of the prooftree *)
- if subtree_solved () then begin
- Flags.if_verbose msgnl (str "Subgoal proved");
- make_focus 0;
- reset_top_of_script ()
- end;
+ begin try while Proof.no_focused_goal p do
+ Proof.unfocus command_focus p
+ done
+ with Util.UserError _ -> () end;
print_subgoals();
if !pcoq <> None then (Option.get !pcoq).solve n
+
(* A command which should be a tactic. It has been
added by Christine to patch an error in the design of the proof
@@ -659,32 +673,6 @@ let vernac_set_end_tac tac =
if tac <> (Tacexpr.TacId []) then set_end_tac (Tacinterp.interp tac) else ()
(* TO DO verifier s'il faut pas mettre exist s | TacId s ici*)
-(***********************)
-(* Proof Language Mode *)
-
-let vernac_decl_proof () =
- check_not_proof_mode "Already in Proof Mode";
- if tree_solved () then
- error "Nothing left to prove here."
- else
- begin
- Decl_proof_instr.go_to_proof_mode ();
- print_subgoals ()
- end
-
-let vernac_return () =
- match get_current_mode () with
- Mode_tactic ->
- Decl_proof_instr.return_from_tactic_mode ();
- print_subgoals ()
- | Mode_proof ->
- error "\"return\" is only used after \"escape\"."
- | Mode_none ->
- error "There is no proof to end."
-
-let vernac_proof_instr instr =
- Decl_proof_instr.proof_instr instr;
- print_subgoals ()
(*****************************)
(* Auxiliary file management *)
@@ -825,14 +813,6 @@ let _ =
optread = Impargs.is_contextual_implicit_args;
optwrite = Impargs.make_contextual_implicit_args }
-(* let _ = *)
-(* declare_bool_option *)
-(* { optsync = true; *)
-(* optname = "forceable implicit arguments"; *)
-(* optkey = ["Forceable";"Implicit")); *)
-(* optread = Impargs.is_forceable_implicit_args; *)
-(* optwrite = Impargs.make_forceable_implicit_args } *)
-
let _ =
declare_bool_option
{ optsync = true;
@@ -1223,41 +1203,56 @@ let vernac_backtrack snum pnum naborts =
vernac_backto snum;
Pp.flush_all();
(* there may be no proof in progress, even if no abort *)
- (try print_subgoals () with UserError _ -> ())
+ (try print_subgoals () with Proof_global.NoCurrentProof | UserError _ -> ())
let vernac_focus gln =
- check_not_proof_mode "No focussing or Unfocussing in Proof Mode.";
+ let p = Proof_global.give_me_the_proof () in
match gln with
- | None -> traverse_nth_goal 1; print_subgoals ()
- | Some n -> traverse_nth_goal n; print_subgoals ()
+ | None -> Proof.focus focus_command_cond 1 p; print_subgoals ()
+ | Some n -> Proof.focus focus_command_cond n p; print_subgoals ()
+
- (* Reset the focus to the top of the tree *)
+ (* Unfocuses one step in the focus stack. *)
let vernac_unfocus () =
- check_not_proof_mode "No focussing or Unfocussing in Proof Mode.";
- make_focus 0; reset_top_of_script (); print_subgoals ()
-
-let vernac_go = function
- | GoTo n -> Pfedit.traverse n;show_node()
- | GoTop -> Pfedit.reset_top_of_tree ();show_node()
- | GoNext -> Pfedit.traverse_next_unproven ();show_node()
- | GoPrev -> Pfedit.traverse_prev_unproven ();show_node()
-
-let apply_subproof f occ =
- let pts = get_pftreestate() in
- let evc = evc_of_pftreestate pts in
- let rec aux pts = function
- | [] -> pts
- | (n::l) -> aux (Tacmach.traverse n pts) occ in
- let pts = aux pts (occ@[-1]) in
- let pf = proof_of_pftreestate pts in
- f evc (Global.named_context()) pf
+ let p = Proof_global.give_me_the_proof () in
+ Proof.unfocus command_focus p; print_subgoals ()
+
+(* BeginSubproof / EndSubproof.
+ BeginSubproof (vernac_subproof) focuses on the first goal, or the goal
+ given as argument.
+ EndSubproof (vernac_end_subproof) unfocuses from a BeginSubproof, provided
+ that the proof of the goal has been completed.
+*)
+let subproof_kind = Proof.new_focus_kind ()
+let subproof_cond = Proof.done_cond subproof_kind
+
+let vernac_subproof gln =
+ let p = Proof_global.give_me_the_proof () in
+ begin match gln with
+ | None -> Proof.focus subproof_cond 1 p
+ | Some n -> Proof.focus subproof_cond n p
+ end ;
+ print_subgoals ()
+
+let vernac_end_subproof () =
+ let p = Proof_global.give_me_the_proof () in
+ Proof.unfocus subproof_kind p ; print_subgoals ()
+
+let vernac_go _ =
+ (* spiwack: don't know what it's supposed to do. Undocumented.
+ Deactivated and candidate for removal. (Feb. 2010) *)
+ ()
let explain_proof occ =
- msg (apply_subproof (fun evd _ -> print_treescript evd) occ)
+ (* spiwack: don't know what it's supposed to do. Undocumented.
+ Deactivated and candidate for removal. (Feb. 2010) *)
+ ()
let explain_tree occ =
- msg (apply_subproof print_proof occ)
+ (* spiwack: don't know what it's supposed to do. Undocumented.
+ Deactivated and candidate for removeal. (Feb. 2010) *)
+ ()
let vernac_show = function
| ShowGoal nopt ->
@@ -1285,11 +1280,11 @@ let vernac_show = function
let vernac_check_guard () =
let pts = get_pftreestate () in
- let pf = proof_of_pftreestate pts in
- let (pfterm,_) = extract_open_pftreestate pts in
+ let pfterm = List.hd (Proof.partial_proof pts) in
let message =
try
- Inductiveops.control_only_guard (Evd.evar_env (goal_of_proof pf))
+ let { Evd.it=gl ; sigma=sigma } = Proof.V82.top_goal pts in
+ Inductiveops.control_only_guard (Goal.V82.env sigma gl)
pfterm;
(str "The condition holds up to here")
with UserError(_,s) ->
@@ -1352,17 +1347,9 @@ let interp c = match c with
| VernacDeclareClass id -> vernac_declare_class id
(* Solving *)
- | VernacSolve (n,tac,b) -> vernac_solve n tac b
+ | VernacSolve (n,bullet,tac,b) -> vernac_solve n bullet tac b
| VernacSolveExistential (n,c) -> vernac_solve_existential n c
- (* MMode *)
-
- | VernacDeclProof -> vernac_decl_proof ()
- | VernacReturn -> vernac_return ()
- | VernacProofInstr stp -> vernac_proof_instr stp
-
- (* /MMode *)
-
(* Auxiliary file and library management *)
| VernacRequireFrom (exp,spec,f) -> vernac_require_from exp spec f
| VernacAddLoadPath (isrec,s,alias) -> vernac_add_loadpath isrec s alias
@@ -1418,10 +1405,13 @@ let interp c = match c with
| VernacBacktrack (snum,pnum,naborts) -> vernac_backtrack snum pnum naborts
| VernacFocus n -> vernac_focus n
| VernacUnfocus -> vernac_unfocus ()
+ | VernacSubproof n -> vernac_subproof n
+ | VernacEndSubproof -> vernac_end_subproof ()
| VernacGo g -> vernac_go g
| VernacShow s -> vernac_show s
| VernacCheckGuard -> vernac_check_guard ()
| VernacProof tac -> vernac_set_end_tac tac
+ | VernacProofMode mn -> Proof_global.set_proof_mode mn
(* Toplevel control *)
| VernacToplevelControl e -> raise e
diff --git a/toplevel/vernacentries.mli b/toplevel/vernacentries.mli
index 44e8b7ab46..58df8a9065 100644
--- a/toplevel/vernacentries.mli
+++ b/toplevel/vernacentries.mli
@@ -55,3 +55,14 @@ val abort_refine : ('a -> unit) -> 'a -> unit;;
val interp : Vernacexpr.vernac_expr -> unit
val vernac_reset_name : identifier Util.located -> unit
+
+(* Print subgoals when the verbose flag is on. Meant to be used inside
+ vernac commands from plugins. *)
+val print_subgoals : unit -> unit
+
+
+(* Handles focusing/defocusing with bullets:
+ - If this bullet follows another one of its kind, defocuses then focuses
+ (which fails if the focused subproof is not complete).
+ - If it is the first bullet of its kind, then focuses a new subproof. *)
+val put_bullet : Proof.proof -> bullet -> unit
diff --git a/toplevel/vernacexpr.ml b/toplevel/vernacexpr.ml
index 0db000a485..e216f25203 100644
--- a/toplevel/vernacexpr.ml
+++ b/toplevel/vernacexpr.ml
@@ -190,6 +190,11 @@ type syntax_modifier =
| SetOnlyParsing
| SetFormat of string located
+type bullet =
+ | Dash
+ | Star
+ | Plus
+
type proof_end =
| Admitted
| Proved of opacity_flag * (lident * theorem_kind option) option
@@ -274,16 +279,9 @@ type vernac_expr =
(* Solving *)
- | VernacSolve of int * raw_tactic_expr * bool
+ | VernacSolve of int * bullet option * raw_tactic_expr * bool
| VernacSolveExistential of int * constr_expr
- (* Proof Mode *)
-
- | VernacDeclProof
- | VernacReturn
- | VernacProofInstr of Decl_expr.raw_proof_instr
-
-
(* Auxiliary file and library management *)
| VernacRequireFrom of export_flag option * specif_flag option * string
| VernacAddLoadPath of rec_flag * string * dir_path option
@@ -343,10 +341,13 @@ type vernac_expr =
| VernacBacktrack of int*int*int
| VernacFocus of int option
| VernacUnfocus
+ | VernacSubproof of int option
+ | VernacEndSubproof
| VernacGo of goable
| VernacShow of showable
| VernacCheckGuard
| VernacProof of raw_tactic_expr
+ | VernacProofMode of string
(* Toplevel control *)
| VernacToplevelControl of exn
diff --git a/toplevel/whelp.ml4 b/toplevel/whelp.ml4
index 98a79a9ce2..9d4b5ec0cf 100644
--- a/toplevel/whelp.ml4
+++ b/toplevel/whelp.ml4
@@ -196,7 +196,8 @@ let whelp_elim ind =
send_whelp "elim" (make_string uri_of_global (IndRef ind))
let on_goal f =
- let gls = nth_goal_of_pftreestate 1 (get_pftreestate ()) in
+ let { Evd.it=goals ; sigma=sigma } = Proof.V82.subgoals (get_pftreestate ()) in
+ let gls = { Evd.it=List.hd goals ; sigma = sigma } in
f (it_mkNamedProd_or_LetIn (pf_concl gls) (pf_hyps gls))
type whelp_request =