aboutsummaryrefslogtreecommitdiffhomepage
path: root/proofs/clenv.mli
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/clenv.mli
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/clenv.mli')
-rw-r--r--proofs/clenv.mli143
1 files changed, 143 insertions, 0 deletions
diff --git a/proofs/clenv.mli b/proofs/clenv.mli
new file mode 100644
index 000000000..2533fc537
--- /dev/null
+++ b/proofs/clenv.mli
@@ -0,0 +1,143 @@
+(************************************************************************)
+(* 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 Environ
+open Evd
+open Evarutil
+open Mod_subst
+open Rawterm
+open Unification
+(*i*)
+
+(***************************************************************)
+(* The Type of Constructions clausale environments. *)
+
+(* [env] is the typing context
+ * [evd] is the mapping from metavar and evar numbers to their types
+ * and values.
+ * [templval] is the template which we are trying to fill out.
+ * [templtyp] is its type.
+ *)
+type clausenv = {
+ env : env;
+ evd : evar_map;
+ templval : constr freelisted;
+ templtyp : constr freelisted }
+
+(* Substitution is not applied on templenv (because [subst_clenv] is
+ applied only on hints which typing env is overwritten by the
+ goal env) *)
+val subst_clenv : substitution -> clausenv -> clausenv
+
+(* subject of clenv (instantiated) *)
+val clenv_value : clausenv -> constr
+(* type of clenv (instantiated) *)
+val clenv_type : clausenv -> types
+(* substitute resolved metas *)
+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 : Goal.goal sigma -> constr * types -> clausenv
+val mk_clenv_from_n :
+ 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 : Goal.goal sigma -> clausenv -> clausenv
+val clenv_fchain :
+ ?allow_K:bool -> ?flags:unify_flags -> metavariable -> clausenv -> clausenv -> clausenv
+
+(***************************************************************)
+(* Unification with clenvs *)
+
+(* Unifies two terms in a clenv. The boolean is [allow_K] (see [Unification]) *)
+val clenv_unify :
+ bool -> ?flags:unify_flags -> conv_pb -> constr -> constr -> clausenv -> clausenv
+
+(* unifies the concl of the goal with the type of the clenv *)
+val clenv_unique_resolver :
+ 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 -> Goal.goal sigma -> clausenv
+
+val clenv_dependent : bool -> clausenv -> metavariable list
+
+val clenv_pose_metas_as_evars : clausenv -> metavariable list -> clausenv
+
+(***************************************************************)
+(* Bindings *)
+
+type arg_bindings = constr explicit_bindings
+
+(* bindings where the key is the position in the template of the
+ clenv (dependent or not). Positions can be negative meaning to
+ start from the rightmost argument of the template. *)
+val clenv_independent : clausenv -> metavariable list
+val clenv_missing : clausenv -> metavariable list
+
+val clenv_constrain_last_binding : constr -> clausenv -> clausenv
+
+(* defines metas corresponding to the name of the bindings *)
+val clenv_match_args : arg_bindings -> clausenv -> clausenv
+
+val clenv_unify_meta_types : ?flags:unify_flags -> clausenv -> clausenv
+
+(* start with a clenv to refine with a given term with bindings *)
+
+(* the arity of the lemma is fixed *)
+(* 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 :
+ Goal.goal sigma -> int option -> constr * constr -> constr bindings ->
+ clausenv
+val make_clenv_binding :
+ 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
+ it produces the unification pattern [ccl]; [sigma'] is [sigma]
+ extended with [lmetas]; if [n] is defined, it limits the size of
+ the list even if [ccl] is still a product; otherwise, it stops when
+ [ccl] is not a product; example: if [t] is [forall x y, x=y -> y=x]
+ and [n] is [None], then [lmetas] is [Meta n1;Meta n2;Meta n3] and
+ [ccl] is [Meta n1=Meta n2]; if [n] is [Some 1], [lmetas] is [Meta n1]
+ and [ccl] is [forall y, Meta n1=y -> y=Meta n1] *)
+val clenv_environments :
+ evar_map -> int option -> types -> evar_map * constr list * types
+
+(* [clenv_environments_evars env sigma n t] does the same but returns
+ a list of Evar's defined in [env] and extends [sigma] accordingly *)
+val clenv_environments_evars :
+ env -> evar_map -> int option -> types -> evar_map * constr list * types
+
+(* [clenv_conv_leq env sigma t c n] looks for c1...cn s.t. [t <= c c1...cn] *)
+val clenv_conv_leq :
+ env -> evar_map -> types -> constr -> int -> constr list
+
+(* if the clause is a product, add an extra meta for this product *)
+exception NotExtensibleClause
+val clenv_push_prod : clausenv -> clausenv
+
+(***************************************************************)
+(* Pretty-print *)
+val pr_clenv : clausenv -> Pp.std_ppcmds
+