aboutsummaryrefslogtreecommitdiffhomepage
path: root/proofs/proof_global.ml
diff options
context:
space:
mode:
authorGravatar aspiwack <aspiwack@85f007b7-540e-0410-9357-904b9bb8a0f7>2010-04-22 19:20:00 +0000
committerGravatar aspiwack <aspiwack@85f007b7-540e-0410-9357-904b9bb8a0f7>2010-04-22 19:20:00 +0000
commitaa99fc9ed78a0246d11d53dde502773a915b1022 (patch)
treed2ead3a9cf896fff6a49cfef72b6d5a52e928b41 /proofs/proof_global.ml
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
Diffstat (limited to 'proofs/proof_global.ml')
-rw-r--r--proofs/proof_global.ml295
1 files changed, 295 insertions, 0 deletions
diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml
new file mode 100644
index 000000000..420ff8432
--- /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
+