summaryrefslogtreecommitdiff
path: root/proofs
diff options
context:
space:
mode:
authorGravatar Enrico Tassi <gareuselesinge@debian.org>2015-01-25 14:42:51 +0100
committerGravatar Enrico Tassi <gareuselesinge@debian.org>2015-01-25 14:42:51 +0100
commit7cfc4e5146be5666419451bdd516f1f3f264d24a (patch)
treee4197645da03dc3c7cc84e434cc31d0a0cca7056 /proofs
parent420f78b2caeaaddc6fe484565b2d0e49c66888e5 (diff)
Imported Upstream version 8.5~beta1+dfsg
Diffstat (limited to 'proofs')
-rw-r--r--proofs/clenv.ml295
-rw-r--r--proofs/clenv.mli108
-rw-r--r--proofs/clenvtac.ml81
-rw-r--r--proofs/clenvtac.mli16
-rw-r--r--proofs/evar_refiner.ml58
-rw-r--r--proofs/evar_refiner.mli9
-rw-r--r--proofs/goal.ml579
-rw-r--r--proofs/goal.mli193
-rw-r--r--proofs/logic.ml515
-rw-r--r--proofs/logic.mli16
-rw-r--r--proofs/logic_monad.ml326
-rw-r--r--proofs/logic_monad.mli157
-rw-r--r--proofs/pfedit.ml179
-rw-r--r--proofs/pfedit.mli110
-rw-r--r--proofs/proof.ml507
-rw-r--r--proofs/proof.mli131
-rw-r--r--proofs/proof_global.ml662
-rw-r--r--proofs/proof_global.mli154
-rw-r--r--proofs/proof_type.ml91
-rw-r--r--proofs/proof_type.mli104
-rw-r--r--proofs/proof_using.ml166
-rw-r--r--proofs/proof_using.mli26
-rw-r--r--proofs/proofs.mllib10
-rw-r--r--proofs/proofview.ml1502
-rw-r--r--proofs/proofview.mli634
-rw-r--r--proofs/proofview_monad.ml270
-rw-r--r--proofs/proofview_monad.mli144
-rw-r--r--proofs/redexpr.ml156
-rw-r--r--proofs/redexpr.mli12
-rw-r--r--proofs/refiner.ml270
-rw-r--r--proofs/refiner.mli54
-rw-r--r--proofs/tacexpr.ml345
-rw-r--r--proofs/tacmach.ml185
-rw-r--r--proofs/tacmach.mli130
-rw-r--r--proofs/tactic_debug.ml271
-rw-r--r--proofs/tactic_debug.mli40
36 files changed, 4928 insertions, 3578 deletions
diff --git a/proofs/clenv.ml b/proofs/clenv.ml
index 8c66269e..2c9c695b 100644
--- a/proofs/clenv.ml
+++ b/proofs/clenv.ml
@@ -1,31 +1,29 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Pp
+open Errors
open Util
open Names
open Nameops
open Term
+open Vars
open Termops
open Namegen
-open Sign
open Environ
open Evd
open Reduction
open Reductionops
-open Glob_term
-open Pattern
open Tacred
open Pretype_errors
open Evarutil
open Unification
-open Mod_subst
-open Coercion.Default
+open Misctypes
(* Abbreviations *)
@@ -44,10 +42,10 @@ type clausenv = {
let cl_env ce = ce.env
let cl_sigma ce = ce.evd
-let subst_clenv sub clenv =
- { templval = map_fl (subst_mps sub) clenv.templval;
- templtyp = map_fl (subst_mps sub) clenv.templtyp;
- evd = subst_evar_defs_light sub clenv.evd;
+let map_clenv sub clenv =
+ { templval = map_fl sub clenv.templval;
+ templtyp = map_fl sub clenv.templtyp;
+ evd = cmap sub clenv.evd;
env = clenv.env }
let clenv_nf_meta clenv c = nf_meta clenv.evd c
@@ -56,6 +54,15 @@ let clenv_meta_type clenv mv = Typing.meta_type clenv.evd mv
let clenv_value clenv = meta_instance clenv.evd clenv.templval
let clenv_type clenv = meta_instance clenv.evd clenv.templtyp
+let refresh_undefined_univs clenv =
+ match kind_of_term clenv.templval.rebus with
+ | Var _ -> clenv, Univ.empty_level_subst
+ | App (f, args) when isVar f -> clenv, Univ.empty_level_subst
+ | _ ->
+ let evd', subst = Evd.refresh_undefined_universes clenv.evd in
+ let map_freelisted f = { f with rebus = subst_univs_level_constr subst f.rebus } in
+ { clenv with evd = evd'; templval = map_freelisted clenv.templval;
+ templtyp = map_freelisted clenv.templtyp }, subst
let clenv_hnf_constr ce t = hnf_constr (cl_env ce) (cl_sigma ce) t
@@ -84,6 +91,16 @@ let clenv_push_prod cl =
(* Instantiate the first [bound] products of [t] with metas (all products if
[bound] is [None]; unfold local defs *)
+(** [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] *)
+
let clenv_environments evd bound t =
let rec clrec (e,metas) n t =
match n, kind_of_term t with
@@ -101,41 +118,13 @@ let clenv_environments evd bound t =
in
clrec (evd,[]) bound t
-(* Instantiate the first [bound] products of [t] with evars (all products if
- [bound] is [None]; unfold local defs *)
-
-let clenv_environments_evars env evd bound t =
- let rec clrec (e,ts) n t =
- match n, kind_of_term t with
- | (Some 0, _) -> (e, List.rev ts, t)
- | (n, Cast (t,_,_)) -> clrec (e,ts) n t
- | (n, Prod (na,t1,t2)) ->
- let e',constr = Evarutil.new_evar e env t1 in
- let dep = dependent (mkRel 1) t2 in
- clrec (e', constr::ts) (Option.map ((+) (-1)) n)
- (if dep then (subst1 constr t2) else t2)
- | (n, LetIn (na,b,_,t)) -> clrec (e,ts) n (subst1 b t)
- | (n, _) -> (e, List.rev ts, t)
- in
- clrec (evd,[]) bound t
-
-let clenv_conv_leq env sigma t c bound =
- let ty = Retyping.get_type_of env sigma c in
- let evd = Evd.create_goal_evar_defs sigma in
- let evars,args,_ = clenv_environments_evars env evd (Some bound) ty in
- let evars = Evarconv.the_conv_x_leq env t (applist (c,args)) evars in
- let evars = Evarconv.consider_remaining_unif_problems env evars in
- let args = List.map (whd_evar evars) args in
- check_evars env sigma evars (applist (c,args));
- args
-
-let mk_clenv_from_env environ sigma n (c,cty) =
+let mk_clenv_from_env env sigma n (c,cty) =
let evd = create_goal_evar_defs sigma in
let (evd,args,concl) = clenv_environments evd n cty in
- { templval = mk_freelisted (match args with [] -> c | _ -> applist (c,args));
+ { templval = mk_freelisted (applist (c,args));
templtyp = mk_freelisted concl;
evd = evd;
- env = environ }
+ env = env }
let mk_clenv_from_n gls n (c,cty) =
mk_clenv_from_env (pf_env gls) gls.sigma n (c, cty)
@@ -152,13 +141,13 @@ let mk_clenv_type_of gls t = mk_clenv_from gls (t,pf_type_of gls t)
let mentions clenv mv0 =
let rec menrec mv1 =
- mv0 = mv1 ||
+ Int.equal mv0 mv1 ||
let mlist =
try match meta_opt_fvalue clenv.evd mv1 with
| Some (b,_) -> b.freemetas
| None -> Metaset.empty
with Not_found -> Metaset.empty in
- meta_exists menrec mlist
+ Metaset.exists menrec mlist
in menrec
let error_incompatible_inst clenv mv =
@@ -169,16 +158,16 @@ let error_incompatible_inst clenv mv =
(str "An incompatible instantiation has already been found for " ++
pr_id id)
| _ ->
- anomaly "clenv_assign: non dependent metavar already assigned"
+ anomaly ~label:"clenv_assign" (Pp.str "non dependent metavar already assigned")
(* TODO: replace by clenv_unify (mkMeta mv) rhs ? *)
let clenv_assign mv rhs clenv =
let rhs_fls = mk_freelisted rhs in
- if meta_exists (mentions clenv mv) rhs_fls.freemetas then
+ if Metaset.exists (mentions clenv mv) rhs_fls.freemetas then
error "clenv_assign: circularity in unification";
try
if meta_defined clenv.evd mv then
- if not (eq_constr (fst (meta_fvalue clenv.evd mv)).rebus rhs) then
+ if not (Term.eq_constr (fst (meta_fvalue clenv.evd mv)).rebus rhs) then
error_incompatible_inst clenv mv
else
clenv
@@ -265,22 +254,49 @@ let clenv_dependent ce = clenv_dependent_gen false ce
(******************************************************************)
-let clenv_unify ?(flags=default_unify_flags) cv_pb t1 t2 clenv =
+let clenv_unify ?(flags=default_unify_flags ()) cv_pb t1 t2 clenv =
{ clenv with
evd = w_unify ~flags clenv.env clenv.evd cv_pb t1 t2 }
-let clenv_unify_meta_types ?(flags=default_unify_flags) clenv =
+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 ?(flags=default_unify_flags) clenv gl =
+let clenv_unique_resolver ?(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
+ if isMeta (fst (decompose_appvect (whd_nored clenv.evd clenv.templtyp.rebus))) then
clenv_unify CUMUL ~flags (clenv_type clenv) concl
(clenv_unify_meta_types ~flags clenv)
else
clenv_unify CUMUL ~flags
(meta_reducible_instance clenv.evd clenv.templtyp) concl clenv
+let adjust_meta_source evd mv = function
+ | loc,Evar_kinds.VarInstance id ->
+ let rec match_name c l =
+ match kind_of_term c, l with
+ | Lambda (Name id,_,c), a::l when Constr.equal a (mkMeta mv) -> Some id
+ | Lambda (_,_,c), a::l -> match_name c l
+ | _ -> None in
+ (* This is very ad hoc code so that an evar inherits the name of the binder
+ in situations like "ex_intro (fun x => P) ?ev p" *)
+ let f = function (mv',(Cltyp (_,t) | Clval (_,_,t))) ->
+ if Metaset.mem mv t.freemetas then
+ let f,l = decompose_app t.rebus in
+ match kind_of_term f with
+ | Meta mv'' ->
+ (match meta_opt_fvalue evd mv'' with
+ | Some (c,_) -> match_name c.rebus l
+ | None -> None)
+ | Evar ev ->
+ (match existential_opt_value evd ev with
+ | Some c -> match_name c l
+ | None -> None)
+ | _ -> None
+ else None in
+ let id = try List.find_map f (Evd.meta_list evd) with Not_found -> id in
+ loc,Evar_kinds.VarInstance id
+ | src -> src
+
(* [clenv_pose_metas_as_evars clenv dep_mvs]
* For each dependent evar in the clause-env which does not have a value,
* pose a value for it by constructing a fresh evar. We do this in
@@ -317,14 +333,13 @@ let clenv_pose_metas_as_evars clenv dep_mvs =
(* This assumes no cycle in the dependencies - is it correct ? *)
if occur_meta ty then fold clenv (mvs@[mv])
else
- let (evd,evar) =
- new_evar clenv.evd (cl_env clenv) ~src:(dummy_loc,GoalEvar) ty in
+ let src = evar_source_of_meta mv clenv.evd in
+ let src = adjust_meta_source clenv.evd mv src in
+ let (evd,evar) = new_evar (cl_env clenv) clenv.evd ~src ty in
let clenv = clenv_assign mv evar {clenv with evd=evd} in
fold clenv mvs in
fold clenv dep_mvs
-let evar_clenv_unique_resolver = clenv_unique_resolver
-
(******************************************************************)
let connect_clenv gls clenv =
@@ -333,6 +348,9 @@ let connect_clenv gls clenv =
evd = evd ;
env = Goal.V82.env evd (sig_it gls) }
+(* let connect_clenv_key = Profile.declare_profile "connect_clenv";; *)
+(* let connect_clenv = Profile.profile2 connect_clenv_key connect_clenv *)
+
(* [clenv_fchain mv clenv clenv']
*
* Resolves the value of "mv" (which must be undefined) in clenv to be
@@ -357,11 +375,11 @@ let connect_clenv gls clenv =
In particular, it assumes that [env'] and [sigma'] extend [env] and [sigma].
*)
-let fchain_flags =
- { default_unify_flags with
+let fchain_flags () =
+ { (default_unify_flags ()) with
allow_K_in_toplevel_higher_order_unification = true }
-let clenv_fchain ?(flags=fchain_flags) mv clenv nextclenv =
+let clenv_fchain ?(flags=fchain_flags ()) mv clenv nextclenv =
(* Add the metavars of [nextclenv] to [clenv], with their name-environment *)
let clenv' =
{ templval = clenv.templval;
@@ -370,7 +388,7 @@ let clenv_fchain ?(flags=fchain_flags) mv clenv nextclenv =
env = nextclenv.env } in
(* unify the type of the template of [nextclenv] with the type of [mv] *)
let clenv'' =
- clenv_unify ~flags:flags CUMUL
+ clenv_unify ~flags CUMUL
(clenv_term clenv' nextclenv.templtyp)
(clenv_meta_type clenv' mv)
clenv' in
@@ -397,8 +415,13 @@ let clenv_independent clenv =
let deps = Metaset.union (dependent_in_type_of_metas clenv mvs) ctyp_mvs in
List.filter (fun mv -> not (Metaset.mem mv deps)) mvs
+let qhyp_eq h1 h2 = match h1, h2 with
+| NamedHyp n1, NamedHyp n2 -> Id.equal n1 n2
+| AnonHyp i1, AnonHyp i2 -> Int.equal i1 i2
+| _ -> false
+
let check_bindings bl =
- match list_duplicates (List.map pi2 bl) with
+ match List.duplicates qhyp_eq (List.map pi2 bl) with
| NamedHyp s :: _ ->
errorlabstrm ""
(str "The variable " ++ pr_id s ++
@@ -423,11 +446,11 @@ let error_already_defined b =
(str "Binder name \"" ++ pr_id id ++
str"\" already defined with incompatible value.")
| AnonHyp n ->
- anomalylabstrm ""
+ anomaly
(str "Position " ++ int n ++ str" already defined.")
let clenv_unify_binding_type clenv c t u =
- if isMeta (fst (whd_stack clenv.evd u)) then
+ if isMeta (fst (decompose_appvect (whd_nored clenv.evd u))) then
(* Not enough information to know if some subtyping is needed *)
CoerceToType, clenv, c
else
@@ -436,7 +459,8 @@ let clenv_unify_binding_type clenv c t u =
let evd,c = w_coerce_to_type (cl_env clenv) clenv.evd c t u in
TypeProcessed, { clenv with evd = evd }, c
with
- | PretypeError (_,_,NotClean _) as e -> raise e
+ | PretypeError (_,_,ActualTypeNotCoercible (_,_,NotClean _)) as e ->
+ raise e
| e when precatchable_exception e ->
TypeNotProcessed, clenv, c
@@ -447,7 +471,7 @@ let clenv_assign_binding clenv k c =
{ clenv' with evd = meta_assign k (c,(Conv,status)) clenv'.evd }
let clenv_match_args bl clenv =
- if bl = [] then
+ if List.is_empty bl then
clenv
else
let mvs = clenv_independent clenv in
@@ -456,7 +480,7 @@ let clenv_match_args bl clenv =
(fun clenv (loc,b,c) ->
let k = meta_of_binder clenv loc mvs b in
if meta_defined clenv.evd k then
- if eq_constr (fst (meta_fvalue clenv.evd k)).rebus c then clenv
+ if Term.eq_constr (fst (meta_fvalue clenv.evd k)).rebus c then clenv
else error_already_defined b
else
clenv_assign_binding clenv k c)
@@ -466,7 +490,7 @@ exception NoSuchBinding
let clenv_constrain_last_binding c clenv =
let all_mvs = collect_metas clenv.templval.rebus in
- let k = try list_last all_mvs with Failure _ -> raise NoSuchBinding in
+ let k = try List.last all_mvs with Failure _ -> raise NoSuchBinding in
clenv_assign_binding clenv k c
let error_not_right_number_missing_arguments n =
@@ -475,17 +499,17 @@ let error_not_right_number_missing_arguments n =
int n ++ str ").")
let clenv_constrain_dep_args hyps_only bl clenv =
- if bl = [] then
+ if List.is_empty bl then
clenv
else
let occlist = clenv_dependent_gen hyps_only clenv in
- if List.length occlist = List.length bl then
+ if Int.equal (List.length occlist) (List.length bl) then
List.fold_left2 clenv_assign_binding clenv occlist bl
else
if hyps_only then
(* Tolerance for compatibility <= 8.3 *)
let occlist' = clenv_dependent_gen hyps_only ~iter:false clenv in
- if List.length occlist' = List.length bl then
+ if Int.equal (List.length occlist') (List.length bl) then
List.fold_left2 clenv_assign_binding clenv occlist' bl
else
error_not_right_number_missing_arguments (List.length occlist)
@@ -508,9 +532,12 @@ let make_clenv_binding_gen hyps_only n env sigma (c,t) = function
let make_clenv_binding_env_apply env sigma n =
make_clenv_binding_gen true n env sigma
+
+let make_clenv_binding_env env sigma =
+ make_clenv_binding_gen false None env sigma
-let make_clenv_binding_apply gls n = make_clenv_binding_gen true n (pf_env gls) gls.sigma
-let make_clenv_binding gls = make_clenv_binding_gen false None (pf_env gls) gls.sigma
+let make_clenv_binding_apply env sigma n = make_clenv_binding_gen true n env sigma
+let make_clenv_binding env sigma = make_clenv_binding_gen false None env sigma
(****************************************************************)
(* Pretty-print *)
@@ -520,3 +547,131 @@ let pr_clenv clenv =
(str"TEMPL: " ++ print_constr clenv.templval.rebus ++
str" : " ++ print_constr clenv.templtyp.rebus ++ fnl () ++
pr_evar_map (Some 2) clenv.evd)
+
+(****************************************************************)
+(** Evar version of mk_clenv *)
+
+type hole = {
+ hole_evar : constr;
+ hole_type : types;
+ hole_deps : bool;
+ hole_name : Name.t;
+}
+
+type clause = {
+ cl_holes : hole list;
+ cl_concl : types;
+}
+
+let make_evar_clause env sigma ?len t =
+ let bound = match len with
+ | None -> -1
+ | Some n -> assert (0 <= n); n
+ in
+ (** FIXME: do the renaming online *)
+ let t = rename_bound_vars_as_displayed [] [] t in
+ let rec clrec (sigma, holes) n t =
+ if n = 0 then (sigma, holes, t)
+ else match kind_of_term t with
+ | Cast (t, _, _) -> clrec (sigma, holes) n t
+ | Prod (na, t1, t2) ->
+ let store = Typeclasses.set_resolvable Evd.Store.empty false in
+ let sigma, ev = new_evar ~store env sigma t1 in
+ let dep = dependent (mkRel 1) t2 in
+ let hole = {
+ hole_evar = ev;
+ hole_type = t1;
+ hole_deps = dep;
+ (* We fix it later *)
+ hole_name = na;
+ } in
+ let t2 = if dep then subst1 ev t2 else t2 in
+ clrec (sigma, hole :: holes) (pred n) t2
+ | LetIn (na, b, _, t) -> clrec (sigma, holes) n (subst1 b t)
+ | _ -> (sigma, holes, t)
+ in
+ let (sigma, holes, t) = clrec (sigma, []) bound t in
+ let holes = List.rev holes in
+ let clause = { cl_concl = t; cl_holes = holes } in
+ (sigma, clause)
+
+let explain_no_such_bound_variable holes id =
+ let fold h accu = match h.hole_name with
+ | Anonymous -> accu
+ | Name id -> id :: accu
+ in
+ let mvl = List.fold_right fold holes [] in
+ let expl = match mvl with
+ | [] -> str " (no bound variables at all in the expression)."
+ | [id] -> str "(possible name is: " ++ pr_id id ++ str ")."
+ | _ -> str "(possible names are: " ++ pr_enum pr_id mvl ++ str ")."
+ in
+ errorlabstrm "" (str "No such bound variable " ++ pr_id id ++ expl)
+
+let evar_with_name holes id =
+ let map h = match h.hole_name with
+ | Anonymous -> None
+ | Name id' -> if Id.equal id id' then Some h else None
+ in
+ let hole = List.map_filter map holes in
+ match hole with
+ | [] -> explain_no_such_bound_variable holes id
+ | [h] -> h.hole_evar
+ | _ ->
+ errorlabstrm ""
+ (str "Binder name \"" ++ pr_id id ++
+ str "\" occurs more than once in clause.")
+
+let evar_of_binder holes = function
+| NamedHyp s -> evar_with_name holes s
+| AnonHyp n ->
+ try
+ let h = List.nth holes (pred n) in
+ h.hole_evar
+ with e when Errors.noncritical e ->
+ errorlabstrm "" (str "No such binder.")
+
+let define_with_type sigma env ev c =
+ let t = Retyping.get_type_of env sigma ev in
+ let ty = Retyping.get_type_of env sigma c in
+ let j = Environ.make_judge c ty in
+ let (sigma, j) = Coercion.inh_conv_coerce_to true (Loc.ghost) env sigma j t in
+ let (ev, _) = destEvar ev in
+ let sigma = Evd.define ev j.Environ.uj_val sigma in
+ sigma
+
+let solve_evar_clause env sigma hyp_only clause = function
+| NoBindings -> sigma
+| ImplicitBindings largs ->
+ let fold holes h =
+ if h.hole_deps then
+ (** Some subsequent term uses the hole *)
+ let (ev, _) = destEvar h.hole_evar in
+ let is_dep hole = occur_evar ev hole.hole_type in
+ let in_hyp = List.exists is_dep holes in
+ let in_ccl = occur_evar ev clause.cl_concl in
+ let dep = if hyp_only then in_hyp && not in_ccl else in_hyp || in_ccl in
+ let h = { h with hole_deps = dep } in
+ h :: holes
+ else
+ (** The hole does not occur anywhere *)
+ h :: holes
+ in
+ let holes = List.fold_left fold [] (List.rev clause.cl_holes) in
+ let map h = if h.hole_deps then Some h.hole_evar else None in
+ let evs = List.map_filter map holes in
+ let len = List.length evs in
+ if Int.equal len (List.length largs) then
+ let fold sigma ev arg = define_with_type sigma env ev arg in
+ let sigma = List.fold_left2 fold sigma evs largs in
+ sigma
+ else
+ error_not_right_number_missing_arguments len
+| ExplicitBindings lbind ->
+ let () = check_bindings lbind in
+ let fold sigma (_, binder, c) =
+ let ev = evar_of_binder clause.cl_holes binder in
+ define_with_type sigma env ev c
+ in
+ let sigma = List.fold_left fold sigma lbind in
+ sigma
diff --git a/proofs/clenv.mli b/proofs/clenv.mli
index f7817611..9b671bcf 100644
--- a/proofs/clenv.mli
+++ b/proofs/clenv.mli
@@ -1,21 +1,18 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Util
open Names
open Term
-open Sign
open Environ
open Evd
-open Evarutil
open Mod_subst
-open Glob_term
open Unification
+open Misctypes
(** {6 The Type of Constructions clausale environments.} *)
@@ -27,10 +24,8 @@ type clausenv = {
out *)
templtyp : constr freelisted (** its type *)}
-(** 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
+
+val map_clenv : (constr -> constr) -> clausenv -> clausenv
(** subject of clenv (instantiated) *)
val clenv_value : clausenv -> constr
@@ -50,6 +45,9 @@ val mk_clenv_from_n :
val mk_clenv_type_of : Goal.goal sigma -> constr -> clausenv
val mk_clenv_from_env : env -> evar_map -> int option -> constr * types -> clausenv
+(** Refresh the universes in a clenv *)
+val refresh_undefined_univs : clausenv -> clausenv * Univ.universe_level_subst
+
(** {6 linking of clenvs } *)
val connect_clenv : Goal.goal sigma -> clausenv -> clausenv
@@ -66,11 +64,6 @@ val clenv_unify :
val clenv_unique_resolver :
?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 :
- ?flags:unify_flags -> clausenv -> Goal.goal sigma -> clausenv
-
val clenv_dependent : clausenv -> metavariable list
val clenv_pose_metas_as_evars : clausenv -> metavariable list -> clausenv
@@ -89,9 +82,6 @@ val clenv_missing : clausenv -> metavariable list
exception NoSuchBinding
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 *)
@@ -104,31 +94,14 @@ val make_clenv_binding_env_apply :
clausenv
val make_clenv_binding_apply :
- Goal.goal sigma -> int option -> constr * constr -> constr bindings ->
+ env -> evar_map -> int option -> constr * constr -> constr bindings ->
clausenv
+
+val make_clenv_binding_env :
+ env -> evar_map -> 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
+ env -> evar_map -> constr * constr -> constr bindings -> clausenv
(** if the clause is a product, add an extra meta for this product *)
exception NotExtensibleClause
@@ -137,3 +110,58 @@ val clenv_push_prod : clausenv -> clausenv
(** {6 Pretty-print (debug only) } *)
val pr_clenv : clausenv -> Pp.std_ppcmds
+(** {6 Evar-based clauses} *)
+
+(** The following code is an adaptation of the [make_clenv_*] functions above,
+ except that it uses evars instead of metas, and naturally fits in the new
+ refinement monad. It should eventually replace all uses of the
+ aforementioned functions.
+
+ A clause is constructed as follows: assume a type [t := forall (x1 : A1) ...
+ (xn : An), T], we instantiate all the [xi] with a fresh evar [ei] and
+ return [T(xi := ei)] together with the [ei] enriched with a bit of
+ additional data. This is the simple part done by [make_evar_clause].
+
+ The problem lies in the fact we want to solve such holes with some
+ [constr bindings]. This entails some subtleties, because the provided terms
+ may only be well-typed up to a coercion, which we can only infer if we have
+ enough typing information. The meta machinery could insert coercions through
+ tricky instantiation delays. The only solution we have now is to delay the
+ tentative resolution of clauses by providing the [solve_evar_clause]
+ function, to be called at a smart enough time.
+*)
+
+type hole = {
+ hole_evar : constr;
+ (** The hole itself. Guaranteed to be an evar. *)
+ hole_type : types;
+ (** Type of the hole in the current environment. *)
+ hole_deps : bool;
+ (** Whether the remainder of the clause was dependent in the hole. Note that
+ because let binders are substituted, it does not mean that it actually
+ appears somewhere in the returned clause. *)
+ hole_name : Name.t;
+ (** Name of the hole coming from its binder. *)
+}
+
+type clause = {
+ cl_holes : hole list;
+ cl_concl : types;
+}
+
+val make_evar_clause : env -> evar_map -> ?len:int -> types ->
+ (evar_map * clause)
+(** An evar version of {!make_clenv_binding}. Given a type [t],
+ [evar_environments env sigma ~len t bl] tries to eliminate at most [len]
+ products of the type [t] by filling it with evars. It returns the resulting
+ type together with the list of holes generated. Assumes that [t] is
+ well-typed in the environment. *)
+
+val solve_evar_clause : env -> evar_map -> bool -> clause -> constr bindings ->
+ evar_map
+(** [solve_evar_clause env sigma hyps cl bl] tries to solve the holes contained
+ in [cl] according to the [bl] argument. Assumes that [bl] are well-typed in
+ the environment. The boolean [hyps] is a compatibility flag that allows to
+ consider arguments to be dependent only when they appear in hypotheses and
+ not in the conclusion. This boolean is only used when [bl] is of the form
+ [ImplicitBindings _]. *)
diff --git a/proofs/clenvtac.ml b/proofs/clenvtac.ml
index b54e2323..18883df2 100644
--- a/proofs/clenvtac.ml
+++ b/proofs/clenvtac.ml
@@ -1,30 +1,20 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Pp
open Util
open Names
-open Nameops
open Term
open Termops
-open Sign
-open Environ
open Evd
-open Evarutil
-open Proof_type
open Refiner
open Logic
open Reduction
-open Reductionops
open Tacmach
-open Glob_term
-open Pattern
-open Tacexpr
open Clenv
@@ -39,6 +29,7 @@ let clenv_cast_meta clenv =
match kind_of_term u with
| App _ | Case _ -> crec_hd u
| Cast (c,_,_) when isMeta c -> u
+ | Proj (p, c) -> mkProj (p, crec_hd c)
| _ -> map_constr crec u
and crec_hd u =
@@ -53,6 +44,7 @@ let clenv_cast_meta clenv =
| App(f,args) -> mkApp (crec_hd f, Array.map crec args)
| Case(ci,p,c,br) ->
mkCase (ci, crec_hd p, crec_hd c, Array.map crec br)
+ | Proj (p, c) -> mkProj (p, crec_hd c)
| _ -> u
in
crec
@@ -62,67 +54,78 @@ let clenv_value_cast_meta clenv =
let clenv_pose_dependent_evars with_evars clenv =
let dep_mvs = clenv_dependent clenv in
- if dep_mvs <> [] & not with_evars then
+ if not (List.is_empty dep_mvs) && not with_evars then
raise
(RefinerError (UnresolvedBindings (List.map (meta_name clenv.evd) dep_mvs)));
clenv_pose_metas_as_evars clenv dep_mvs
-let clenv_refine with_evars ?(with_classes=true) clenv gls =
+let clenv_refine with_evars ?(with_classes=true) clenv =
+ (** ppedrot: a Goal.enter here breaks things, because the tactic below may
+ solve goals by side effects, while the compatibility layer keeps those
+ useless goals. That deserves a FIXME. *)
+ Proofview.V82.tactic begin fun gl ->
let clenv = clenv_pose_dependent_evars with_evars clenv in
let evd' =
if with_classes then
- Typeclasses.resolve_typeclasses ~filter:Typeclasses.all_evars
+ let evd' = Typeclasses.resolve_typeclasses ~filter:Typeclasses.all_evars
~fail:(not with_evars) clenv.env clenv.evd
+ in Typeclasses.mark_unresolvables ~filter:Typeclasses.all_goals evd'
else clenv.evd
in
let clenv = { clenv with evd = evd' } in
tclTHEN
- (tclEVARS evd')
- (refine (clenv_cast_meta clenv (clenv_value clenv)))
- gls
+ (tclEVARS (Evd.clear_metas evd'))
+ (refine_no_check (clenv_cast_meta clenv (clenv_value clenv))) gl
+ end
open Unification
let dft = default_unify_flags
-let res_pf clenv ?(with_evars=false) ?(flags=dft) gls =
- clenv_refine with_evars (clenv_unique_resolver ~flags clenv gls) gls
-
-let elim_res_pf_THEN_i clenv tac gls =
- let clenv' = (clenv_unique_resolver ~flags:elim_flags clenv gls) in
- tclTHENLASTn (clenv_refine false clenv') (tac clenv') gls
-
-let e_res_pf clenv = res_pf clenv ~with_evars:true ~flags:dft
-
+let res_pf ?(with_evars=false) ?(with_classes=true) ?(flags=dft ()) clenv =
+ Proofview.Goal.enter begin fun gl ->
+ let clenv gl = clenv_unique_resolver ~flags clenv gl in
+ clenv_refine with_evars ~with_classes (Tacmach.New.of_old clenv (Proofview.Goal.assume gl))
+ end
(* [unifyTerms] et [unify] ne semble pas gérer les Meta, en
particulier ne semblent pas vérifier que des instances différentes
d'une même Meta sont compatibles. D'ailleurs le "fst" jette les metas
provenant de w_Unify. (Utilisé seulement dans prolog.ml) *)
-let fail_quick_unif_flags = {
+let fail_quick_core_unif_flags = {
modulo_conv_on_closed_terms = Some full_transparent_state;
use_metas_eagerly_in_conv_on_closed_terms = false;
+ use_evars_eagerly_in_conv_on_closed_terms = false;
modulo_delta = empty_transparent_state;
modulo_delta_types = full_transparent_state;
- modulo_delta_in_merge = None;
check_applied_meta_types = false;
- resolve_evars = false;
use_pattern_unification = false;
use_meta_bound_pattern_unification = true; (* ? *)
- frozen_evars = ExistentialSet.empty;
+ frozen_evars = Evar.Set.empty;
restrict_conv_on_strict_subterms = false; (* ? *)
modulo_betaiota = false;
modulo_eta = true;
- allow_K_in_toplevel_higher_order_unification = false
}
-(* let unifyTerms m n = walking (fun wc -> fst (w_Unify CONV m n [] wc)) *)
-let unifyTerms ?(flags=fail_quick_unif_flags) m n gls =
- let env = pf_env gls in
- let evd = create_goal_evar_defs (project gls) in
- let evd' = w_unify env evd CONV ~flags m n in
- tclIDTAC {it = gls.it; sigma = evd'}
+let fail_quick_unif_flags = {
+ core_unify_flags = fail_quick_core_unif_flags;
+ merge_unify_flags = fail_quick_core_unif_flags;
+ subterm_unify_flags = fail_quick_core_unif_flags;
+ allow_K_in_toplevel_higher_order_unification = false;
+ resolve_evars = false
+}
-let unify ?(flags=fail_quick_unif_flags) m gls =
- let n = pf_concl gls in unifyTerms ~flags m n gls
+(* let unifyTerms m n = walking (fun wc -> fst (w_Unify CONV m n [] wc)) *)
+let unify ?(flags=fail_quick_unif_flags) m =
+ Proofview.Goal.enter begin fun gl ->
+ let env = Tacmach.New.pf_env gl in
+ let n = Tacmach.New.pf_nf_concl gl in
+ let evd = create_goal_evar_defs (Proofview.Goal.sigma gl) in
+ try
+ let evd' = w_unify env evd CONV ~flags m n in
+ Proofview.Unsafe.tclEVARSADVANCE evd'
+ with e when Errors.noncritical e ->
+ (** This is Tacticals.tclFAIL *)
+ Proofview.tclZERO (FailError (0, lazy (Errors.print e)))
+ end
diff --git a/proofs/clenvtac.mli b/proofs/clenvtac.mli
index 4ed18d2e..da40427c 100644
--- a/proofs/clenvtac.mli
+++ b/proofs/clenvtac.mli
@@ -1,29 +1,21 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Util
-open Names
open Term
-open Sign
-open Evd
open Clenv
open Proof_type
open Tacexpr
open Unification
(** Tactics *)
-val unify : ?flags:unify_flags -> constr -> tactic
-val clenv_refine : evars_flag -> ?with_classes:bool -> clausenv -> tactic
-val res_pf : clausenv -> ?with_evars:evars_flag -> ?flags:unify_flags -> tactic
-val elim_res_pf_THEN_i : clausenv -> (clausenv -> tactic array) -> tactic
+val unify : ?flags:unify_flags -> constr -> unit Proofview.tactic
+val clenv_refine : evars_flag -> ?with_classes:bool -> clausenv -> unit Proofview.tactic
+val res_pf : ?with_evars:evars_flag -> ?with_classes:bool -> ?flags:unify_flags -> clausenv -> unit Proofview.tactic
val clenv_pose_dependent_evars : evars_flag -> clausenv -> clausenv
val clenv_value_cast_meta : clausenv -> constr
-
-(** Compatibility, use res_pf ?with_evars:true instead *)
-val e_res_pf : clausenv -> tactic
diff --git a/proofs/evar_refiner.ml b/proofs/evar_refiner.ml
index 6548a1dd..c8cb1d1c 100644
--- a/proofs/evar_refiner.ml
+++ b/proofs/evar_refiner.ml
@@ -1,61 +1,71 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Errors
open Util
open Names
-open Term
open Evd
open Evarutil
-open Sign
-open Refiner
+open Evarsolve
(******************************************)
(* Instantiation of existential variables *)
(******************************************)
let depends_on_evar evk _ (pbty,_,t1,t2) =
- try head_evar t1 = evk
+ try Evar.equal (head_evar t1) evk
with NoHeadEvar ->
- try head_evar t2 = evk
+ try Evar.equal (head_evar t2) evk
with NoHeadEvar -> false
-let define_and_solve_constraints evk c evd =
- try
- let evd = define evk c evd in
- let (evd,pbs) = extract_changed_conv_pbs evd (depends_on_evar evk) in
- fst (List.fold_left
- (fun (evd,b as p) (pbty,env,t1,t2) ->
- if b then Evarconv.evar_conv_x full_transparent_state env evd pbty t1 t2 else p) (evd,true)
- pbs)
- with e when Pretype_errors.precatchable_exception e ->
- error "Instance does not satisfy constraints."
+let define_and_solve_constraints evk c env evd =
+ if Termops.occur_evar evk c then
+ Pretype_errors.error_occur_check env evd evk c;
+ let evd = define evk c evd in
+ let (evd,pbs) = extract_changed_conv_pbs evd (depends_on_evar evk) in
+ match
+ List.fold_left
+ (fun p (pbty,env,t1,t2) -> match p with
+ | Success evd -> Evarconv.evar_conv_x full_transparent_state env evd pbty t1 t2
+ | UnifFailure _ as x -> x) (Success evd)
+ pbs
+ with
+ | Success evd -> evd
+ | UnifFailure _ -> error "Instance does not satisfy the constraints."
let w_refine (evk,evi) (ltac_var,rawc) sigma =
if Evd.is_defined sigma evk then
error "Instantiate called on already-defined evar";
- let env = Evd.evar_env evi in
+ let env = Evd.evar_filtered_env evi in
let sigma',typed_c =
- try Pretyping.Default.understand_ltac ~resolve_classes:true true sigma env ltac_var
- (Pretyping.OfType (Some evi.evar_concl)) rawc
+ let flags = {
+ Pretyping.use_typeclasses = true;
+ Pretyping.use_unif_heuristics = true;
+ Pretyping.use_hook = None;
+ Pretyping.fail_evar = false;
+ Pretyping.expand_evars = true } in
+ try Pretyping.understand_ltac flags
+ env sigma ltac_var (Pretyping.OfType evi.evar_concl) rawc
with e when Errors.noncritical e ->
- let loc = Glob_term.loc_of_glob_constr rawc in
+ let loc = Glob_ops.loc_of_glob_constr rawc in
user_err_loc
(loc,"",Pp.str ("Instance is not well-typed in the environment of " ^
string_of_existential evk))
in
- define_and_solve_constraints evk typed_c (evars_reset_evd sigma' sigma)
+ define_and_solve_constraints evk typed_c env (evars_reset_evd sigma' sigma)
(* vernac command Existential *)
(* 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
+ let env = Evd.evar_filtered_env evi in
+ let rawc = Constrintern.intern_constr env com in
+ let ltac_vars = Pretyping.empty_lvar in
+ let sigma' = w_refine (evk, evi) (ltac_vars, rawc) sigma in
sigma'
diff --git a/proofs/evar_refiner.mli b/proofs/evar_refiner.mli
index 9fb9c09b..673dad55 100644
--- a/proofs/evar_refiner.mli
+++ b/proofs/evar_refiner.mli
@@ -1,18 +1,13 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Names
-open Term
-open Environ
open Evd
-open Refiner
open Pretyping
-open Glob_term
(** Refinement of existential variables. *)
@@ -20,6 +15,6 @@ val w_refine : evar * evar_info ->
glob_constr_ltac_closure -> evar_map -> evar_map
val instantiate_pf_com :
- Evd.evar -> Topconstr.constr_expr -> Evd.evar_map -> Evd.evar_map
+ Evd.evar -> Constrexpr.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
index 20527c62..e3570242 100644
--- a/proofs/goal.ml
+++ b/proofs/goal.ml
@@ -1,545 +1,105 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Util
open Pp
open Term
+open Vars
+open Context
(* This module implements the abstract interface to goals *)
-(* A general invariant of the module, is that a goal whose associated
+(* 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. *)
+type goal = Evd.evar
+let pr_goal e = str "GOAL:" ++ Pp.int (Evar.repr e)
-let pr_goal {content = e} = str "GOAL:" ++ Pp.int e
+let uid e = string_of_int (Evar.repr e)
+let get_by_uid u = Evar.unsafe_of_int (int_of_string u)
-(* 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 = []
- }
-
-
-let uid {content = e} = string_of_int e
-let get_by_uid u =
- (* this necessarily forget about tags.
- when tags are to be implemented, they
- should be done another way.
- We could use the store in evar_extra,
- for instance. *)
- build (int_of_string u)
-
-(* 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 true (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 ?filter ?split ?(fail=false) () env rdefs _ _ =
- rdefs:=Typeclasses.resolve_typeclasses ?filter ?split ~fail env !rdefs;
- ()
-
-
-
- (* a pessimistic (i.e : there won't be many positive answers) filter
- over evar_maps, acting only on undefined evars *)
- let evar_map_filter_undefined f evm =
- Evd.fold_undefined
- (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)
-
- let update_handle handle init_defs post_defs =
- (* [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_undefined
- (fun ev _ -> 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;
- delta_evars
-
- (* [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 [glob_constr] 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
- ignore(update_handle handle init_defs !rdefs);
- open_constr
-
- let constr_of_open_constr handle check_type (evars, c) env rdefs gl info =
- let delta = update_handle handle !rdefs evars in
- rdefs := Evd.fold (fun ev evi evd -> Evd.add evd ev evi) delta !rdefs;
- if check_type then with_type c (Evd.evar_concl (content !rdefs gl)) env rdefs gl info
- else c
-
-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 e when Errors.noncritical e ->
- 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
-(* no handler: this should never be allowed to reach toplevel *)
-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.
+(* 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
+ let env evars gl =
+ let evi = Evd.find evars gl in
+ Evd.evar_filtered_env evi
(* Old style hyps primitive *)
let hyps evars gl =
- let evi = content evars gl in
+ let evi = Evd.find evars gl in
Evd.evar_filtered_hyps evi
+ (* same as [hyps], but ensures that existential variables are
+ normalised. *)
+ let nf_hyps evars gl =
+ let hyps = Environ.named_context_of_val (hyps evars gl) in
+ Environ.val_of_named_context (Evarutil.nf_named_context_evar evars hyps)
+
(* Access to ".evar_concl" *)
let concl evars gl =
- let evi = content evars gl in
+ let evi = Evd.find evars gl in
evi.Evd.evar_concl
(* Access to ".evar_extra" *)
let extra evars gl =
- let evi = content evars gl in
+ let evi = Evd.find 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_candidates = None;
- Evd.evar_extra = extra }
+ (* A goal created that way will not be used by refine and will not
+ be shelved. It must not appear as a future_goal, so the future
+ goals are restored to their initial value after the evar is
+ created. *)
+ let prev_future_goals = Evd.future_goals evars in
+ let prev_principal_goal = Evd.principal_future_goal evars in
+ let evi = { Evd.evar_hyps = hyps;
+ Evd.evar_concl = concl;
+ Evd.evar_filter = Evd.Filter.identity;
+ Evd.evar_body = Evd.Evar_empty;
+ Evd.evar_source = (Loc.ghost,Evar_kinds.GoalEvar);
+ Evd.evar_candidates = None;
+ 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 (evars, evk) = Evarutil.new_pure_evar_full evars evi in
+ let evars = Evd.restore_future_goals evars prev_future_goals prev_principal_goal in
+ let ctxt = Environ.named_context_of_val hyps in
+ let inst = Array.map_of_list (fun (id, _, _) -> mkVar id) ctxt 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
+ (evk, ev, evars)
(* Instantiates a goal with an open term *)
- let partial_solution sigma { content=evk } c =
+ let partial_solution sigma evk c =
+ (* Check that the goal itself does not appear in the refined term *)
+ let _ =
+ if not (Evarutil.occur_evar_upto sigma evk c) then ()
+ else Pretype_errors.error_occur_check Environ.empty_env sigma evk c
+ in
Evd.define evk c sigma
-
+
+ (* Instantiates a goal with an open term, using name of goal for evk' *)
+ let partial_solution_to sigma evk evk' c =
+ let id = Evd.evar_ident evk sigma in
+ Evd.rename evk' id (partial_solution sigma evk c)
+
(* Parts of the progress tactical *)
let same_goal evars1 gl1 evars2 gl2 =
- let evi1 = content evars1 gl1 in
- let evi2 = content evars2 gl2 in
+ let evi1 = Evd.find evars1 gl1 in
+ let evi2 = Evd.find 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)
@@ -548,41 +108,40 @@ module V82 = struct
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)
+ (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 and change *)
+
+ (* Used for congruence closure *)
let new_goal_with sigma gl extra_hyps =
- let evi = content sigma gl in
+ let evi = Evd.find sigma gl in
let hyps = evi.Evd.evar_hyps in
let new_hyps =
List.fold_right Environ.push_named_context_val extra_hyps hyps in
- let extra_filter = List.map (fun _ -> true) extra_hyps in
- let new_filter = extra_filter @ evi.Evd.evar_filter in
+ let filter = evi.Evd.evar_filter in
+ let new_filter = Evd.Filter.extend (List.length extra_hyps) filter in
let new_evi =
{ evi with Evd.evar_hyps = new_hyps; Evd.evar_filter = new_filter } 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 }
+ let (new_sigma, evk) = Evarutil.new_pure_evar_full Evd.empty new_evi in
+ { Evd.it = evk ; sigma = new_sigma; }
(* Used by the compatibility layer and typeclasses *)
let nf_evar sigma gl =
- let evi = content sigma gl in
+ let evi = Evd.find sigma gl in
let evi = Evarutil.nf_evar_info sigma evi in
- let sigma = Evd.add sigma gl.content evi in
- (gl,sigma)
+ let sigma = Evd.add sigma gl evi in
+ (gl, sigma)
(* Goal represented as a type, doesn't take into account section variables *)
- let abstract_type sigma gl =
+ 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
+ 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
diff --git a/proofs/goal.mli b/proofs/goal.mli
index b7b20272..a00a95a2 100644
--- a/proofs/goal.mli
+++ b/proofs/goal.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,12 +8,7 @@
(* 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
+type goal = Evar.t
(* Gives a unique identifier to each goal. The identifier is
guaranteed to contain no space. *)
@@ -25,155 +20,6 @@ val get_by_uid : string -> goal
(* Debugging help *)
val pr_goal : goal -> Pp.std_ppcmds
-(* [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 : ?filter:(Evd.hole_kind -> 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 [glob_constr] 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 -> Glob_term.glob_constr -> Term.constr sensitive
-
- (* [constr_of_open_constr h check_type] transforms an open constr into a
- goal-sensitive constr, adding the undefined variables to the set of subgoals.
- If [check_type] is true, the term is coerced to the conclusion of the goal.
- It allows to do refinement with already-built terms with holes.
- *)
- val constr_of_open_constr : handle -> bool -> Evd.open_constr -> 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. *)
@@ -182,20 +28,18 @@ 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
+ (* same as [hyps], but ensures that existential variables are
+ normalised. *)
+ val nf_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
+ val extra : Evd.evar_map -> goal -> Evd.Store.t
(* Old style mk_goal primitive, returns a new goal with corresponding
hypotheses and conclusion, together with a term which is precisely
@@ -203,25 +47,16 @@ module V82 : sig
val mk_goal : Evd.evar_map ->
Environ.named_context_val ->
Term.constr ->
- Store.t ->
+ Evd.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
-
+
+ (* Instantiates a goal with an open term, reusing name of goal for
+ second goal *)
+ val partial_solution_to : Evd.evar_map -> goal -> goal -> Term.constr -> Evd.evar_map
+
(* Principal part of the weak-progress tactical *)
val weak_progress : goal list Evd.sigma -> goal Evd.sigma -> bool
@@ -232,7 +67,7 @@ module V82 : sig
val same_goal : Evd.evar_map -> goal -> Evd.evar_map -> goal -> bool
(* Used for congruence closure *)
- val new_goal_with : Evd.evar_map -> goal -> Sign.named_context -> goal Evd.sigma
+ val new_goal_with : Evd.evar_map -> goal -> Context.named_context -> goal Evd.sigma
(* Used by the compatibility layer and typeclasses *)
val nf_evar : Evd.evar_map -> goal -> goal * Evd.evar_map
diff --git a/proofs/logic.ml b/proofs/logic.ml
index 52ca0e00..53f8093e 100644
--- a/proofs/logic.ml
+++ b/proofs/logic.ml
@@ -1,37 +1,34 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Compat
open Pp
+open Errors
open Util
open Names
open Nameops
-open Evd
open Term
+open Vars
+open Context
open Termops
-open Sign
open Environ
open Reductionops
-open Inductive
open Inductiveops
open Typing
open Proof_type
-open Typeops
open Type_errors
open Retyping
-open Evarutil
-open Tacexpr
+open Misctypes
type refiner_error =
(* Errors raised by the refiner *)
| BadType of constr * constr * constr
- | UnresolvedBindings of name list
+ | UnresolvedBindings of Name.t list
| CannotApply of constr * constr
| NotWellTyped of constr
| NonLinearProof of constr
@@ -39,31 +36,39 @@ type refiner_error =
(* Errors raised by the tactics *)
| IntroNeedsProduct
- | DoesNotOccurIn of constr * identifier
+ | DoesNotOccurIn of constr * Id.t
+ | NoSuchHyp of Id.t
exception RefinerError of refiner_error
open Pretype_errors
-let rec catchable_exception = function
- | Loc.Exc_located(_,e) -> catchable_exception e
- | LtacLocated(_,e) -> catchable_exception e
- | Util.UserError _ | TypeError _ | PretypeError (_,_,TypingError _)
+(** FIXME: this is quite brittle. Why not accept any PretypeError? *)
+let is_typing_error = function
+| UnexpectedType (_, _) | NotProduct _
+| VarNotFound _ | TypingError _ -> true
+| _ -> false
+
+let is_unification_error = function
+| CannotUnify _ | CannotUnifyLocal _| CannotGeneralize _
+| NoOccurrenceFound _ | CannotUnifyBindingType _
+| ActualTypeNotCoercible _ | UnifOccurCheck _
+| CannotFindWellTypedAbstraction _ | WrongAbstractionType _
+| UnsolvableImplicit _| AbstractionOverMeta _
+| UnsatisfiableConstraints _ -> true
+| _ -> false
+
+let catchable_exception = function
+ | Errors.UserError _ | TypeError _
| RefinerError _ | Indrec.RecursionSchemeError _
- | Nametab.GlobalizationError _ | PretypeError (_,_,VarNotFound _)
+ | Nametab.GlobalizationError _
(* reduction errors *)
- | Tacred.ReductionTacticError _
- (* unification errors *)
- | PretypeError(_,_,(CannotUnify _|CannotUnifyLocal _|CannotGeneralize _
- |NoOccurrenceFound _|CannotUnifyBindingType _|NotClean _
- |CannotFindWellTypedAbstraction _|OccurCheck _
- |UnsolvableImplicit _|AbstractionOverMeta _)) -> true
- | Typeclasses_errors.TypeClassError
- (_, Typeclasses_errors.UnsatisfiableConstraints _) -> true
+ | Tacred.ReductionTacticError _ -> true
+ (* unification and typing errors *)
+ | PretypeError(_,_, e) -> is_unification_error e || is_typing_error e
| _ -> false
-let error_no_such_hypothesis id =
- error ("No such hypothesis: " ^ string_of_id id ^ ".")
+let error_no_such_hypothesis id = raise (RefinerError (NoSuchHyp id))
(* Tells if the refiner should check that the submitted rules do not
produce invalid subgoals *)
@@ -75,13 +80,13 @@ let with_check = Flags.with_option check
let apply_to_hyp sign id f =
try apply_to_hyp sign id f
with Hyp_not_found ->
- if !check then error "No such assumption."
+ if !check then error_no_such_hypothesis id
else sign
let apply_to_hyp_and_dependent_on sign id f g =
try apply_to_hyp_and_dependent_on sign id f g
with Hyp_not_found ->
- if !check then error "No such assumption."
+ if !check then error_no_such_hypothesis id
else sign
let check_typability env sigma c =
@@ -96,41 +101,17 @@ let check_typability env sigma c =
(instead of iterating on the list of identifier to be removed, which
forces the user to give them in order). *)
-let clear_hyps sigma ids sign cl =
+let clear_hyps env sigma ids sign cl =
let evdref = ref (Evd.create_goal_evar_defs sigma) in
- let (hyps,concl) = Evarutil.clear_hyps_in_evi evdref sign cl ids in
- (hyps,concl, !evdref)
+ let (hyps,cl) = Evarutil.clear_hyps_in_evi env evdref sign cl ids in
+ (hyps, cl, !evdref)
-(* The ClearBody tactic *)
+let clear_hyps2 env sigma ids sign t cl =
+ let evdref = ref (Evd.create_goal_evar_defs sigma) in
+ let (hyps,t,cl) = Evarutil.clear_hyps2_in_evi env evdref sign t cl ids in
+ (hyps, t, cl, !evdref)
-let recheck_typability (what,id) env sigma t =
- try check_typability env sigma t
- with e when Errors.noncritical e ->
- let s = match what with
- | None -> "the conclusion"
- | Some id -> "hypothesis "^(string_of_id id) in
- error
- ("The correctness of "^s^" relies on the body of "^(string_of_id id))
-
-let remove_hyp_body env sigma id =
- let sign =
- apply_to_hyp_and_dependent_on (named_context_val env) id
- (fun (_,c,t) _ ->
- match c with
- | None -> error ((string_of_id id)^" is not a local definition.")
- | Some c ->(id,None,t))
- (fun (id',c,t as d) sign ->
- (if !check then
- begin
- let env = 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
- reset_with_named_context sign env
+(* The ClearBody tactic *)
(* Reordering of the context *)
@@ -139,36 +120,36 @@ let remove_hyp_body env sigma id =
(* pas echangees. Choix: les hyps mentionnees ne peuvent qu'etre *)
(* reculees par rapport aux autres (faire le contraire!) *)
-let mt_q = (Idmap.empty,[])
+let mt_q = (Id.Map.empty,[])
let push_val y = function
(_,[] as q) -> q
- | (m, (x,l)::q) -> (m, (x,Idset.add y l)::q)
+ | (m, (x,l)::q) -> (m, (x,Id.Set.add y l)::q)
let push_item x v (m,l) =
- (Idmap.add x v m, (x,Idset.empty)::l)
-let mem_q x (m,_) = Idmap.mem x m
-let rec find_q x (m,q) =
- let v = Idmap.find x m in
- let m' = Idmap.remove x m in
+ (Id.Map.add x v m, (x,Id.Set.empty)::l)
+let mem_q x (m,_) = Id.Map.mem x m
+let find_q x (m,q) =
+ let v = Id.Map.find x m in
+ let m' = Id.Map.remove x m in
let rec find accs acc = function
[] -> raise Not_found
| [(x',l)] ->
- if x=x' then ((v,Idset.union accs l),(m',List.rev acc))
+ if Id.equal x x' then ((v,Id.Set.union accs l),(m',List.rev acc))
else raise Not_found
| (x',l as i)::((x'',l'')::q as itl) ->
- if x=x' then
- ((v,Idset.union accs l),
- (m',List.rev acc@(x'',Idset.add x (Idset.union l l''))::q))
- else find (Idset.union l accs) (i::acc) itl in
- find Idset.empty [] q
+ if Id.equal x x' then
+ ((v,Id.Set.union accs l),
+ (m',List.rev acc@(x'',Id.Set.add x (Id.Set.union l l''))::q))
+ else find (Id.Set.union l accs) (i::acc) itl in
+ find Id.Set.empty [] q
let occur_vars_in_decl env hyps d =
- if Idset.is_empty hyps then false else
+ if Id.Set.is_empty hyps then false else
let ohyps = global_vars_set_of_decl env d in
- Idset.exists (fun h -> Idset.mem h ohyps) hyps
+ Id.Set.exists (fun h -> Id.Set.mem h ohyps) hyps
let reorder_context env sign ord =
- let ords = List.fold_right Idset.add ord Idset.empty in
- if List.length ord <> Idset.cardinal ords then
+ let ords = List.fold_right Id.Set.add ord Id.Set.empty in
+ if not (Int.equal (List.length ord) (Id.Set.cardinal ords)) then
error "Order list has duplicates";
let rec step ord expected ctxt_head moved_hyps ctxt_tail =
match ord with
@@ -179,16 +160,16 @@ let reorder_context env sign ord =
errorlabstrm "reorder_context"
(str "Cannot move declaration " ++ pr_id top ++ spc() ++
str "before " ++
- prlist_with_sep pr_spc pr_id
- (Idset.elements (Idset.inter h
+ pr_sequence pr_id
+ (Id.Set.elements (Id.Set.inter h
(global_vars_set_of_decl env d))));
step ord' expected ctxt_head mh (d::ctxt_tail)
| _ ->
(match ctxt_head with
| [] -> error_no_such_hypothesis (List.hd ord)
| (x,_,_ as d) :: ctxt ->
- if Idset.mem x expected then
- step ord (Idset.remove x expected)
+ if Id.Set.mem x expected then
+ step ord (Id.Set.remove x expected)
ctxt (push_item x d moved_hyps) ctxt_tail
else
step ord expected
@@ -204,8 +185,8 @@ let reorder_val_context env sign ord =
let check_decl_position env sign (x,_,_ as d) =
let needed = global_vars_set_of_decl env d in
let deps = dependency_closure env (named_context_of_val sign) needed in
- if List.mem x deps then
- error ("Cannot create self-referring hypothesis "^string_of_id x);
+ if Id.List.mem x deps then
+ error ("Cannot create self-referring hypothesis "^Id.to_string x);
x::deps
(* Auxiliary functions for primitive MOVE tactic
@@ -216,11 +197,18 @@ let check_decl_position env sign (x,_,_ as d) =
* on the right side [right] if [toleft=false].
* If [with_dep] then dependent hypotheses are moved accordingly. *)
+let move_location_eq m1 m2 = match m1, m2 with
+| MoveAfter id1, MoveAfter id2 -> Id.equal id1 id2
+| MoveBefore id1, MoveBefore id2 -> Id.equal id1 id2
+| MoveLast, MoveLast -> true
+| MoveFirst, MoveFirst -> true
+| _ -> false
+
let rec get_hyp_after h = function
| [] -> error_no_such_hypothesis h
| (hyp,_,_) :: right ->
- if hyp = h then
- match right with (id,_,_)::_ -> MoveBefore id | [] -> MoveToEnd false
+ if Id.equal hyp h then
+ match right with (id,_,_)::_ -> MoveBefore id | [] -> MoveFirst
else
get_hyp_after h right
@@ -228,11 +216,14 @@ let split_sign hfrom hto l =
let rec splitrec left toleft = function
| [] -> error_no_such_hypothesis hfrom
| (hyp,c,typ) as d :: right ->
- if hyp = hfrom then
- (left,right,d, toleft or hto = MoveToEnd true)
+ if Id.equal hyp hfrom then
+ (left,right,d, toleft || move_location_eq hto MoveLast)
else
- splitrec (d::left)
- (toleft or hto = MoveAfter hyp or hto = MoveBefore hyp)
+ let is_toleft = match hto with
+ | MoveAfter h' | MoveBefore h' -> Id.equal hyp h'
+ | _ -> false
+ in
+ splitrec (d::left) (toleft || is_toleft)
right
in
splitrec [] false l
@@ -242,7 +233,7 @@ let hyp_of_move_location = function
| MoveBefore id -> id
| _ -> assert false
-let move_hyp with_dep toleft (left,(idfrom,_,_ as declfrom),right) hto =
+let move_hyp toleft (left,(idfrom,_,_ as declfrom),right) hto =
let env = Global.env() in
let test_dep (hyp,c,typ as d) (hyp2,c,typ2 as d2) =
if toleft
@@ -251,25 +242,25 @@ let move_hyp with_dep toleft (left,(idfrom,_,_ as declfrom),right) hto =
in
let rec moverec first middle = function
| [] ->
- if match hto with MoveToEnd _ -> false | _ -> true then
+ if match hto with MoveFirst | MoveLast -> false | _ -> true then
error_no_such_hypothesis (hyp_of_move_location hto);
List.rev first @ List.rev middle
- | (hyp,_,_) :: _ as right when hto = MoveBefore hyp ->
+ | (hyp,_,_) :: _ as right when move_location_eq hto (MoveBefore hyp) ->
List.rev first @ List.rev middle @ right
| (hyp,_,_) as d :: right ->
let (first',middle') =
if List.exists (test_dep d) middle then
- if with_dep & hto <> MoveAfter hyp then
+ if not (move_location_eq hto (MoveAfter hyp)) then
(first, d::middle)
else
errorlabstrm "move_hyp" (str "Cannot move " ++ pr_id idfrom ++
- pr_move_location pr_id hto ++
+ Miscprint.pr_move_location pr_id hto ++
str (if toleft then ": it occurs in " else ": it depends on ")
++ pr_id hyp ++ str ".")
else
(d::first, middle)
in
- if hto = MoveAfter hyp then
+ if move_location_eq hto (MoveAfter hyp) then
List.rev first' @ List.rev middle' @ right
else
moverec first' middle' right
@@ -291,6 +282,9 @@ let rename_hyp id1 id2 sign =
(fun (_,b,t) _ -> (id2,b,t))
(fun d _ -> map_named_declaration (replace_vars [id1,mkVar id2]) d)
+(**********************************************************************)
+
+
(************************************************************************)
(************************************************************************)
(* Implementation of the logical rules *)
@@ -308,21 +302,34 @@ let collect_meta_variables c =
| Meta mv -> if deep then error_unsupported_deep_meta () else mv::acc
| Cast(c,_,_) -> collrec deep acc c
| (App _| Case _) -> fold_constr (collrec deep) acc c
+ | Proj (_, c) -> collrec deep acc c
| _ -> fold_constr (collrec true) acc c
in
List.rev (collrec false [] c)
let check_meta_variables c =
- if not (list_distinct (collect_meta_variables c)) then
+ if not (List.distinct_f Int.compare (collect_meta_variables c)) then
raise (RefinerError (NonLinearProof c))
let check_conv_leq_goal env sigma arg ty conclty =
- if !check & not (is_conv_leq env sigma ty conclty) then
- raise (RefinerError (BadType (arg,ty,conclty)))
+ if !check then
+ let evm, b = Reductionops.infer_conv env sigma ty conclty in
+ if b then evm
+ else raise (RefinerError (BadType (arg,ty,conclty)))
+ else sigma
+
+exception Stop of constr list
+let meta_free_prefix a =
+ try
+ let _ = Array.fold_left (fun acc a ->
+ if occur_meta a then raise (Stop acc)
+ else a :: acc) [] a
+ in a
+ with Stop acc -> Array.rev_of_list acc
let goal_type_of env sigma c =
if !check then type_of env sigma c
- else Retyping.get_type_of ~refresh:true env sigma c
+ else Retyping.get_type_of env sigma c
let rec mk_refgoals sigma goal goalacc conclty trm =
let env = Goal.V82.env sigma goal in
@@ -330,62 +337,77 @@ let rec mk_refgoals sigma goal goalacc conclty trm =
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 _ ->
+ if (not !check) && not (occur_meta trm) then
+ let t'ty = Retyping.get_type_of env sigma trm in
+ let sigma = check_conv_leq_goal env sigma trm t'ty conclty in
+ (goalacc,t'ty,sigma,trm)
+ else
+ match kind_of_term trm with
+ | Meta _ ->
let conclty = nf_betaiota sigma conclty in
if !check && occur_meta conclty then
raise (RefinerError (MetaInType conclty));
let (gl,ev,sigma) = mk_goal hyps conclty in
gl::goalacc, conclty, sigma, ev
- | Cast (t,k, ty) ->
+ | Cast (t,k, ty) ->
check_typability env sigma ty;
- check_conv_leq_goal env sigma trm ty conclty;
+ let sigma = check_conv_leq_goal env sigma trm ty conclty in
let res = mk_refgoals sigma goal goalacc ty t in
- (** we keep the casts (in particular VMcast) except
+ (** we keep the casts (in particular VMcast and NATIVEcast) except
when they are annotating metas *)
if isMeta t then begin
- assert (k <> VMcast);
+ assert (k != VMcast && k != NATIVEcast);
res
end else
- let (gls,cty,sigma,trm) = res in
- (gls,cty,sigma,mkCast(trm,k,ty))
+ let (gls,cty,sigma,ans) = res in
+ let ans = if ans == t then trm else mkCast(ans,k,ty) in
+ (gls,cty,sigma,ans)
- | App (f,l) ->
+ | App (f,l) ->
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,
- sigma, f
- | _ ->
- mk_hdgoals sigma goal goalacc f
+ if is_template_polymorphic env f then
+ let sigma, ty =
+ (* Template sort-polymorphism of definition and inductive types *)
+ type_of_global_reference_knowing_conclusion env sigma f conclty
+ in
+ goalacc, ty, sigma, f
+ else
+ mk_hdgoals sigma goal goalacc f
in
- 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',sigma, Term.mkApp (applicand, Array.of_list args))
-
- | Case (ci,p,c,lf) ->
+ let ((acc'',conclty',sigma), args) = mk_arggoals sigma goal acc' hdty l in
+ let sigma = check_conv_leq_goal env sigma trm conclty' conclty in
+ let ans = if applicand == f && args == l then trm else Term.mkApp (applicand, args) in
+ (acc'',conclty',sigma, ans)
+
+ | Proj (p,c) ->
+ let (acc',cty,sigma,c') = mk_hdgoals sigma goal goalacc c in
+ let c = mkProj (p, c') in
+ let ty = get_type_of env sigma c in
+ (acc',ty,sigma,c)
+
+ | 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 sigma = check_conv_leq_goal env sigma trm conclty' conclty in
let (acc'',sigma, rbranches) =
- array_fold_left2
+ Array.fold_left2
(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',sigma, Term.mkCase (ci,p',c',Array.of_list (List.rev rbranches)))
+ let lf' = Array.rev_of_list rbranches in
+ let ans =
+ if p' == p && c' == c && Array.equal (==) lf' lf then trm
+ else Term.mkCase (ci,p',c',lf')
+ in
+ (acc'',conclty',sigma, ans)
- | _ ->
+ | _ ->
if occur_meta trm then
- anomaly "refiner called with a meta in non app/case subterm";
-
- let t'ty = goal_type_of env sigma trm in
- check_conv_leq_goal env sigma trm t'ty conclty;
- (goalacc,t'ty,sigma, trm)
+ anomaly (Pp.str "refiner called with a meta in non app/case subterm");
+ let t'ty = goal_type_of env sigma trm in
+ let sigma = check_conv_leq_goal env sigma trm t'ty conclty in
+ (goalacc,t'ty,sigma, trm)
(* Same as mkREFGOALS but without knowing the type of the term. Therefore,
* Metas should be casted. *)
@@ -407,44 +429,57 @@ and mk_hdgoals sigma goal goalacc trm =
| App (f,l) ->
let (acc',hdty,sigma,applicand) =
- if isInd f or isConst f
- & not (array_exists occur_meta l) (* we could be finer *)
+ if is_template_polymorphic env f
then
- (goalacc,type_of_global_reference_knowing_parameters env sigma f l,sigma,f)
+ let l' = meta_free_prefix l in
+ (goalacc,type_of_global_reference_knowing_parameters env sigma f l',sigma,f)
else mk_hdgoals sigma goal goalacc f
in
- 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))
+ let ((acc'',conclty',sigma), args) = mk_arggoals sigma goal acc' hdty l in
+ let ans = if applicand == f && args == l then trm else Term.mkApp (applicand, args) in
+ (acc'',conclty',sigma, ans)
| 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
+ Array.fold_left2
(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',sigma, Term.mkCase (ci,p',c',Array.of_list (List.rev rbranches)))
+ let lf' = Array.rev_of_list rbranches in
+ let ans =
+ if p' == p && c' == c && Array.equal (==) lf' lf then trm
+ else Term.mkCase (ci,p',c',lf')
+ in
+ (acc'',conclty',sigma, ans)
+
+ | Proj (p,c) ->
+ let (acc',cty,sigma,c') = mk_hdgoals sigma goal goalacc c in
+ let c = mkProj (p, c') in
+ let ty = get_type_of env sigma c in
+ (acc',ty,sigma,c)
| _ ->
if !check && occur_meta trm then
- anomaly "refine called with a dependent meta";
+ anomaly (Pp.str "refine called with a dependent meta");
goalacc, goal_type_of env sigma trm, sigma, trm
-and mk_arggoals sigma goal goalacc funty = function
- | [] -> goalacc,funty,sigma, []
- | harg::tlargs as allargs ->
- let t = whd_betadeltaiota (Goal.V82.env sigma goal) sigma funty in
- match kind_of_term t with
- | Prod (_,c1,b) ->
- 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_arggoals sigma goal goalacc funty allargs =
+ let foldmap (goalacc, funty, sigma) harg =
+ let t = whd_betadeltaiota (Goal.V82.env sigma goal) sigma funty in
+ let rec collapse t = match kind_of_term t with
+ | LetIn (_, c1, _, b) -> collapse (subst1 c1 b)
+ | _ -> t
+ in
+ let t = collapse t in
+ match kind_of_term t with
+ | Prod (_, c1, b) ->
+ let (acc, hargty, sigma, arg) = mk_refgoals sigma goal goalacc c1 harg in
+ (acc, subst1 harg b, sigma), arg
+ | _ -> raise (RefinerError (CannotApply (t, harg)))
+ in
+ Array.smartfoldmap foldmap (goalacc, funty, sigma) allargs
and mk_casegoals sigma goal goalacc p c =
let env = Goal.V82.env sigma goal in
@@ -452,24 +487,23 @@ and mk_casegoals sigma goal goalacc p c =
let (acc'',pt,sigma,p') = mk_hdgoals sigma goal acc' p in
let indspec =
try Tacred.find_hnf_rectype env sigma ct
- with Not_found -> anomaly "mk_casegoals" in
- let (lbrty,conclty) =
- type_case_branches_with_names env indspec p c in
+ with Not_found -> anomaly (Pp.str "mk_casegoals") in
+ let (lbrty,conclty) = type_case_branches_with_names env indspec p c in
(acc'',lbrty,conclty,sigma,p',c')
-let convert_hyp sign sigma (id,b,bt as d) =
+let convert_hyp check sign sigma (id,b,bt as d) =
let env = Global.env() in
let reorder = ref [] in
let sign' =
apply_to_hyp sign id
(fun _ (_,c,ct) _ ->
let env = Global.env_of_context sign in
- if !check && not (is_conv env sigma bt ct) then
- error ("Incorrect change of the type of "^(string_of_id id)^".");
- if !check && not (Option.Misc.compare (is_conv env sigma) b c) then
- error ("Incorrect change of the body of "^(string_of_id id)^".");
- if !check then reorder := check_decl_position env sign d;
+ if check && not (is_conv env sigma bt ct) then
+ error ("Incorrect change of the type of "^(Id.to_string id)^".");
+ if check && not (Option.equal (is_conv env sigma) b c) then
+ error ("Incorrect change of the body of "^(Id.to_string id)^".");
+ if check then reorder := check_decl_position env sign d;
d) in
reorder_val_context env sign' !reorder
@@ -488,50 +522,31 @@ let prim_refiner r sigma goal =
in
match r with
(* Logical rules *)
- | Intro id ->
- if !check && mem_named_context id (named_context_of_val sign) then
- error ("Variable " ^ string_of_id id ^ " is already declared.");
- (match kind_of_term (strip_outer_cast cl) with
- | Prod (_,c1,b) ->
- 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,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) ->
+(* if !check && not (Retyping.get_sort_of env sigma t) then*)
let (sg1,ev1,sigma) = mk_goal sign (nf_betaiota sigma t) in
- let sign,cl,sigma =
+ let sign,t,cl,sigma =
if replace then
let nexthyp = get_hyp_after id (named_context_of_val sign) in
- let sign,cl,sigma = clear_hyps sigma [id] sign cl in
- move_hyp true false ([],(id,None,t),named_context_of_val sign)
+ let sign,t,cl,sigma = clear_hyps2 env sigma (Id.Set.singleton id) sign t cl in
+ move_hyp false ([],(id,None,t),named_context_of_val sign)
nexthyp,
- cl,sigma
+ t,cl,sigma
else
(if !check && mem_named_context id (named_context_of_val sign) then
- error ("Variable " ^ string_of_id id ^ " is already declared.");
- push_named_context_val (id,None,t) sign,cl,sigma) in
+ error ("Variable " ^ Id.to_string id ^ " is already declared.");
+ push_named_context_val (id,None,t) sign,t,cl,sigma) 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
+ let oterm = Term.mkApp (mkNamedLambda id t ev2 , [| ev1 |]) in
+ let sigma = Goal.V82.partial_solution_to sigma goal sg2 oterm in
if b then ([sg1;sg2],sigma) else ([sg2;sg1],sigma)
| FixRule (f,n,rest,j) ->
let rec check_ind env k cl =
match kind_of_term (strip_outer_cast cl) with
| Prod (na,c1,b) ->
- if k = 1 then
+ if Int.equal k 1 then
try
fst (find_inductive env sigma c1)
with Not_found ->
@@ -540,31 +555,30 @@ let prim_refiner r sigma goal =
check_ind (push_rel (na,None,c1) env) (k-1) b
| _ -> error "Not enough products."
in
- let (sp,_) = check_ind env n cl in
- let firsts,lasts = list_chop j rest in
+ let ((sp,_),u) = check_ind env n cl in
+ let firsts,lasts = List.chop j rest in
let all = firsts@(f,n,cl)::lasts in
let rec mk_sign sign = function
| (f,n,ar)::oth ->
- let (sp',_) = check_ind env n ar in
- if not (sp=sp') then
+ let ((sp',_),u') = check_ind env n ar in
+ if not (eq_mind sp sp') then
error ("Fixpoints should be on the same " ^
"mutual inductive declaration.");
if !check && mem_named_context f (named_context_of_val sign) then
error
- ("Name "^string_of_id f^" already used in the environment");
+ ("Name "^Id.to_string f^" already used in the environment");
mk_sign (push_named_context_val (f,None,ar) sign) oth
| [] ->
- 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
+ Evd.Monad.List.map (fun (_,_,c) sigma ->
+ let gl,ev,sig' =
+ Goal.V82.mk_goal sigma sign c (Goal.V82.extra sigma goal) in
+ (gl,ev),sig')
+ all sigma
in
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 evs = List.map (Vars.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
@@ -585,7 +599,7 @@ let prim_refiner r sigma goal =
error ("All methods must construct elements " ^
"in coinductive types.")
in
- let firsts,lasts = list_chop j others in
+ let firsts,lasts = List.chop j others in
let all = firsts@(f,cl)::lasts in
List.iter (fun (_,c) -> check_is_coind env c) all;
let rec mk_sign sign = function
@@ -596,18 +610,17 @@ let prim_refiner r sigma goal =
with
| Not_found ->
mk_sign (push_named_context_val (f,None,ar) sign) oth)
- | [] -> 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
+ | [] ->
+ Evd.Monad.List.map (fun (_,c) sigma ->
+ let gl,ev,sigma =
+ Goal.V82.mk_goal sigma sign c (Goal.V82.extra sigma goal) in
+ (gl,ev),sigma)
+ all sigma
in
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 evs = List.map (Vars.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
@@ -622,87 +635,21 @@ let prim_refiner r sigma goal =
let sigma = Goal.V82.partial_solution sigma goal oterm in
(sgl, sigma)
- (* Conversion rules *)
- | Convert_concl (cl',k) ->
- check_typability env sigma cl';
- if (not !check) || is_conv_leq env sigma cl' cl then
- let (sg,ev,sigma) = mk_goal sign cl' in
- let ev = if k<>DEFAULTcast then mkCast(ev,k,cl) else ev 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) ->
- 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
+ let ids = List.fold_left (fun accu x -> Id.Set.add x accu) Id.Set.empty ids in
+ let (hyps,concl,nsigma) = clear_hyps env sigma ids sign cl in
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
+ let sigma = Goal.V82.partial_solution_to sigma goal gl ev in
([gl], sigma)
- | ThinBody ids ->
- let clear_aux env id =
- let env' = remove_hyp_body env sigma id in
- if !check then recheck_typability (None,id) env' sigma cl;
- env'
- in
- let sign' = named_context_val (List.fold_left clear_aux env ids) 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) ->
+ | Move (hfrom, hto) ->
let (left,right,declfrom,toleft) =
split_sign hfrom hto (named_context_of_val sign) in
let hyps' =
- move_hyp withdep toleft (left,declfrom,right) hto in
+ move_hyp toleft (left,declfrom,right) hto in
let (gl,ev,sigma) = mk_goal hyps' cl in
- let sigma = Goal.V82.partial_solution sigma goal ev in
+ let sigma = Goal.V82.partial_solution_to sigma goal gl ev in
([gl], sigma)
-
- | Order ord ->
- let hyps' = reorder_val_context env sign ord in
- 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 &&
- List.mem id2 (ids_of_named_context (named_context_of_val sign)) then
- 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
- 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 ->
- (* Normalises evars in goals. Used by instantiate. *)
- let (goal,sigma) = Goal.V82.nf_evar sigma goal in
- ([goal],sigma)
-
-(************************************************************************)
-(************************************************************************)
-(* Extracting a proof term from the proof tree *)
-
-(* Util *)
-
-type variable_proof_status = ProofVar | SectionVar of identifier
-
-type proof_variable = name * variable_proof_status
-
-let proof_variable_index x =
- let rec aux n = function
- | (Name id,ProofVar)::l when x = id -> n
- | _::l -> aux (n+1) l
- | [] -> raise Not_found
- in
- aux 1
diff --git a/proofs/logic.mli b/proofs/logic.mli
index 1044b59e..d034b73c 100644
--- a/proofs/logic.mli
+++ b/proofs/logic.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,9 +8,7 @@
open Names
open Term
-open Sign
open Evd
-open Environ
open Proof_type
(** This suppresses check done in [prim_refiner] for the tactic given in
@@ -32,10 +30,6 @@ val with_check : tactic -> tactic
val prim_refiner : prim_rule -> evar_map -> goal -> goal list * evar_map
-type proof_variable
-
-
-val proof_variable_index : identifier -> proof_variable list -> int
(** {6 Refiner errors. } *)
@@ -43,7 +37,7 @@ type refiner_error =
(*i Errors raised by the refiner i*)
| BadType of constr * constr * constr
- | UnresolvedBindings of name list
+ | UnresolvedBindings of Name.t list
| CannotApply of constr * constr
| NotWellTyped of constr
| NonLinearProof of constr
@@ -51,8 +45,12 @@ type refiner_error =
(*i Errors raised by the tactics i*)
| IntroNeedsProduct
- | DoesNotOccurIn of constr * identifier
+ | DoesNotOccurIn of constr * Id.t
+ | NoSuchHyp of Id.t
exception RefinerError of refiner_error
val catchable_exception : exn -> bool
+
+val convert_hyp : bool -> Environ.named_context_val -> evar_map ->
+ Context.named_declaration -> Environ.named_context_val
diff --git a/proofs/logic_monad.ml b/proofs/logic_monad.ml
new file mode 100644
index 00000000..d509670e
--- /dev/null
+++ b/proofs/logic_monad.ml
@@ -0,0 +1,326 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** This file defines the low-level monadic operations used by the
+ tactic monad. The monad is divided into two layers: a non-logical
+ layer which consists in operations which will not (or cannot) be
+ backtracked in case of failure (input/output or persistent state)
+ and a logical layer which handles backtracking, proof
+ manipulation, and any other effect which needs to backtrack. *)
+
+
+(** {6 Exceptions} *)
+
+
+(** To help distinguish between exceptions raised by the IO monad from
+ the one used natively by Coq, the former are wrapped in
+ [Exception]. It is only used internally so that [catch] blocks of
+ the IO monad would only catch exceptions raised by the [raise]
+ function of the IO monad, and not for instance, by system
+ interrupts. Also used in [Proofview] to avoid capturing exception
+ from the IO monad ([Proofview] catches errors in its compatibility
+ layer, and when lifting goal-level expressions). *)
+exception Exception of exn
+(** This exception is used to signal abortion in [timeout] functions. *)
+exception Timeout
+(** This exception is used by the tactics to signal failure by lack of
+ successes, rather than some other exceptions (like system
+ interrupts). *)
+exception TacticFailure of exn
+
+let _ = Errors.register_handler begin function
+ | Timeout -> Errors.errorlabstrm "Some timeout function" (Pp.str"Timeout!")
+ | Exception e -> Errors.print e
+ | TacticFailure e -> Errors.print e
+ | _ -> Pervasives.raise Errors.Unhandled
+end
+
+(** {6 Non-logical layer} *)
+
+(** The non-logical monad is a simple [unit -> 'a] (i/o) monad. The
+ operations are simple wrappers around corresponding usual
+ operations and require little documentation. *)
+module NonLogical =
+struct
+
+ (* The functions in this module follow the pattern that they are
+ defined with the form [(); fun ()->...]. This is an optimisation
+ which signals to the compiler that the function is usually partially
+ applied up to the [();]. Without this annotation, partial
+ applications can be significantly slower.
+
+ Documentation of this behaviour can be found at:
+ https://ocaml.janestreet.com/?q=node/30 *)
+
+ include Monad.Make(struct
+ type 'a t = unit -> 'a
+
+ let return a = (); fun () -> a
+ let (>>=) a k = (); fun () -> k (a ()) ()
+ let (>>) a k = (); fun () -> a (); k ()
+ let map f a = (); fun () -> f (a ())
+ end)
+
+ type 'a ref = 'a Pervasives.ref
+
+ let ignore a = (); fun () -> ignore (a ())
+
+ let ref a = (); fun () -> Pervasives.ref a
+
+ (** [Pervasives.(:=)] *)
+ let (:=) r a = (); fun () -> r := a
+
+ (** [Pervasives.(!)] *)
+ let (!) = fun r -> (); fun () -> ! r
+
+ (** [Pervasives.raise]. Except that exceptions are wrapped with
+ {!Exception}. *)
+ let raise ?info = fun e -> (); fun () -> Exninfo.raise ?info (Exception e)
+
+ (** [try ... with ...] but restricted to {!Exception}. *)
+ let catch = fun s h -> ();
+ fun () -> try s ()
+ with Exception e as src ->
+ let (src, info) = Errors.push src in
+ h (e, info) ()
+
+ let read_line = fun () -> try Pervasives.read_line () with e ->
+ let (e, info) = Errors.push e in raise ~info e ()
+
+ let print_char = fun c -> (); fun () -> print_char c
+
+ (** {!Pp.pp}. The buffer is also flushed. *)
+ let print = fun s -> (); fun () -> try Pp.msg_info s; Pp.pp_flush () with e ->
+ let (e, info) = Errors.push e in raise ~info e ()
+
+ let timeout = fun n t -> (); fun () ->
+ Control.timeout n t (Exception Timeout)
+
+ let make f = (); fun () ->
+ try f ()
+ with e when Errors.noncritical e ->
+ let (e, info) = Errors.push e in
+ Util.iraise (Exception e, info)
+
+ let run = fun x ->
+ try x () with Exception e as src ->
+ let (src, info) = Errors.push src in
+ Util.iraise (e, info)
+end
+
+(** {6 Logical layer} *)
+
+(** The logical monad is a backtracking monad on top of which is
+ layered a state monad (which is used to implement all of read/write,
+ read only, and write only effects). The state monad being layered on
+ top of the backtracking monad makes it so that the state is
+ backtracked on failure.
+
+ Backtracking differs from regular exception in that, writing (+)
+ for exception catching and (>>=) for bind, we require the
+ following extra distributivity laws:
+
+ x+(y+z) = (x+y)+z
+
+ zero+x = x
+
+ x+zero = x
+
+ (x+y)>>=k = (x>>=k)+(y>>=k) *)
+
+(** A view type for the logical monad, which is a form of list, hence
+ we can decompose it with as a list. *)
+type ('a, 'b) list_view =
+ | Nil of Exninfo.iexn
+ | Cons of 'a * 'b
+
+module type Param = sig
+
+ (** Read only *)
+ type e
+
+ (** Write only *)
+ type w
+
+ (** [w] must be a monoid *)
+ val wunit : w
+ val wprod : w -> w -> w
+
+ (** Read-write *)
+ type s
+
+ (** Update-only. Essentially a writer on [u->u]. *)
+ type u
+
+ (** [u] must be pointed. *)
+ val uunit : u
+
+end
+
+
+module Logical (P:Param) =
+struct
+
+ (** All three of environment, writer and state are coded as a single
+ state-passing-style monad.*)
+ type state = {
+ rstate : P.e;
+ ustate : P.u;
+ wstate : P.w;
+ sstate : P.s;
+ }
+
+ (** Double-continuation backtracking monads are reasonable folklore
+ for "search" implementations (including the 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. Shan, D. Friedman, and A. Sabry. The peculiar
+ shape of the monadic type is reminiscent of that of the
+ continuation monad transformer.
+
+ The paper also contains the rational for the [split] abstraction.
+
+ An explanation of how to derive such a monad from mathematical
+ principles can be found in "Kan Extensions for Program
+ Optimisation" by Ralf Hinze.
+
+ A somewhat concrete view is that the type ['a iolist] is, in fact
+ the impredicative encoding of the following stream type:
+
+ [type 'a _iolist' = Nil of exn | Cons of 'a*'a iolist'
+ and 'a iolist = 'a _iolist NonLogical.t]
+
+ Using impredicative encoding avoids intermediate allocation and
+ is, empirically, very efficient in Ocaml. It also has the
+ practical benefit that the monadic operation are independent of
+ the underlying monad, which simplifies the code and side-steps
+ the limited inlining of Ocaml.
+
+ In that vision, [bind] is simply [concat_map] (though the cps
+ version is significantly simpler), [plus] is concatenation, and
+ [split] is pattern-matching. *)
+ type rich_exn = Exninfo.iexn
+
+ type 'a iolist =
+ { iolist : 'r. (rich_exn -> 'r NonLogical.t) ->
+ ('a -> (rich_exn -> 'r NonLogical.t) -> 'r NonLogical.t) ->
+ 'r NonLogical.t }
+
+ include Monad.Make(struct
+ type 'a t = state -> ('a * state) iolist
+
+ let return x : 'a t = (); fun s ->
+ { iolist = fun nil cons -> cons (x, s) nil }
+
+ let (>>=) (m : 'a t) (f : 'a -> 'b t) : 'b t = (); fun s ->
+ let m = m s in
+ { iolist = fun nil cons ->
+ m.iolist nil (fun (x, s) next -> (f x s).iolist next cons) }
+
+ let (>>) (m : unit t) (f : 'a t) : 'a t = (); fun s ->
+ let m = m s in
+ { iolist = fun nil cons ->
+ m.iolist nil (fun ((), s) next -> (f s).iolist next cons) }
+
+ let map (f : 'a -> 'b) (m : 'a t) : 'b t = (); fun s ->
+ let m = m s in
+ { iolist = fun nil cons -> m.iolist nil (fun (x, s) next -> cons (f x, s) next) }
+
+ end)
+
+ let zero e : 'a t = (); fun s ->
+ { iolist = fun nil cons -> nil e }
+
+ let plus m1 m2 : 'a t = (); fun s ->
+ let m1 = m1 s in
+ { iolist = fun nil cons -> m1.iolist (fun e -> (m2 e s).iolist nil cons) cons }
+
+ let ignore (m : 'a t) : unit t = (); fun s ->
+ let m = m s in
+ { iolist = fun nil cons -> m.iolist nil (fun (_, s) next -> cons ((), s) next) }
+
+ let lift (m : 'a NonLogical.t) : 'a t = (); fun s ->
+ { iolist = fun nil cons -> NonLogical.(m >>= fun x -> cons (x, s) nil) }
+
+ (** State related *)
+
+ let get : P.s t = (); fun s ->
+ { iolist = fun nil cons -> cons (s.sstate, s) nil }
+
+ let set (sstate : P.s) : unit t = (); fun s ->
+ { iolist = fun nil cons -> cons ((), { s with sstate }) nil }
+
+ let modify (f : P.s -> P.s) : unit t = (); fun s ->
+ { iolist = fun nil cons -> cons ((), { s with sstate = f s.sstate }) nil }
+
+ let current : P.e t = (); fun s ->
+ { iolist = fun nil cons -> cons (s.rstate, s) nil }
+
+ let local (type a) (e:P.e) (m:a t) : a t = (); fun s ->
+ let m = m { s with rstate = e } in
+ { iolist = fun nil cons ->
+ m.iolist nil (fun (x,s') next -> cons (x,{s' with rstate=s.rstate}) next) }
+
+ let put (w : P.w) : unit t = (); fun s ->
+ { iolist = fun nil cons -> cons ((), { s with wstate = P.wprod s.wstate w }) nil }
+
+ let update (f : P.u -> P.u) : unit t = (); fun s ->
+ { iolist = fun nil cons -> cons ((), { s with ustate = f s.ustate }) nil }
+
+ (** List observation *)
+
+ let once (m : 'a t) : 'a t = (); fun s ->
+ let m = m s in
+ { iolist = fun nil cons -> m.iolist nil (fun x _ -> cons x nil) }
+
+ let break (f : rich_exn -> rich_exn option) (m : 'a t) : 'a t = (); fun s ->
+ let m = m s in
+ { iolist = fun nil cons ->
+ m.iolist nil (fun x next -> cons x (fun e -> match f e with None -> next e | Some e -> nil e))
+ }
+
+ (** For [reflect] and [split] see the "Backtracking, Interleaving,
+ and Terminating Monad Transformers" paper. *)
+ type 'a reified = ('a, rich_exn -> 'a reified) list_view NonLogical.t
+
+ let rec reflect (m : 'a reified) : 'a iolist =
+ { iolist = fun nil cons ->
+ let next = function
+ | Nil e -> nil e
+ | Cons (x, l) -> cons x (fun e -> (reflect (l e)).iolist nil cons)
+ in
+ NonLogical.(m >>= next)
+ }
+
+ let split (m : 'a t) : ('a, rich_exn -> 'a t) list_view t = (); fun s ->
+ let m = m s in
+ let rnil e = NonLogical.return (Nil e) in
+ let rcons p l = NonLogical.return (Cons (p, l)) in
+ { iolist = fun nil cons ->
+ let open NonLogical in
+ m.iolist rnil rcons >>= begin function
+ | Nil e -> cons (Nil e, s) nil
+ | Cons ((x, s), l) ->
+ let l e = (); fun _ -> reflect (l e) in
+ cons (Cons (x, l), s) nil
+ end }
+
+ let run m r s =
+ let s = { wstate = P.wunit; ustate = P.uunit; rstate = r; sstate = s } in
+ let m = m s in
+ let rnil e = NonLogical.return (Nil e) in
+ let rcons (x, s) l =
+ let p = (x, s.sstate, s.wstate, s.ustate) in
+ NonLogical.return (Cons (p, l))
+ in
+ m.iolist rnil rcons
+
+ let repr x = x
+
+ end
diff --git a/proofs/logic_monad.mli b/proofs/logic_monad.mli
new file mode 100644
index 00000000..ab729aff
--- /dev/null
+++ b/proofs/logic_monad.mli
@@ -0,0 +1,157 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** This file defines the low-level monadic operations used by the
+ tactic monad. The monad is divided into two layers: a non-logical
+ layer which consists in operations which will not (or cannot) be
+ backtracked in case of failure (input/output or persistent state)
+ and a logical layer which handles backtracking, proof
+ manipulation, and any other effect which needs to backtrack. *)
+
+
+(** {6 Exceptions} *)
+
+
+(** To help distinguish between exceptions raised by the IO monad from
+ the one used natively by Coq, the former are wrapped in
+ [Exception]. It is only used internally so that [catch] blocks of
+ the IO monad would only catch exceptions raised by the [raise]
+ function of the IO monad, and not for instance, by system
+ interrupts. Also used in [Proofview] to avoid capturing exception
+ from the IO monad ([Proofview] catches errors in its compatibility
+ layer, and when lifting goal-level expressions). *)
+exception Exception of exn
+(** This exception is used to signal abortion in [timeout] functions. *)
+exception Timeout
+(** This exception is used by the tactics to signal failure by lack of
+ successes, rather than some other exceptions (like system
+ interrupts). *)
+exception TacticFailure of exn
+
+
+(** {6 Non-logical layer} *)
+
+(** The non-logical monad is a simple [unit -> 'a] (i/o) monad. The
+ operations are simple wrappers around corresponding usual
+ operations and require little documentation. *)
+module NonLogical : sig
+
+ include Monad.S
+
+ val ignore : 'a t -> unit t
+
+ type 'a ref
+
+ val ref : 'a -> 'a ref t
+ (** [Pervasives.(:=)] *)
+ val (:=) : 'a ref -> 'a -> unit t
+ (** [Pervasives.(!)] *)
+ val (!) : 'a ref -> 'a t
+
+ val read_line : string t
+ val print_char : char -> unit t
+ (** {!Pp.pp}. The buffer is also flushed. *)
+ val print : Pp.std_ppcmds -> unit t
+
+ (** [Pervasives.raise]. Except that exceptions are wrapped with
+ {!Exception}. *)
+ val raise : ?info:Exninfo.info -> exn -> 'a t
+ (** [try ... with ...] but restricted to {!Exception}. *)
+ val catch : 'a t -> (Exninfo.iexn -> 'a t) -> 'a t
+ val timeout : int -> 'a t -> 'a t
+
+ (** Construct a monadified side-effect. Exceptions raised by the argument are
+ wrapped with {!Exception}. *)
+ val make : (unit -> 'a) -> 'a t
+
+ (** [run] performs effects. *)
+ val run : 'a t -> 'a
+
+end
+
+
+(** {6 Logical layer} *)
+
+(** The logical monad is a backtracking monad on top of which is
+ layered a state monad (which is used to implement all of read/write,
+ read only, and write only effects). The state monad being layered on
+ top of the backtracking monad makes it so that the state is
+ backtracked on failure.
+
+ Backtracking differs from regular exception in that, writing (+)
+ for exception catching and (>>=) for bind, we require the
+ following extra distributivity laws:
+
+ x+(y+z) = (x+y)+z
+
+ zero+x = x
+
+ x+zero = x
+
+ (x+y)>>=k = (x>>=k)+(y>>=k) *)
+
+(** A view type for the logical monad, which is a form of list, hence
+ we can decompose it with as a list. *)
+type ('a, 'b) list_view =
+| Nil of Exninfo.iexn
+| Cons of 'a * 'b
+
+(** The monad is parametrised in the types of state, environment and
+ writer. *)
+module type Param = sig
+
+ (** Read only *)
+ type e
+
+ (** Write only *)
+ type w
+
+ (** [w] must be a monoid *)
+ val wunit : w
+ val wprod : w -> w -> w
+
+ (** Read-write *)
+ type s
+
+ (** Update-only. Essentially a writer on [u->u]. *)
+ type u
+
+ (** [u] must be pointed. *)
+ val uunit : u
+
+end
+
+module Logical (P:Param) : sig
+
+ include Monad.S
+
+ val ignore : 'a t -> unit t
+
+ val set : P.s -> unit t
+ val get : P.s t
+ val modify : (P.s -> P.s) -> unit t
+ val put : P.w -> unit t
+ val current : P.e t
+ val local : P.e -> 'a t -> 'a t
+ val update : (P.u -> P.u) -> unit t
+
+ val zero : Exninfo.iexn -> 'a t
+ val plus : 'a t -> (Exninfo.iexn -> 'a t) -> 'a t
+ val split : 'a t -> (('a,(Exninfo.iexn->'a t)) list_view) t
+ val once : 'a t -> 'a t
+ val break : (Exninfo.iexn -> Exninfo.iexn option) -> 'a t -> 'a t
+
+ val lift : 'a NonLogical.t -> 'a t
+
+ type 'a reified
+
+ val repr : 'a reified -> ('a, Exninfo.iexn -> 'a reified) list_view NonLogical.t
+
+ val run : 'a t -> P.e -> P.s -> ('a * P.s * P.w * P.u) reified
+
+end
diff --git a/proofs/pfedit.ml b/proofs/pfedit.ml
index 77e7b324..fdc93bcb 100644
--- a/proofs/pfedit.ml
+++ b/proofs/pfedit.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -9,19 +9,9 @@
open Pp
open Util
open Names
-open Nameops
-open Sign
-open Term
-open Declarations
open Entries
open Environ
open Evd
-open Typing
-open Refiner
-open Tacexpr
-open Proof_type
-open Lib
-open Safe_typing
let refining = Proof_global.there_are_pending_proofs
let check_no_pending_proofs = Proof_global.check_no_pending_proof
@@ -35,58 +25,28 @@ let delete_proof = Proof_global.discard
let delete_current_proof = Proof_global.discard_current
let delete_all_proofs = Proof_global.discard_all
-let undo n =
- let p = Proof_global.give_me_the_proof () in
- let d = Proof.V82.depth p in
- if n >= d then raise Proof.EmptyUndoStack;
- for i = 1 to n do
- Proof.undo p
- done
-
-let current_proof_depth () =
- try
- let p = Proof_global.give_me_the_proof () in
- Proof.V82.depth p
- with Proof_global.NoCurrentProof -> -1
-
-(* [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
- undo ((current_proof_depth ()) - n )
- with Proof_global.NoCurrentProof when n=0 -> ()
-
-let set_undo _ = ()
-let get_undo _ = None
-
-
-let start_proof id str hyps c ?init_tac ?compute_guard hook =
+let start_proof (id : Id.t) str sigma hyps c ?init_tac terminator =
let goals = [ (Global.env_of_context hyps , c) ] in
- Proof_global.start_proof id str goals ?compute_guard hook;
- let tac = match init_tac with
- | Some tac -> Proofview.V82.tactic tac
- | None -> Proofview.tclUNIT ()
- in
- try Proof_global.run_tactic tac
- with reraise -> Proof_global.discard_current (); raise reraise
-
-let restart_proof () = undo_todepth 1
-
-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
-
+ Proof_global.start_proof sigma id str goals terminator;
+ let env = Global.env () in
+ ignore (Proof_global.with_current_proof (fun _ p ->
+ match init_tac with
+ | None -> p,(true,[])
+ | Some tac -> Proof.run_tactic env tac p))
+
+let cook_this_proof p =
+ match p with
+ | { Proof_global.id;entries=[constr];persistence;universes } ->
+ (id,(constr,universes,persistence))
+ | _ -> Errors.anomaly ~label:"Pfedit.cook_proof" (Pp.str "more than one proof term.")
+
+let cook_proof () =
+ cook_this_proof (fst
+ (Proof_global.close_proof ~keep_body_ucst_sepatate:false (fun x -> x)))
let get_pftreestate () =
Proof_global.give_me_the_proof ()
let set_end_tac tac =
- let tac = Proofview.V82.tactic tac in
Proof_global.set_endline_tactic tac
let set_used_variables l =
@@ -96,25 +56,25 @@ let get_used_variables () =
exception NoSuchGoal
let _ = Errors.register_handler begin function
- | NoSuchGoal -> Util.error "No such goal."
+ | NoSuchGoal -> Errors.error "No such goal."
| _ -> raise Errors.Unhandled
end
let get_nth_V82_goal i =
let p = Proof_global.give_me_the_proof () in
- let { it=goals ; sigma = sigma } = Proof.V82.subgoals p in
+ let { it=goals ; sigma = sigma; } = Proof.V82.subgoals p in
try
- { it=(List.nth goals (i-1)) ; sigma=sigma }
+ { it=(List.nth goals (i-1)) ; sigma=sigma; }
with Failure _ -> raise NoSuchGoal
let get_goal_context_gen i =
try
- let { it=goal ; sigma=sigma } = get_nth_V82_goal i in
- (sigma, Refiner.pf_env { it=goal ; sigma=sigma })
- with Proof_global.NoCurrentProof -> Util.error "No focused proof."
+let { it=goal ; sigma=sigma; } = get_nth_V82_goal i in
+(sigma, Refiner.pf_env { it=goal ; sigma=sigma; })
+ with Proof_global.NoCurrentProof -> Errors.error "No focused proof."
let get_goal_context i =
try get_goal_context_gen i
- with NoSuchGoal -> Util.error "No such goal."
+ with NoSuchGoal -> Errors.error "No such goal."
let get_current_goal_context () =
try get_goal_context_gen 1
@@ -125,29 +85,44 @@ let get_current_goal_context () =
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"
+ | (id,([concl],strength)) -> id,strength,concl
+ | _ -> Errors.anomaly ~label:"Pfedit.current_proof_statement" (Pp.str "more than one statement")
-let solve_nth ?(with_end_tac=false) gi tac =
+let solve ?with_end_tac gi info_lvl tac pr =
try
- let tac = Proofview.V82.tactic tac in
- let tac = if with_end_tac then
- Proof_global.with_end_tac tac
- else
- tac
+ let tac = match with_end_tac with
+ | None -> tac
+ | Some etac -> Proofview.tclTHEN tac etac in
+ let tac = match info_lvl with
+ | None -> tac
+ | Some _ -> Proofview.Trace.record_info_trace 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 tac = match gi with
+ | Vernacexpr.SelectNth i -> Proofview.tclFOCUS i i tac
+ | Vernacexpr.SelectId id -> Proofview.tclFOCUSID id tac
+ | Vernacexpr.SelectAll -> tac
+ | Vernacexpr.SelectAllParallel ->
+ Errors.anomaly(str"SelectAllParallel not handled by Stm")
+ in
+ let (p,(status,info)) = Proof.run_tactic (Global.env ()) tac pr in
+ let () =
+ match info_lvl with
+ | None -> ()
+ | Some i -> Pp.ppnl (hov 0 (Proofview.Trace.pr_info ~lvl:i info))
+ in
+ (p,status)
+ with
+ | Proof_global.NoCurrentProof -> Errors.error "No focused proof"
+ | CList.IndexOutOfRange ->
+ match gi with
+ | Vernacexpr.SelectNth i -> let msg = str "No such goal: " ++ int i ++ str "." in
+ Errors.errorlabstrm "" msg
+ | _ -> assert false
-let by = solve_nth 1
+let by tac = Proof_global.with_current_proof (fun _ -> solve (Vernacexpr.SelectNth 1) None tac)
let instantiate_nth_evar_com n com =
- let pf = Proof_global.give_me_the_proof () in
- Proof.V82.instantiate_evar n com pf
+ Proof_global.simple_with_current_proof (fun _ p -> Proof.V82.instantiate_evar n com p)
(**********************************************************************)
@@ -157,21 +132,29 @@ open Decl_kinds
let next = let n = ref 0 in fun () -> incr n; !n
-let build_constant_by_tactic id sign typ tac =
- start_proof id (Global,Proof Theorem) sign typ (fun _ _ -> ());
+let build_constant_by_tactic id ctx sign ?(goal_kind = Global, false, Proof Theorem) typ tac =
+ let evd = Evd.from_env ~ctx Environ.empty_env in
+ start_proof id goal_kind evd sign typ (fun _ -> ());
try
- by tac;
- let _,(const,_,_,_) = cook_proof (fun _ -> ()) in
+ let status = by tac in
+ let _,(const,univs,_) = cook_proof () in
delete_current_proof ();
- const
+ const, status, univs
with reraise ->
+ let reraise = Errors.push reraise in
delete_current_proof ();
- raise reraise
+ iraise reraise
-let build_by_tactic env typ tac =
- let id = id_of_string ("temporary_proof"^string_of_int (next())) in
+let build_by_tactic env ctx ?(poly=false) typ tac =
+ let id = Id.of_string ("temporary_proof"^string_of_int (next())) in
let sign = val_of_named_context (named_context env) in
- (build_constant_by_tactic id sign typ tac).const_entry_body
+ let gk = Global, poly, Proof Theorem in
+ let ce, status, univs = build_constant_by_tactic id ctx sign ~goal_kind:gk typ tac in
+ let ce = Term_typing.handle_entry_side_effects env ce in
+ let (cb, ctx), se = Future.force ce.const_entry_body in
+ assert(Declareops.side_effects_is_empty se);
+ assert(Univ.ContextSet.is_empty ctx);
+ cb, status, univs
(**********************************************************************)
(* Support for resolution of evars in tactic interpretation, including
@@ -181,13 +164,19 @@ let implicit_tactic = ref None
let declare_implicit_tactic tac = implicit_tactic := Some tac
-let solve_by_implicit_tactic env sigma (evk,args) =
+let clear_implicit_tactic () = implicit_tactic := None
+
+let solve_by_implicit_tactic env sigma evk =
let evi = Evd.find_undefined sigma evk in
match (!implicit_tactic, snd (evar_source evk sigma)) with
- | Some tac, (ImplicitArg _ | QuestionMark _)
+ | Some tac, (Evar_kinds.ImplicitArg _ | Evar_kinds.QuestionMark _)
when
- Sign.named_context_equal (Environ.named_context_of_val evi.evar_hyps)
+ Context.named_context_equal (Environ.named_context_of_val evi.evar_hyps)
(Environ.named_context env) ->
- (try build_by_tactic env evi.evar_concl (tclCOMPLETE tac)
+ let tac = Proofview.tclTHEN tac (Proofview.tclEXTEND [] (Proofview.tclZERO (Errors.UserError ("",Pp.str"Proof is not complete."))) []) in
+ (try
+ let (ans, _, _) =
+ build_by_tactic env (Evd.evar_universe_context sigma) evi.evar_concl tac in
+ ans
with e when Logic.catchable_exception e -> raise Exit)
| _ -> raise Exit
diff --git a/proofs/pfedit.mli b/proofs/pfedit.mli
index ec083e41..edbc18a3 100644
--- a/proofs/pfedit.mli
+++ b/proofs/pfedit.mli
@@ -1,20 +1,16 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Util
-open Pp
+open Loc
open Names
open Term
-open Sign
open Environ
open Decl_kinds
-open Tacmach
-open Tacexpr
(** Several proofs can be opened simultaneously but at most one is
focused at some time. The following functions work by side-effect
@@ -38,7 +34,7 @@ val check_no_pending_proofs : unit -> unit
(** [delete_proof name] deletes proof of name [name] or fails if no proof
has this name *)
-val delete_proof : identifier located -> unit
+val delete_proof : Id.t located -> unit
(** [delete_current_proof ()] deletes current focused proof or fails if
no proof is focused *)
@@ -50,21 +46,6 @@ val delete_current_proof : unit -> unit
val delete_all_proofs : unit -> unit
(** {6 ... } *)
-(** [undo n] undoes the effect of the last [n] tactics applied to the
- current proof; it fails if no proof is focused or if the ``undo''
- stack is exhausted *)
-
-val undo : 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
-
-(** {6 ... } *)
(** [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
@@ -75,14 +56,9 @@ val current_proof_depth: unit -> int
type lemma_possible_guards = Proof_global.lemma_possible_guards
val start_proof :
- identifier -> goal_kind -> named_context_val -> constr ->
- ?init_tac:tactic -> ?compute_guard:lemma_possible_guards ->
- declaration_hook -> unit
-
-(** [restart_proof ()] restarts the current focused proof from the beginning
- or fails if no proof is focused *)
-
-val restart_proof : unit -> unit
+ Id.t -> goal_kind -> Evd.evar_map -> named_context_val -> constr ->
+ ?init_tac:unit Proofview.tactic ->
+ Proof_global.proof_terminator -> unit
(** {6 ... } *)
(** [cook_proof opacity] turns the current proof (assumed completed) into
@@ -90,17 +66,18 @@ val restart_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 : (Proof.proof -> unit) ->
- identifier *
- (Entries.definition_entry * lemma_possible_guards * goal_kind *
- declaration_hook)
+val cook_this_proof :
+ Proof_global.proof_object ->
+ (Id.t *
+ (Entries.definition_entry * Proof_global.proof_universes * goal_kind))
-(** To export completed proofs to xml *)
-val set_xml_cook_proof : (goal_kind * Proof.proof -> unit) -> unit
+val cook_proof : unit ->
+ (Id.t *
+ (Entries.definition_entry * Proof_global.proof_universes * goal_kind))
(** {6 ... } *)
-(** [get_Proof.proof ()] returns the current focused pending proof or
- raises [UserError "no focused proof"] *)
+(** [get_pftreestate ()] returns the current focused pending proof.
+ @raise NoCurrentProof if there is no pending proof. *)
val get_pftreestate : unit -> Proof.proof
@@ -117,61 +94,76 @@ val get_current_goal_context : unit -> Evd.evar_map * env
(** [current_proof_statement] *)
val current_proof_statement :
- unit -> identifier * goal_kind * types * declaration_hook
+ unit -> Id.t * goal_kind * types
(** {6 ... } *)
(** [get_current_proof_name ()] return the name of the current focused
proof or failed if no proof is focused *)
-val get_current_proof_name : unit -> identifier
+val get_current_proof_name : unit -> Id.t
(** [get_all_proof_names ()] returns the list of all pending proof names.
The first name is the current proof, the other names may come in
any order. *)
-val get_all_proof_names : unit -> identifier list
+val get_all_proof_names : unit -> Id.t list
(** {6 ... } *)
(** [set_end_tac tac] applies tactic [tac] to all subgoal generate
- by [solve_nth] *)
+ by [solve] *)
-val set_end_tac : tactic -> unit
+val set_end_tac : Tacexpr.raw_tactic_expr -> unit
(** {6 ... } *)
(** [set_used_variables l] declares that section variables [l] will be
used in the proof *)
-val set_used_variables : identifier list -> unit
-val get_used_variables : unit -> Sign.section_context option
+val set_used_variables : Id.t list -> Context.section_context
+val get_used_variables : unit -> Context.section_context option
(** {6 ... } *)
-(** [solve_nth n tac] applies tactic [tac] to the [n]th subgoal of the
- current focused proof or raises a UserError if no proof is focused or
- if there is no [n]th subgoal *)
+(** [solve (SelectNth n) tac] applies tactic [tac] to the [n]th
+ subgoal of the current focused proof or raises a [UserError] if no
+ proof is focused or if there is no [n]th subgoal. [solve SelectAll
+ tac] applies [tac] to all subgoals. *)
-val solve_nth : ?with_end_tac:bool -> int -> tactic -> unit
+val solve : ?with_end_tac:unit Proofview.tactic ->
+ Vernacexpr.goal_selector -> int option -> unit Proofview.tactic ->
+ Proof.proof -> Proof.proof*bool
(** [by tac] applies tactic [tac] to the 1st subgoal of the current
- focused proof or raises a UserError if there is no focused proof or
- if there is no more subgoals *)
+ focused proof or raises a UserError if there is no focused proof or
+ if there is no more subgoals.
+ Returns [false] if an unsafe tactic has been used. *)
-val by : tactic -> unit
+val by : unit Proofview.tactic -> bool
(** [instantiate_nth_evar_com n c] instantiate the [n]th undefined
existential variable of the current focused proof by [c] or raises a
UserError if no proof is focused or if there is no such [n]th
existential variable *)
-val instantiate_nth_evar_com : int -> Topconstr.constr_expr -> unit
+val instantiate_nth_evar_com : int -> Constrexpr.constr_expr -> unit
-(** [build_by_tactic typ tac] returns a term of type [typ] by calling [tac] *)
+(** [build_by_tactic typ tac] returns a term of type [typ] by calling
+ [tac]. The return boolean, if [false] indicates the use of an unsafe
+ tactic. *)
-val build_constant_by_tactic : identifier -> named_context_val -> types -> tactic ->
- Entries.definition_entry
-val build_by_tactic : env -> types -> tactic -> constr
+val build_constant_by_tactic :
+ Id.t -> Evd.evar_universe_context -> named_context_val -> ?goal_kind:goal_kind ->
+ types -> unit Proofview.tactic ->
+ Entries.definition_entry * bool * Evd.evar_universe_context
+
+val build_by_tactic : env -> Evd.evar_universe_context -> ?poly:polymorphic ->
+ types -> unit Proofview.tactic ->
+ constr * bool * Evd.evar_universe_context
(** Declare the default tactic to fill implicit arguments *)
-val declare_implicit_tactic : tactic -> unit
+val declare_implicit_tactic : unit Proofview.tactic -> unit
+
+(** To remove the default tactic *)
+val clear_implicit_tactic : unit -> unit
(* Raise Exit if cannot solve *)
-val solve_by_implicit_tactic : env -> Evd.evar_map -> existential -> constr
+(* FIXME: interface: it may incur some new universes etc... *)
+val solve_by_implicit_tactic : env -> Evd.evar_map -> Evd.evar -> constr
diff --git a/proofs/proof.ml b/proofs/proof.ml
index bd185b99..828f9fa7 100644
--- a/proofs/proof.ml
+++ b/proofs/proof.ml
@@ -1,16 +1,15 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
(* 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.
+ A proof deals with the focusing commands (including the braces and bullets),
+ the shelf (see the [shelve] tactic) and given up goal (see the [give_up]
+ tactic). A proof is made of the following:
- 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
@@ -23,24 +22,27 @@
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.
+ - Shelf: A list of goals which have been put aside during the proof. They can be
+ retrieved with the [Unshelve] command, or solved by side effects
+ - Given up goals: as long as there is a given up goal, the proof is not completed.
+ Given up goals cannot be retrieved, the user must go back where the tactic
+ [give_up] was run and solve the goal there.
*)
-open Term
+open Util
type _focus_kind = int
type 'a focus_kind = _focus_kind
type focus_info = Obj.t
+type reason = NotThisWay | AlreadyNoFocus
type unfocusable =
- | Cannot of exn
+ | Cannot of reason
| Loose
| Strict
-type _focus_condition =
- (_focus_kind -> Proofview.proofview -> unfocusable) *
- (_focus_kind -> bool)
+type _focus_condition =
+ | CondNo of bool * _focus_kind
+ | CondDone of bool * _focus_kind
+ | CondEndStack of _focus_kind (* loose_end is false here *)
type 'a focus_condition = _focus_condition
let next_kind = ref 0
@@ -49,107 +51,71 @@ let new_focus_kind () =
incr next_kind;
r
-(* Auxiliary function to define conditions. *)
-let check kind1 kind2 = kind1=kind2
-
(* 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]. *)
-
-module Cond = struct
- (* first attempt at an algebra of condition *)
- (* semantics:
- - [Cannot] means that the condition is not met
- - [Strict] that the condition is met
- - [Loose] that the condition is not quite met
- but authorises to unfocus provided a condition
- of a previous focus on the stack is (strictly)
- met. [Loose] focuses are those, like bullets,
- which do not have a closing command and
- are hence closed by unfocusing actions unrelated
- to their focus_kind.
- *)
- let bool e b =
- if b then fun _ _ -> Strict
- else fun _ _ -> Cannot e
- let loose c k p = match c k p with
- | Cannot _ -> Loose
- | c -> c
- let cloose l c =
- if l then loose c
- else c
- let (&&&) c1 c2 k p=
- match c1 k p , c2 k p with
- | Cannot e , _
- | _ , Cannot e -> Cannot e
- | Strict, Strict -> Strict
- | _ , _ -> Loose
- let kind e k0 k p = bool e (k0=k) k p
- let pdone e k p = bool e (Proofview.finished p) k p
-end
-
-(* Unfocus command.
- Fails if the proof is not focused. *)
exception CannotUnfocusThisWay
+
+(* Cannot focus on non-existing subgoals *)
+exception NoSuchGoals of int * int
+
+
+exception FullyUnfocused
+
let _ = Errors.register_handler begin function
| CannotUnfocusThisWay ->
- Util.error "This proof is focused, but cannot be unfocused this way"
+ Errors.error "This proof is focused, but cannot be unfocused this way"
+ | NoSuchGoals (i,j) when Int.equal i j ->
+ Errors.errorlabstrm "Focus" Pp.(str"No such goal (" ++ int i ++ str").")
+ | NoSuchGoals (i,j) ->
+ Errors.errorlabstrm "Focus" Pp.(
+ str"Not every goal in range ["++ int i ++ str","++int j++str"] exist."
+ )
+ | FullyUnfocused -> Errors.error "The proof is not focused"
| _ -> raise Errors.Unhandled
end
-open Cond
-let no_cond_gen e ~loose_end k0 =
- cloose loose_end (kind e k0)
-let no_cond_gen e ?(loose_end=false) k = no_cond_gen e ~loose_end k , check k
-let no_cond ?loose_end = no_cond_gen CannotUnfocusThisWay ?loose_end
-(* [done_cond] checks that the unfocusing command uses the right [focus_kind]
- and that the focused proofview is complete. *)
-let done_cond_gen e ~loose_end k0 =
- (cloose loose_end (kind e k0)) &&& pdone e
-let done_cond_gen e ?(loose_end=false) k = done_cond_gen e ~loose_end k , check k
-let done_cond ?loose_end = done_cond_gen CannotUnfocusThisWay ?loose_end
-
+let check_cond_kind c k =
+ let kind_of_cond = function
+ | CondNo (_,k) | CondDone(_,k) | CondEndStack k -> k in
+ Int.equal (kind_of_cond c) k
+
+let test_cond c k1 pw =
+ match c with
+ | CondNo(_, k) when Int.equal k k1 -> Strict
+ | CondNo(true, _) -> Loose
+ | CondNo(false, _) -> Cannot NotThisWay
+ | CondDone(_, k) when Int.equal k k1 && Proofview.finished pw -> Strict
+ | CondDone(true, _) -> Loose
+ | CondDone(false, _) -> Cannot NotThisWay
+ | CondEndStack k when Int.equal k k1 -> Strict
+ | CondEndStack _ -> Cannot AlreadyNoFocus
+
+let no_cond ?(loose_end=false) k = CondNo (loose_end, k)
+let done_cond ?(loose_end=false) k = CondDone (loose_end,k)
(* Subpart of the type of proofs. It contains the parts of the proof which
are under control of the undo mechanism *)
-type proof_state = {
+type proof = {
(* Current focused proofview *)
proofview: Proofview.proofview;
+ (* Entry for the proofview *)
+ entry : Proofview.entry;
(* History of the focusings, provides information on how
to unfocus the proof and the extra information stored while focusing.
The list is empty when the proof is fully unfocused. *)
focus_stack: (_focus_condition*focus_info*Proofview.focus_context) list;
- (* Extra information which can be freely used to create new behaviours. *)
- intel: Store.t
+ (* List of goals that have been shelved. *)
+ shelf : Goal.goal list;
+ (* List of goals that have been given up *)
+ given_up : Goal.goal list;
}
-type proof_info = {
- mutable endline_tactic : unit Proofview.tactic ;
- mutable section_vars : Sign.section_context option;
- 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;
- (* secondary undo stacks used for transactions *)
- mutable transactions : undo_action list list;
- info : proof_info
- }
-
-
(*** General proof functions ***)
-let proof { state = p } =
+let proof p =
let (goals,sigma) = Proofview.proofview p.proofview in
(* spiwack: beware, the bottom of the stack is used by [Proof]
internally, and should not be exposed. *)
@@ -161,322 +127,261 @@ let proof { state = p } =
let stack =
map_minus_one (fun (_,_,c) -> Proofview.focus_context c) p.focus_stack
in
- (goals,stack,sigma)
+ let shelf = p.shelf in
+ let given_up = p.given_up in
+ (goals,stack,shelf,given_up,sigma)
+
+type 'a pre_goals = {
+ fg_goals : 'a list;
+ (** List of the focussed goals *)
+ bg_goals : ('a list * 'a list) list;
+ (** Zipper representing the unfocussed background goals *)
+ shelved_goals : 'a list;
+ (** List of the goals on the shelf. *)
+ given_up_goals : 'a list;
+ (** List of the goals that have been given up *)
+}
+
+let map_structured_proof pfts process_goal: 'a pre_goals =
+ let (goals, zipper, shelf, given_up, sigma) = proof pfts in
+ let fg = List.map (process_goal sigma) goals in
+ let map_zip (lg, rg) =
+ let lg = List.map (process_goal sigma) lg in
+ let rg = List.map (process_goal sigma) rg in
+ (lg, rg)
+ in
+ let bg = List.map map_zip zipper in
+ let shelf = List.map (process_goal sigma) shelf in
+ let given_up = List.map (process_goal sigma) given_up in
+ { fg_goals = fg;
+ bg_goals = bg;
+ shelved_goals = shelf;
+ given_up_goals = given_up; }
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.
Unfocusing is handled in {!return}. *)
let is_done p =
- Proofview.finished p.state.proofview &&
- Proofview.finished (unroll_focus p.state.proofview p.state.focus_stack)
+ Proofview.finished p.proofview &&
+ Proofview.finished (unroll_focus p.proofview p.focus_stack)
(* spiwack: for compatibility with <= 8.2 proof engine *)
let has_unresolved_evar p =
- Proofview.V82.has_unresolved_evar p.state.proofview
+ Proofview.V82.has_unresolved_evar p.proofview
(* Returns the list of partial proofs to initial goals *)
-let partial_proof p =
- List.map fst (Proofview.return p.state.proofview)
-
-
+let partial_proof p = Proofview.partial_proof p.entry p.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 cond inf context pr =
- pr.state <- { pr.state with focus_stack = (cond,inf,context)::pr.state.focus_stack }
+ { pr with focus_stack = (cond,inf,context)::pr.focus_stack }
-exception FullyUnfocused
-let _ = Errors.register_handler begin function
- | FullyUnfocused -> Util.error "The proof is not focused"
- | _ -> raise Errors.Unhandled
-end
(* An auxiliary function to read the kind of the next focusing step *)
let cond_of_focus pr =
- match pr.state.focus_stack with
+ match pr.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
- | _ ->
+ match pr.focus_stack with
+ | focus::other_focuses ->
+ { pr with focus_stack = other_focuses }, focus
+ | _ ->
raise FullyUnfocused
-(* Auxiliary function to push a [proof_state] onto the undo stack. *)
-let push_undo save pr =
- match pr.transactions with
- | [] -> pr.undo_stack <- save::pr.undo_stack
- | stack::trans' -> pr.transactions <- (save::stack)::trans'
-
-(* Auxiliary function to pop and read a [save_state] from the undo stack. *)
-exception EmptyUndoStack
-let _ = Errors.register_handler begin function
- | EmptyUndoStack -> Util.error "Cannot undo: no more undo information"
- | _ -> raise Errors.Unhandled
-end
-let pop_undo pr =
- match pr.transactions , pr.undo_stack with
- | [] , state::stack -> pr.undo_stack <- stack; state
- | (state::stack)::trans', _ -> pr.transactions <- stack::trans'; state
- | _ -> raise EmptyUndoStack
-
-
(* This function focuses the proof [pr] between indices [i] and [j] *)
let _focus cond inf i j pr =
- let (focused,context) = Proofview.focus i j pr.state.proofview in
- push_focus cond inf context pr;
- pr.state <- { pr.state with proofview = focused }
+ let focused, context = Proofview.focus i j pr.proofview in
+ let pr = push_focus cond inf context pr in
+ { pr with proofview = focused }
(* This function unfocuses the proof [pr], it raises [FullyUnfocused],
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 }
-
-
-let set_used_variables l p =
- p.info.section_vars <- Some l
-
-let get_used_variables p = p.info.section_vars
-
-(*** 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
-
-
-
-(*** Transactions ***)
-
-let init_transaction pr =
- pr.transactions <- []::pr.transactions
-
-let commit_stack pr stack =
- let push s = push_undo s pr in
- List.iter push (List.rev stack)
-
-(* Invariant: [commit] should be called only when a transaction
- is open. It closes the current transaction. *)
-let commit pr =
- match pr.transactions with
- | stack::trans' ->
- pr.transactions <- trans';
- commit_stack pr stack
- | [] -> assert false
-
-(* Invariant: [rollback] should be called only when a transaction
- is open. It closes the current transaction. *)
-let rec rollback pr =
- try
- undo pr;
- rollback pr
- with EmptyUndoStack ->
- match pr.transactions with
- | []::trans' -> pr.transactions <- trans'
- | _ -> assert false
-
-
-let transaction pr t =
- init_transaction pr;
- try t (); commit pr
- with reraise -> rollback pr; raise reraise
-
+ let pr, (_,_,fc) = pop_focus pr in
+ { pr with proofview = Proofview.unfocus fc pr.proofview }
(* 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 inf i pr =
- save pr;
- _focus cond (Obj.repr inf) i i pr
+ try _focus cond (Obj.repr inf) i i pr
+ with CList.IndexOutOfRange -> raise (NoSuchGoals (i,i))
let rec unfocus kind pr () =
- let starting_point = save_state pr in
let cond = cond_of_focus pr in
- match fst cond kind pr.state.proofview with
- | Cannot e -> raise e
- | Strict ->
- (_unfocus pr;
- push_undo starting_point pr)
+ match test_cond cond kind pr.proofview with
+ | Cannot NotThisWay -> raise CannotUnfocusThisWay
+ | Cannot AlreadyNoFocus -> raise FullyUnfocused
+ | Strict ->
+ let pr = _unfocus pr in
+ pr
| Loose ->
begin try
- _unfocus pr;
- push_undo starting_point pr;
+ let pr = _unfocus pr in
unfocus kind pr ()
with FullyUnfocused -> raise CannotUnfocusThisWay
end
-let unfocus kind pr =
- transaction pr (unfocus kind pr)
-
exception NoSuchFocus
(* no handler: should not be allowed to reach toplevel. *)
let rec get_in_focus_stack kind stack =
match stack with
- | ((_,check),inf,_)::stack ->
- if check kind then inf
+ | (cond,inf,_)::stack ->
+ if check_cond_kind cond kind then inf
else get_in_focus_stack kind stack
| [] -> raise NoSuchFocus
let get_at_focus kind pr =
- Obj.magic (get_in_focus_stack kind pr.state.focus_stack)
+ Obj.magic (get_in_focus_stack kind pr.focus_stack)
let is_last_focus kind pr =
- let ((_,check),_,_) = List.hd pr.state.focus_stack in
- check kind
+ let (cond,_,_) = List.hd pr.focus_stack in
+ check_cond_kind cond kind
let no_focused_goal p =
- Proofview.finished p.state.proofview
+ Proofview.finished p.proofview
+
+let rec maximal_unfocus k p =
+ if no_focused_goal p then
+ try maximal_unfocus k (unfocus k p ())
+ with FullyUnfocused | CannotUnfocusThisWay -> p
+ else p
(*** Proof Creation/Termination ***)
(* [end_of_stack] is unfocused by return to close every loose focus. *)
let end_of_stack_kind = new_focus_kind ()
-let end_of_stack = done_cond_gen FullyUnfocused end_of_stack_kind
+let end_of_stack = CondEndStack end_of_stack_kind
let unfocused = is_last_focus end_of_stack_kind
-let start goals =
- let pr =
- { state = { proofview = Proofview.init goals ;
- focus_stack = [] ;
- intel = Store.empty} ;
- undo_stack = [] ;
- transactions = [] ;
- info = { endline_tactic = Proofview.tclUNIT ();
- initial_conclusions = List.map snd goals;
- section_vars = None }
- }
- in
- _focus end_of_stack (Obj.repr ()) 1 (List.length goals) pr;
- pr
+let start sigma goals =
+ let entry, proofview = Proofview.init sigma goals in
+ let pr = {
+ proofview;
+ entry;
+ focus_stack = [] ;
+ shelf = [] ;
+ given_up = [] } in
+ _focus end_of_stack (Obj.repr ()) 1 (List.length goals) pr
+let dependent_start goals =
+ let entry, proofview = Proofview.dependent_init goals in
+ let pr = {
+ proofview;
+ entry;
+ focus_stack = [] ;
+ shelf = [] ;
+ given_up = [] } in
+ let number_of_goals = List.length (Proofview.initial_goals pr.entry) in
+ _focus end_of_stack (Obj.repr ()) 1 number_of_goals pr
exception UnfinishedProof
+exception HasShelvedGoals
+exception HasGivenUpGoals
exception HasUnresolvedEvar
let _ = Errors.register_handler begin function
- | UnfinishedProof -> Util.error "Some goals have not been solved."
- | HasUnresolvedEvar -> Util.error "Some existential variables are uninstantiated."
+ | UnfinishedProof -> Errors.error "Some goals have not been solved."
+ | HasShelvedGoals -> Errors.error "Some goals have been left on the shelf."
+ | HasGivenUpGoals -> Errors.error "Some goals have been given up."
+ | HasUnresolvedEvar -> Errors.error "Some existential variables are uninstantiated."
| _ -> raise Errors.Unhandled
end
+
let return p =
if not (is_done p) then
raise UnfinishedProof
- else if has_unresolved_evar p then
+ else if not (CList.is_empty (p.shelf)) then
+ raise HasShelvedGoals
+ else if not (CList.is_empty (p.given_up)) then
+ raise HasGivenUpGoals
+ else if has_unresolved_evar p then
(* spiwack: for compatibility with <= 8.3 proof engine *)
raise HasUnresolvedEvar
else
- unfocus end_of_stack_kind p;
- Proofview.return p.state.proofview
-
-(*** Function manipulation proof extra informations ***)
+ let p = unfocus end_of_stack_kind p () in
+ Proofview.return p.proofview
-let get_proof_info pr =
- pr.state.intel
+let initial_goals p = Proofview.initial_goals p.entry
-let set_proof_info info pr =
- save pr;
- pr.state <- { pr.state with intel=info }
+let compact p =
+ let entry, proofview = Proofview.compact p.entry p.proofview in
+ { p with proofview; entry }
+(*** Function manipulation proof extra informations ***)
(*** 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 reraise ->
- restore_state starting_point pr;
- raise reraise
+ let sp = pr.proofview in
+ let (_,tacticced_proofview,(status,to_shelve,give_up),info_trace) =
+ Proofview.apply env tac sp
+ in
+ let sigma = Proofview.return tacticced_proofview in
+ (* Already solved goals are not to be counted as shelved. Nor are
+ they to be marked as unresolvable. *)
+ let undef l = List.filter (fun g -> Evd.is_undefined sigma g) l in
+ let retrieved = undef (List.rev (Evd.future_goals sigma)) in
+ let shelf = (undef pr.shelf)@retrieved@(undef to_shelve) in
+ let proofview =
+ List.fold_left
+ Proofview.Unsafe.mark_as_goal
+ tacticced_proofview
+ retrieved
+ in
+ let given_up = pr.given_up@give_up in
+ let proofview = Proofview.Unsafe.reset_future_goals proofview in
+ { pr with proofview ; shelf ; given_up },(status,info_trace)
(*** Commands ***)
-let in_proof p k =
- Proofview.in_proofview p.state.proofview k
+let in_proof p k = k (Proofview.return p.proofview)
+(* Remove all the goals from the shelf and adds them at the end of the
+ focused goals. *)
+let unshelve p =
+ { p with proofview = Proofview.unshelve (p.shelf) (p.proofview) ; shelf = [] }
(*** Compatibility layer with <=v8.2 ***)
module V82 = struct
let subgoals p =
- Proofview.V82.goals p.state.proofview
+ Proofview.V82.goals p.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
+ Proofview.V82.goals (unroll_focus p.proofview p.focus_stack)
- let top_goal p =
- let { Evd.it=gls ; sigma=sigma } =
- Proofview.V82.top_goals p.state.proofview
+ let top_goal p =
+ let { Evd.it=gls ; sigma=sigma; } =
+ Proofview.V82.top_goals p.entry p.proofview
in
- { Evd.it=List.hd gls ; sigma=sigma }
+ { Evd.it=List.hd gls ; sigma=sigma; }
let top_evars p =
- Proofview.V82.top_evars p.state.proofview
+ Proofview.V82.top_evars p.entry
let grab_evars p =
if not (is_done p) then
raise UnfinishedProof
else
- save p;
- p.state <- { p.state with proofview = Proofview.V82.grab p.state.proofview }
-
-
- 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 reraise ->
- restore_state starting_point pr;
- raise reraise
+ { p with proofview = Proofview.V82.grab p.proofview }
+
+
+ let instantiate_evar n com pr =
+ let sp = pr.proofview in
+ let proofview = Proofview.V82.instantiate_evar n com sp in
+ let shelf =
+ List.filter begin fun g ->
+ Evd.is_undefined (Proofview.return proofview) g
+ end pr.shelf
+ in
+ { pr with proofview ; shelf }
+
end
diff --git a/proofs/proof.mli b/proofs/proof.mli
index d6873790..4ae64ae6 100644
--- a/proofs/proof.mli
+++ b/proofs/proof.mli
@@ -1,16 +1,15 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
(* 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.
+ A proof deals with the focusing commands (including the braces and bullets),
+ the shelf (see the [shelve] tactic) and given up goal (see the [give_up]
+ tactic). A proof is made of the following:
- 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
@@ -23,14 +22,13 @@
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.
+ - Shelf: A list of goals which have been put aside during the proof. They can be
+ retrieved with the [Unshelve] command, or solved by side effects
+ - Given up goals: as long as there is a given up goal, the proof is not completed.
+ Given up goals cannot be retrieved, the user must go back where the tactic
+ [give_up] was run and solve the goal there.
*)
-open Term
-
(* Type of a proof. *)
type proof
@@ -40,13 +38,37 @@ type proof
functions to ide-s. This would be better than spawning a new nearly
identical function everytime. Hence the generic name. *)
(* In this version: returns the focused goals, a representation of the
- focus stack (the number of goals at each level) and the underlying
+ focus stack (the goals at each level), a representation of the
+ shelf (the list of goals on the shelf), a representation of the
+ given up goals (the list of the given up goals) and the underlying
evar_map *)
-val proof : proof -> Goal.goal list * (Goal.goal list * Goal.goal list) list * Evd.evar_map
+val proof : proof ->
+ Goal.goal list
+ * (Goal.goal list * Goal.goal list) list
+ * Goal.goal list
+ * Goal.goal list
+ * Evd.evar_map
+
+(* Generic records structured like the return type of proof *)
+type 'a pre_goals = {
+ fg_goals : 'a list;
+ (** List of the focussed goals *)
+ bg_goals : ('a list * 'a list) list;
+ (** Zipper representing the unfocussed background goals *)
+ shelved_goals : 'a list;
+ (** List of the goals on the shelf. *)
+ given_up_goals : 'a list;
+ (** List of the goals that have been given up *)
+}
+
+val map_structured_proof : proof -> (Evd.evar_map -> Goal.goal -> 'a) -> ('a pre_goals)
+
(*** General proof functions ***)
-val start : (Environ.env * Term.types) list -> proof
+val start : Evd.evar_map -> (Environ.env * Term.types) list -> proof
+val dependent_start : Proofview.telescope -> proof
+val initial_goals : proof -> (Term.constr * Term.types) list
(* Returns [true] if the considered proof is completed, that is if no goal remain
to be considered (this does not require that all evars have been solved). *)
@@ -55,25 +77,18 @@ val is_done : proof -> bool
(* Returns the list of partial proofs to initial goals. *)
val partial_proof : proof -> Term.constr list
+val compact : proof -> proof
+
(* Returns the proofs (with their type) of the initial goals.
Raises [UnfinishedProof] is some goals remain to be considered.
+ Raises [HasShelvedGoals] if some goals are left on the shelf.
+ Raises [HasGivenUpGoals] if some goals have been given up.
Raises [HasUnresolvedEvar] if some evars have been left undefined. *)
exception UnfinishedProof
+exception HasShelvedGoals
+exception HasGivenUpGoals
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.
- An undo effect is meant to undo an effect on a proof (a canonical example
- of which is {!Proofglobal.set_proof_mode} which changes the current parser for
- tactics). Make sure it will work even if the effects have been only partially
- applied at the time of failure. *)
-val add_undo : (unit -> unit) -> proof -> unit
+val return : proof -> Evd.evar_map
(*** Focusing actions ***)
@@ -113,15 +128,23 @@ val done_cond : ?loose_end:bool -> 'a focus_kind -> 'a focus_condition
(* 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? *)
-val focus : 'a focus_condition -> 'a -> int -> proof -> unit
+val focus : 'a focus_condition -> 'a -> int -> proof -> proof
exception FullyUnfocused
exception CannotUnfocusThisWay
+
+(* This is raised when trying to focus on non-existing subgoals. It is
+ handled by an error message but one may need to catched it and
+ settle a better error message in some case (suggesting a better
+ bullet for example), see proof_global.ml function Bullet.pop and
+ Bullet.push. *)
+exception NoSuchGoals of int * int
+
(* Unfocusing command.
Raises [FullyUnfocused] if the proof is not focused.
Raises [CannotUnfocusThisWay] if the proof the unfocusing condition
is not met. *)
-val unfocus : 'a focus_kind -> proof -> unit
+val unfocus : 'a focus_kind -> proof -> unit -> proof
(* [unfocused p] returns [true] when [p] is fully unfocused. *)
val unfocused : proof -> bool
@@ -138,43 +161,23 @@ val is_last_focus : 'a focus_kind -> proof -> bool
(* 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
-
-(* Sets the section variables assumed by the proof *)
-val set_used_variables : Sign.section_context -> proof -> unit
-val get_used_variables : proof -> Sign.section_context option
-
-(*** 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
-
-
-(*** Transactions ***)
-
-(* A transaction chains several commands into a single one. For instance,
- a focusing command and a tactic. Transactions are such that if
- any of the atomic action fails, the whole transaction fails.
-
- During a transaction, the visible undo stack is constituted only
- of the actions performed done during the transaction.
-
- [transaction p f] can be called on an [f] using, itself, [transaction p].*)
-val transaction : proof -> (unit -> unit) -> unit
+(* the returned boolean signal whether an unsafe tactic has been
+ used. In which case it is [false]. *)
+val run_tactic : Environ.env ->
+ unit Proofview.tactic -> proof -> proof*(bool*Proofview_monad.Info.tree)
+val maximal_unfocus : 'a focus_kind -> proof -> proof
(*** Commands ***)
val in_proof : proof -> (Evd.evar_map -> 'a) -> 'a
+(* Remove all the goals from the shelf and adds them at the end of the
+ focused goals. *)
+val unshelve : proof -> proof
+
(*** Compatibility layer with <=v8.2 ***)
module V82 : sig
val subgoals : proof -> Goal.goal list Evd.sigma
@@ -182,19 +185,15 @@ module V82 : sig
(* 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
-
+
(* returns the existential variable used to start the proof *)
val top_evars : proof -> Evd.evar list
(* Turns the unresolved evars into goals.
Raises [UnfinishedProof] if there are still unsolved goals. *)
- val grab_evars : proof -> unit
+ val grab_evars : proof -> proof
(* Implements the Existential command *)
- val instantiate_evar : int -> Topconstr.constr_expr -> proof -> unit
+ val instantiate_evar : int -> Constrexpr.constr_expr -> proof -> proof
end
diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml
index a58fab0c..f55ab700 100644
--- a/proofs/proof_global.ml
+++ b/proofs/proof_global.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -14,6 +14,7 @@
(* *)
(***********************************************************************)
+open Util
open Pp
open Names
@@ -32,15 +33,17 @@ let proof_modes = Hashtbl.create 6
let find_proof_mode n =
try Hashtbl.find proof_modes n
with Not_found ->
- Util.error (Format.sprintf "No proof mode named \"%s\"." n)
+ Errors.error (Format.sprintf "No proof mode named \"%s\"." n)
+
+let register_proof_mode ({name = n} as m) =
+ Hashtbl.add proof_modes n (Ephemeron.create m)
-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 default_proof_mode = ref (find_proof_mode "No")
let _ =
Goptions.declare_string_option {Goptions.
@@ -48,48 +51,63 @@ let _ =
optdepr = false;
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 := find_proof_mode n
- end
+ optread = begin fun () ->
+ (Ephemeron.default !default_proof_mode standard).name
+ end;
+ optwrite = begin fun n ->
+ default_proof_mode := find_proof_mode 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
+type proof_universes = Evd.evar_universe_context
+
+type proof_object = {
+ id : Names.Id.t;
+ entries : Entries.definition_entry list;
+ persistence : Decl_kinds.goal_kind;
+ universes: proof_universes;
+ (* constraints : Univ.constraints; *)
+}
+
+type proof_ending =
+ | Admitted
+ | Proved of Vernacexpr.opacity_flag *
+ (Vernacexpr.lident * Decl_kinds.theorem_kind option) option *
+ proof_object
+type proof_terminator = proof_ending -> unit
+type closed_proof = proof_object * proof_terminator
+
+type pstate = {
+ pid : Id.t;
+ terminator : proof_terminator Ephemeron.key;
+ endline_tactic : Tacexpr.raw_tactic_expr option;
+ section_vars : Context.section_context option;
+ proof : Proof.proof;
+ strength : Decl_kinds.goal_kind;
+ mode : proof_mode Ephemeron.key;
}
-(* Invariant: the domain of proof_info is current_proof.*)
-(* The head of [!current_proof] is the actual current proof, the other ones are
+(* The head of [!pstates] is the actual current proof, the other ones are
to be resumed when the current proof is closed or aborted. *)
-let current_proof = ref ([]:nproof list)
-let proof_info = ref (Idmap.empty : proof_info Idmap.t)
+let pstates = ref ([] : pstate list)
(* 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 ();
+let update_proof_mode () =
+ match !pstates with
+ | { mode = m } :: _ ->
+ Ephemeron.iter_opt !current_proof_mode (fun x -> x.reset ());
current_proof_mode := m;
- !current_proof_mode.set ()
- | _ ->
- !current_proof_mode.reset ();
- current_proof_mode := standard
+ Ephemeron.iter_opt !current_proof_mode (fun x -> x.set ())
+ | _ ->
+ Ephemeron.iter_opt !current_proof_mode (fun x -> x.reset ());
+ current_proof_mode := find_proof_mode "No"
(* combinators for the current_proof lists *)
let push a l = l := a::!l;
@@ -97,213 +115,271 @@ let push a l = l := a::!l;
exception NoSuchProof
let _ = Errors.register_handler begin function
- | NoSuchProof -> Util.error "No such proof."
+ | NoSuchProof -> Errors.error "No such proof."
| _ -> raise Errors.Unhandled
end
-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 _ = Errors.register_handler begin function
- | NoCurrentProof -> Util.error "No focused proof (No proof-editing in progress)."
+ | NoCurrentProof -> Errors.error "No focused proof (No proof-editing in progress)."
| _ -> raise Errors.Unhandled
end
-let extract_top l =
- match !l with
- | np::l' -> l := l' ; update_proof_mode (); np
- | [] -> raise NoCurrentProof
-let find_top l =
- match !l with
+
+(*** Proof Global manipulation ***)
+
+let get_all_proof_names () =
+ List.map (function { pid = id } -> id) !pstates
+
+let cur_pstate () =
+ match !pstates with
| np::_ -> np
| [] -> raise NoCurrentProof
-let rotate_top l1 l2 =
- let np = extract_top l1 in
- push np l2
+let give_me_the_proof () = (cur_pstate ()).proof
+let get_current_proof_name () = (cur_pstate ()).pid
-let rotate_find id l1 l2 =
- let np = extract id l1 in
- push np l2
-
+let interp_tac = ref (fun _ -> assert false)
+let set_interp_tac f = interp_tac := f
-(* 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 with_current_proof f =
+ match !pstates with
+ | [] -> raise NoCurrentProof
+ | p :: rest ->
+ let et =
+ match p.endline_tactic with
+ | None -> Proofview.tclUNIT ()
+ | Some tac -> !interp_tac tac in
+ let (newpr,ret) = f et p.proof in
+ let p = { p with proof = newpr } in
+ pstates := p :: rest;
+ ret
+let simple_with_current_proof f = with_current_proof (fun t p -> f t p , ())
-let get_all_proof_names () =
- List.map fst !current_proof
+let compact_the_proof () = simple_with_current_proof (fun _ -> Proof.compact)
-let give_me_the_proof () =
- snd (find_top current_proof)
-let get_current_proof_name () =
- fst (find_top current_proof)
+(* Sets the tactic to be used when a tactic line is closed with [...] *)
+let set_endline_tactic tac =
+ match !pstates with
+ | [] -> raise NoCurrentProof
+ | p :: rest -> pstates := { p with endline_tactic = Some tac } :: rest
(* spiwack: it might be considered to move error messages away.
Or else to remove special exceptions from Proof_global.
Arguments for the former: there is no reason Proof_global is only
- accessed directly through vernacular commands. Error message should be
+ accessed directly through vernacular commands. Error message should be
pushed to external layers, and so we should be able to have a finer
control on error message on complex actions. *)
let msg_proofs () =
match get_all_proof_names () with
| [] -> (spc () ++ str"(No proof-editing in progress).")
| l -> (str"." ++ fnl () ++ str"Proofs currently edited:" ++ spc () ++
- (Util.prlist_with_sep Util.pr_spc Nameops.pr_id l)++ str ".")
+ (pr_sequence Nameops.pr_id l) ++ str".")
-let there_is_a_proof () = !current_proof <> []
+let there_is_a_proof () = not (List.is_empty !pstates)
let there_are_pending_proofs () = there_is_a_proof ()
let check_no_pending_proof () =
if not (there_are_pending_proofs ()) then
()
else begin
- Util.error (Pp.string_of_ppcmds
+ Errors.error (Pp.string_of_ppcmds
(str"Proof editing in progress" ++ msg_proofs () ++ fnl() ++
str"Use \"Abort All\" first or complete proof(s)."))
end
let discard_gen id =
- ignore (extract id current_proof);
- remove id proof_info
-
-let discard (loc,id) =
- try
- discard_gen id
- with NoSuchProof ->
- Util.user_err_loc
+ pstates := List.filter (fun { pid = id' } -> not (Id.equal id id')) !pstates
+
+let discard (loc,id) =
+ let n = List.length !pstates in
+ discard_gen id;
+ if Int.equal (List.length !pstates) n then
+ Errors.user_err_loc
(loc,"Pfedit.delete_proof",str"No such proof" ++ msg_proofs ())
let discard_current () =
- let (id,_) = extract_top current_proof in
- remove id proof_info
-
-let discard_all () =
- current_proof := [];
- proof_info := Idmap.empty
+ if List.is_empty !pstates then raise NoCurrentProof else pstates := List.tl !pstates
+
+let discard_all () = pstates := []
(* [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;
+let set_proof_mode m id =
+ pstates :=
+ List.map (function { pid = id' } as p ->
+ if Id.equal id' id then { p with mode = m } else p) !pstates;
update_proof_mode ()
-(* Complete function.
- Handles undo.
- Applies to current proof, and proof mode name [mn]. *)
-let set_proof_mode mn =
- let m = find_proof_mode 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
-
-exception AlreadyExists
-let _ = Errors.register_handler begin function
- | AlreadyExists -> Util.error "Already editing something of that name."
- | _ -> raise Errors.Unhandled
-end
-(* [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 =
- begin
- List.iter begin fun (id_ex,_) ->
- if Names.id_ord id id_ex = 0 then raise AlreadyExists
- end !current_proof
- end;
- 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 set_proof_mode mn =
+ set_proof_mode (find_proof_mode mn) (get_current_proof_name ())
+
+let activate_proof_mode mode =
+ Ephemeron.iter_opt (find_proof_mode mode) (fun x -> x.set ())
+let disactivate_proof_mode mode =
+ Ephemeron.iter_opt (find_proof_mode mode) (fun x -> x.reset ())
+
+(** [start_proof sigma id str goals terminator] starts a proof of name
+ [id] with goals [goals] (a list of pairs of environment and
+ conclusion); [str] describes what kind of theorem/definition this
+ is (spiwack: for potential printing, I believe is used only by
+ closing commands and the xml plugin); [terminator] is used at the
+ end of the proof to close the proof. The proof is started in the
+ evar map [sigma] (which can typically contain universe
+ constraints). *)
+let start_proof sigma id str goals terminator =
+ let initial_state = {
+ pid = id;
+ terminator = Ephemeron.create terminator;
+ proof = Proof.start sigma goals;
+ endline_tactic = None;
+ section_vars = None;
+ strength = str;
+ mode = find_proof_mode "No" } in
+ push initial_state pstates
+
+let start_dependent_proof id str goals terminator =
+ let initial_state = {
+ pid = id;
+ terminator = Ephemeron.create terminator;
+ proof = Proof.dependent_start goals;
+ endline_tactic = None;
+ section_vars = None;
+ strength = str;
+ mode = find_proof_mode "No" } in
+ push initial_state pstates
+
+let get_used_variables () = (cur_pstate ()).section_vars
let set_used_variables l =
- let p = give_me_the_proof () in
let env = Global.env () in
- let ids = List.fold_right Idset.add l Idset.empty in
+ let ids = List.fold_right Id.Set.add l Id.Set.empty in
let ctx = Environ.keep_hyps env ids in
- Proof.set_used_variables ctx p
-
-let get_used_variables () =
- Proof.get_used_variables (give_me_the_proof ())
-
-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 section_vars = Proof.get_used_variables p in
- let entries = List.map
- (fun (c,t) -> { Entries.const_entry_body = c;
- const_entry_secctx = section_vars;
- const_entry_type = Some t;
- const_entry_opaque = true })
- 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"
-
+ match !pstates with
+ | [] -> raise NoCurrentProof
+ | p :: rest ->
+ if not (Option.is_empty p.section_vars) then
+ Errors.error "Used section variables can be declared only once";
+ pstates := { p with section_vars = Some ctx} :: rest;
+ ctx
+
+let get_open_goals () =
+ let gl, gll, shelf , _ , _ = Proof.proof (cur_pstate ()).proof in
+ List.length gl +
+ List.fold_left (+) 0
+ (List.map (fun (l1,l2) -> List.length l1 + List.length l2) gll) +
+ List.length shelf
+
+let close_proof ~keep_body_ucst_sepatate ?feedback_id ~now fpl =
+ let { pid; section_vars; strength; proof; terminator } = cur_pstate () in
+ let poly = pi2 strength (* Polymorphic *) in
+ let initial_goals = Proof.initial_goals proof in
+ let fpl, univs = Future.split2 fpl in
+ let universes =
+ if poly || now then Future.force univs
+ else Proof.in_proof proof (fun x -> Evd.evar_universe_context x)
+ in
+ (* Because of dependent subgoals at the begining of proofs, we could
+ have existential variables in the initial types of goals, we need to
+ normalise them for the kernel. *)
+ let subst_evar k =
+ Proof.in_proof proof (fun m -> Evd.existential_opt_value m k) in
+ let nf = Universes.nf_evars_and_universes_opt_subst subst_evar
+ (Evd.evar_universe_context_subst universes) in
+ let make_body =
+ if poly || now then
+ let make_body t (c, eff) =
+ let open Universes in
+ let body = c and typ = nf t in
+ let used_univs_body = Universes.universes_of_constr body in
+ let used_univs_typ = Universes.universes_of_constr typ in
+ let ctx = Evd.evar_universe_context_set universes in
+ if keep_body_ucst_sepatate then
+ (* For vi2vo compilation proofs are computed now but we need to
+ * completent the univ constraints of the typ with the ones of
+ * the body. So we keep the two sets distinct. *)
+ let ctx_body = restrict_universe_context ctx used_univs_body in
+ let ctx_typ = restrict_universe_context ctx used_univs_typ in
+ let univs_typ = Univ.ContextSet.to_context ctx_typ in
+ (univs_typ, typ), ((body, ctx_body), eff)
+ else
+ (* Since the proof is computed now, we can simply have 1 set of
+ * constraints in which we merge the ones for the body and the ones
+ * for the typ *)
+ let used_univs = Univ.LSet.union used_univs_body used_univs_typ in
+ let ctx = restrict_universe_context ctx used_univs in
+ let univs = Univ.ContextSet.to_context ctx in
+ (univs, typ), ((body, Univ.ContextSet.empty), eff)
+ in
+ fun t p ->
+ Future.split2 (Future.chain ~pure:true p (make_body t))
+ else
+ fun t p ->
+ let initunivs = Evd.evar_context_universe_context universes in
+ Future.from_val (initunivs, nf t),
+ Future.chain ~pure:true p (fun (pt,eff) ->
+ (pt, Evd.evar_universe_context_set (Future.force univs)), eff)
+ in
+ let entries =
+ Future.map2 (fun p (_, t) ->
+ let univstyp, body = make_body t p in
+ let univs, typ = Future.force univstyp in
+ { Entries.
+ const_entry_body = body;
+ const_entry_secctx = section_vars;
+ const_entry_feedback = feedback_id;
+ const_entry_type = Some typ;
+ const_entry_inline_code = false;
+ const_entry_opaque = true;
+ const_entry_universes = univs;
+ const_entry_polymorphic = poly})
+ fpl initial_goals in
+ { id = pid; entries = entries; persistence = strength; universes = universes },
+ fun pr_ending -> Ephemeron.get terminator pr_ending
+
+type closed_proof_output = (Term.constr * Declareops.side_effects) list * Evd.evar_universe_context
+
+let return_proof () =
+ let { proof; strength = (_,poly,_) } = cur_pstate () in
+ let initial_goals = Proof.initial_goals proof in
+ let evd =
+ let error s = raise (Errors.UserError("last tactic before Qed",s)) in
+ try Proof.return proof with
+ | Proof.UnfinishedProof ->
+ error(str"Attempt to save an incomplete proof")
+ | Proof.HasShelvedGoals ->
+ error(str"Attempt to save a proof with shelved goals")
+ | Proof.HasGivenUpGoals ->
+ error(str"Attempt to save a proof with given up goals")
+ | Proof.HasUnresolvedEvar->
+ error(str"Attempt to save a proof with existential " ++
+ str"variables still non-instantiated") in
+ let eff = Evd.eval_side_effects evd in
+ let evd =
+ if poly || !Flags.compilation_mode = Flags.BuildVo
+ then Evd.nf_constraints evd
+ else evd in
+ (** ppedrot: FIXME, this is surely wrong. There is no reason to duplicate
+ side-effects... This may explain why one need to uniquize side-effects
+ thereafter... *)
+ let proofs =
+ List.map (fun (c, _) -> (Evarutil.nf_evars_universes evd c, eff)) initial_goals in
+ proofs, Evd.evar_universe_context evd
+
+let close_future_proof ~feedback_id proof =
+ close_proof ~keep_body_ucst_sepatate:true ~feedback_id ~now:false proof
+let close_proof ~keep_body_ucst_sepatate fix_exn =
+ close_proof ~keep_body_ucst_sepatate ~now:true
+ (Future.from_val ~fix_exn (return_proof ()))
+
+(** Gets the current terminator without checking that the proof has
+ been completed. Useful for the likes of [Admitted]. *)
+let get_terminator () = Ephemeron.get ( cur_pstate() ).terminator
+let set_terminator hook =
+ match !pstates with
+ | [] -> raise NoCurrentProof
+ | p :: ps -> pstates := { p with terminator = Ephemeron.create hook } :: ps
-(**********************************************************)
-(* *)
-(* Utility functions *)
-(* *)
-(**********************************************************)
-let maximal_unfocus k p =
- begin try while Proof.no_focused_goal p do
- Proof.unfocus k p
- done
- with Proof.FullyUnfocused | Proof.CannotUnfocusThisWay -> ()
- end
(**********************************************************)
@@ -314,14 +390,25 @@ let maximal_unfocus k p =
module Bullet = struct
- open Store.Field
+ type t = Vernacexpr.bullet
+
+ let bullet_eq b1 b2 = match b1, b2 with
+ | Vernacexpr.Dash n1, Vernacexpr.Dash n2 -> n1 = n2
+ | Vernacexpr.Star n1, Vernacexpr.Star n2 -> n1 = n2
+ | Vernacexpr.Plus n1, Vernacexpr.Plus n2 -> n1 = n2
+ | _ -> false
+ let pr_bullet b =
+ match b with
+ | Vernacexpr.Dash n -> str (String.make n '-')
+ | Vernacexpr.Star n -> str (String.make n '*')
+ | Vernacexpr.Plus n -> str (String.make n '+')
- type t = Vernacexpr.bullet
type behavior = {
name : string;
- put : Proof.proof -> t -> unit
+ put : Proof.proof -> t -> Proof.proof;
+ suggest: Proof.proof -> string option
}
let behaviors = Hashtbl.create 4
@@ -330,11 +417,55 @@ module Bullet = struct
(*** initial modes ***)
let none = {
name = "None";
- put = fun _ _ -> ()
+ put = (fun x _ -> x);
+ suggest = (fun _ -> None)
}
let _ = register_behavior none
module Strict = struct
+ type suggestion =
+ | Suggest of t (* this bullet is mandatory here *)
+ | Unfinished of t (* no mandatory bullet here, but this bullet is unfinished *)
+ | NoBulletInUse (* No mandatory bullet (or brace) here, no bullet pending,
+ some focused goals exists. *)
+ | NeedClosingBrace (* Some unfocussed goal exists "{" needed to focus them *)
+ | ProofFinished (* No more goal anywhere *)
+
+ (* give a message only if more informative than the standard coq message *)
+ let suggest_on_solved_goal sugg =
+ match sugg with
+ | NeedClosingBrace -> Some "Try unfocusing with \"}\"."
+ | NoBulletInUse -> None
+ | ProofFinished -> None
+ | Suggest b -> Some ("Focus next goal with bullet "
+ ^ Pp.string_of_ppcmds (Pp.(pr_bullet b))
+ ^".")
+ | Unfinished b -> Some ("The current bullet "
+ ^ Pp.string_of_ppcmds (Pp.(pr_bullet b))
+ ^ " is unfinished.")
+
+ (* give always a message. *)
+ let suggest_on_error sugg =
+ match sugg with
+ | NeedClosingBrace -> "Try unfocusing with \"}\"."
+ | NoBulletInUse -> assert false (* This should never raise an error. *)
+ | ProofFinished -> "No more subgoals."
+ | Suggest b -> ("Bullet " ^ Pp.string_of_ppcmds (Pp.(pr_bullet b))
+ ^ " is mandatory here.")
+ | Unfinished b -> ("Current bullet " ^ Pp.string_of_ppcmds (Pp.(pr_bullet b))
+ ^ " is not finished.")
+
+ exception FailedBullet of t * suggestion
+
+ let _ =
+ Errors.register_handler
+ (function
+ | FailedBullet (b,sugg) ->
+ let prefix = "Wrong bullet " ^ Pp.string_of_ppcmds (Pp.(pr_bullet b)) ^ " : " in
+ Errors.errorlabstrm "Focus" (str prefix ++ str (suggest_on_error sugg))
+ | _ -> raise Errors.Unhandled)
+
+
(* spiwack: we need only one focus kind as we keep a stack of (distinct!) bullets *)
let bullet_kind = (Proof.new_focus_kind () : t list Proof.focus_kind)
let bullet_cond = Proof.done_cond ~loose_end:true bullet_kind
@@ -350,42 +481,88 @@ module Bullet = struct
let has_bullet bul pr =
let rec has_bullet = function
- | b'::_ when bul=b' -> true
+ | b'::_ when bullet_eq bul b' -> true
| _::l -> has_bullet l
| [] -> false
in
has_bullet (get_bullets pr)
- (* precondition: the stack is not empty *)
+ (* pop a bullet from proof [pr]. There should be at least one
+ bullet in use. If pop impossible (pending proofs on this level
+ of bullet or higher) then raise [Proof.CannotUnfocusThisWay]. *)
let pop pr =
match get_bullets pr with
- | b::_ ->
- Proof.unfocus bullet_kind pr;
- (*returns*) b
+ | b::_ -> Proof.unfocus bullet_kind pr () , b
| _ -> assert false
- let push b pr =
+ let push (b:t) pr =
Proof.focus bullet_cond (b::get_bullets pr) 1 pr
+ (* Used only in the next function.
+ TODO: use a recursive function instead? *)
+ exception SuggestFound of t
+
+ let suggest_bullet (prf:Proof.proof): suggestion =
+ if Proof.is_done prf then ProofFinished
+ else if not (Proof.no_focused_goal prf)
+ then (* No suggestion if a bullet is not mandatory, look for an unfinished bullet *)
+ match get_bullets prf with
+ | b::_ -> Unfinished b
+ | _ -> NoBulletInUse
+ else (* There is no goal under focus but some are unfocussed,
+ let us look at the bullet needed. If no *)
+ let pcobaye = ref prf in
+ try
+ while true do
+ let pcobaye', b = pop !pcobaye in
+ (* pop went well, this means that there are no more goals
+ *under this* bullet b, see if a new b can be pushed. *)
+ (try let _ = push b pcobaye' in (* push didn't fail so a new b can be pushed. *)
+ raise (SuggestFound b)
+ with SuggestFound _ as e -> raise e
+ | _ -> ()); (* b could not be pushed, so we must look for a outer bullet *)
+ pcobaye := pcobaye'
+ done;
+ assert false
+ with SuggestFound b -> Suggest b
+ | _ -> NeedClosingBrace (* No push was possible, but there are still
+ subgoals somewhere: there must be a "}" to use. *)
+
+
+ let rec pop_until (prf:Proof.proof) bul: Proof.proof =
+ let prf', b = pop prf in
+ if bullet_eq bul b then prf'
+ else pop_until prf' bul
+
let put p bul =
- if has_bullet bul p then
- Proof.transaction p begin fun () ->
- while bul <> pop p do () done;
+ try
+ if not (has_bullet bul p) then
+ (* bullet is not in use, so pushing it is always ok unless
+ no goal under focus. *)
push bul p
- end
- else
- push bul p
+ else
+ match suggest_bullet p with
+ | Suggest suggested_bullet when bullet_eq bul suggested_bullet
+ -> (* suggested_bullet is mandatory and you gave the right one *)
+ let p' = pop_until p bul in
+ push bul p'
+ (* the bullet you gave is in use but not the right one *)
+ | sugg -> raise (FailedBullet (bul,sugg))
+ with Proof.NoSuchGoals _ -> (* push went bad *)
+ raise (FailedBullet (bul,suggest_bullet p))
let strict = {
name = "Strict Subproofs";
- put = put
+ put = put;
+ suggest = (fun prf -> suggest_on_solved_goal (suggest_bullet prf))
+
}
let _ = register_behavior strict
end
(* Current bullet behavior, controled by the option *)
let current_behavior = ref Strict.strict
-
+
let _ =
Goptions.declare_string_option {Goptions.
optsync = true;
@@ -402,16 +579,77 @@ module Bullet = struct
let put p b =
(!current_behavior).put p b
+
+ let suggest p =
+ (!current_behavior).suggest p
end
+let _ =
+ let hook n =
+ let prf = give_me_the_proof () in
+ (Bullet.suggest prf) in
+ Proofview.set_nosuchgoals_hook hook
+
+
+(**********************************************************)
+(* *)
+(* Default goal selector *)
+(* *)
+(**********************************************************)
+
+
+(* Default goal selector: selector chosen when a tactic is applied
+ without an explicit selector. *)
+let default_goal_selector = ref (Vernacexpr.SelectNth 1)
+let get_default_goal_selector () = !default_goal_selector
+
+let print_goal_selector = function
+ | Vernacexpr.SelectAll -> "all"
+ | Vernacexpr.SelectNth i -> string_of_int i
+ | Vernacexpr.SelectId id -> Id.to_string id
+ | Vernacexpr.SelectAllParallel -> "par"
+
+let parse_goal_selector = function
+ | "all" -> Vernacexpr.SelectAll
+ | i ->
+ let err_msg = "A selector must be \"all\" or a natural number." in
+ begin try
+ let i = int_of_string i in
+ if i < 0 then Errors.error err_msg;
+ Vernacexpr.SelectNth i
+ with Failure _ -> Errors.error err_msg
+ end
+
+let _ =
+ Goptions.declare_string_option {Goptions.
+ optsync = true ;
+ optdepr = false;
+ optname = "default goal selector" ;
+ optkey = ["Default";"Goal";"Selector"] ;
+ optread = begin fun () -> print_goal_selector !default_goal_selector end;
+ optwrite = begin fun n ->
+ default_goal_selector := parse_goal_selector n
+ end
+ }
+
+
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))
+ let { pid; strength; proof } = cur_pstate () in
+ let initial = Proof.initial_goals proof in
+ let goals = List.map (fun (o, c) -> c) initial in
+ pid, (goals, strength)
end
+type state = pstate list
+
+let freeze ~marshallable =
+ match marshallable with
+ | `Yes ->
+ Errors.anomaly (Pp.str"full marshalling of proof state not supported")
+ | `Shallow -> !pstates
+ | `No -> !pstates
+let unfreeze s = pstates := s; update_proof_mode ()
+let proof_of_state = function { proof }::_ -> proof | _ -> raise NoCurrentProof
+
diff --git a/proofs/proof_global.mli b/proofs/proof_global.mli
index 08ae7519..2700e901 100644
--- a/proofs/proof_global.mli
+++ b/proofs/proof_global.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -28,14 +28,13 @@ type proof_mode = {
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 get_current_proof_name : unit -> Names.Id.t
+val get_all_proof_names : unit -> Names.Id.t list
-val discard : Names.identifier Util.located -> unit
+val discard : Names.Id.t Loc.located -> unit
val discard_current : unit -> unit
val discard_all : unit -> unit
@@ -45,53 +44,102 @@ val set_proof_mode : string -> unit
exception NoCurrentProof
val give_me_the_proof : unit -> Proof.proof
+(** @raise NoCurrentProof when outside proof mode. *)
+val compact_the_proof : unit -> unit
-(** [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). *)
+(** When a proof is closed, it is reified into a [proof_object], where
+ [id] is the name of the proof, [entries] the list of the proof terms
+ (in a form suitable for definitions). Together with the [terminator]
+ function which takes a [proof_object] together with a [proof_end]
+ (i.e. an proof ending command) and registers the appropriate
+ values. *)
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)
+type proof_universes = Evd.evar_universe_context
+type proof_object = {
+ id : Names.Id.t;
+ entries : Entries.definition_entry list;
+ persistence : Decl_kinds.goal_kind;
+ universes: proof_universes;
+ (* constraints : Univ.constraints; *)
+ (** guards : lemma_possible_guards; *)
+}
+
+type proof_ending =
+ | Admitted
+ | Proved of Vernacexpr.opacity_flag *
+ (Vernacexpr.lident * Decl_kinds.theorem_kind option) option *
+ proof_object
+type proof_terminator = proof_ending -> unit
+type closed_proof = proof_object * proof_terminator
+
+(** [start_proof id str goals terminator] starts a proof of name [id]
+ with goals [goals] (a list of pairs of environment and
+ conclusion); [str] describes what kind of theorem/definition this
+ is (spiwack: for potential printing, I believe is used only by
+ closing commands and the xml plugin); [terminator] is used at the
+ end of the proof to close the proof. *)
+val start_proof :
+ Evd.evar_map -> Names.Id.t -> Decl_kinds.goal_kind -> (Environ.env * Term.types) list ->
+ proof_terminator -> unit
+
+(** Like [start_proof] except that there may be dependencies between
+ initial goals. *)
+val start_dependent_proof :
+ Names.Id.t -> Decl_kinds.goal_kind -> Proofview.telescope ->
+ proof_terminator -> unit
+
+(* Takes a function to add to the exceptions data relative to the
+ state in which the proof was built *)
+val close_proof : keep_body_ucst_sepatate:bool -> Future.fix_exn -> closed_proof
+
+(* Intermediate step necessary to delegate the future.
+ * Both access the current proof state. The formes is supposed to be
+ * chained with a computation that completed the proof *)
+
+type closed_proof_output = (Term.constr * Declareops.side_effects) list * Evd.evar_universe_context
+
+val return_proof : unit -> closed_proof_output
+val close_future_proof : feedback_id:Stateid.t ->
+ closed_proof_output Future.computation -> closed_proof
+
+(** Gets the current terminator without checking that the proof has
+ been completed. Useful for the likes of [Admitted]. *)
+val get_terminator : unit -> proof_terminator
+val set_terminator : proof_terminator -> unit
exception NoSuchProof
-(** Runs a tactic on the current proof. Raises [NoCurrentProof] is there is
- no current proof. *)
-val run_tactic : unit Proofview.tactic -> unit
+val get_open_goals : unit -> int
-(** Sets the tactic to be used when a tactic line is closed with [...] *)
-val set_endline_tactic : unit Proofview.tactic -> unit
+(** Runs a tactic on the current proof. Raises [NoCurrentProof] is there is
+ no current proof.
+ The return boolean is set to [false] if an unsafe tactic has been used. *)
+val with_current_proof :
+ (unit Proofview.tactic -> Proof.proof -> Proof.proof*'a) -> 'a
+val simple_with_current_proof :
+ (unit Proofview.tactic -> Proof.proof -> Proof.proof) -> unit
-(** Sets the section variables assumed by the proof *)
-val set_used_variables : Names.identifier list -> unit
-val get_used_variables : unit -> Sign.section_context option
+(** Sets the tactic to be used when a tactic line is closed with [...] *)
+val set_endline_tactic : Tacexpr.raw_tactic_expr -> unit
+val set_interp_tac :
+ (Tacexpr.raw_tactic_expr -> 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
+(** Sets the section variables assumed by the proof, returns its closure
+ * (w.r.t. type dependencies *)
+val set_used_variables : Names.Id.t list -> Context.section_context
+val get_used_variables : unit -> Context.section_context option
(**********************************************************)
-(* *)
-(* Utility functions *)
-(* *)
+(* *)
+(* Proof modes *)
+(* *)
(**********************************************************)
-(** [maximal_unfocus k p] unfocuses [p] until [p] has at least a
- focused goal or that the last focus isn't of kind [k]. *)
-val maximal_unfocus : 'a Proof.focus_kind -> Proof.proof -> unit
+
+val activate_proof_mode : string -> unit
+val disactivate_proof_mode : string -> unit
(**********************************************************)
(* *)
@@ -103,11 +151,13 @@ module Bullet : sig
type t = Vernacexpr.bullet
(** A [behavior] is the data of a put function which
- is called when a bullet prefixes a tactic, together
- with a name to identify it. *)
+ is called when a bullet prefixes a tactic, a suggest function
+ suggesting the right bullet to use on a given proof, together
+ with a name to identify the behavior. *)
type behavior = {
name : string;
- put : Proof.proof -> t -> unit
+ put : Proof.proof -> t -> Proof.proof;
+ suggest: Proof.proof -> string option
}
(** A registered behavior can then be accessed in Coq
@@ -123,9 +173,25 @@ module Bullet : sig
(** Handles focusing/defocusing with bullets:
*)
- val put : Proof.proof -> t -> unit
+ val put : Proof.proof -> t -> Proof.proof
+ val suggest : Proof.proof -> string option
end
+
+(**********************************************************)
+(* *)
+(* Default goal selector *)
+(* *)
+(**********************************************************)
+
+val get_default_goal_selector : unit -> Vernacexpr.goal_selector
+
module V82 : sig
- val get_current_initial_conclusions : unit -> Names.identifier *(Term.types list * Decl_kinds.goal_kind * Tacexpr.declaration_hook)
+ val get_current_initial_conclusions : unit -> Names.Id.t *(Term.types list *
+ Decl_kinds.goal_kind)
end
+
+type state
+val freeze : marshallable:[`Yes | `No | `Shallow] -> state
+val unfreeze : state -> unit
+val proof_of_state : state -> Proof.proof
diff --git a/proofs/proof_type.ml b/proofs/proof_type.ml
index c56f8a24..26bb78df 100644
--- a/proofs/proof_type.ml
+++ b/proofs/proof_type.ml
@@ -1,104 +1,53 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
(*i*)
-open Environ
open Evd
open Names
-open Libnames
open Term
-open Util
+open Context
open Tacexpr
-(* open Decl_expr *)
open Glob_term
-open Genarg
open Nametab
-open Pattern
+open Misctypes
(*i*)
(* This module defines the structure of proof tree and the tactic type. So, it
is used by Proof_tree and Refiner *)
+(** Types of goals, tactics, rules ... *)
+
type goal = Goal.goal
type tactic = goal sigma -> goal list sigma
type prim_rule =
- | Intro of identifier
- | Cut of bool * bool * identifier * types
- | FixRule of identifier * int * (identifier * int * constr) list * int
- | Cofix of identifier * (identifier * constr) list * int
+ | Cut of bool * bool * Id.t * types
+ | FixRule of Id.t * int * (Id.t * int * constr) list * int
+ | Cofix of Id.t * (Id.t * constr) list * int
| Refine of constr
- | Convert_concl of types * cast_kind
- | Convert_hyp of named_declaration
- | Thin of identifier list
- | ThinBody of identifier list
- | Move of bool * identifier * identifier move_location
- | Order of identifier list
- | Rename of identifier * identifier
- | Change_evars
-
-type proof_tree = {
- goal : goal;
- ref : (rule * proof_tree list) option }
-
-and rule =
- | Prim of prim_rule
- | Nested of compound_rule * proof_tree
- | Decl_proof of bool
- | Daimon
+ | Thin of Id.t list
+ | Move of Id.t * Id.t move_location
-and compound_rule=
- | Tactic of tactic_expr * bool
+(** Nowadays, the only rules we'll consider are the primitive rules *)
-and tactic_expr =
- (constr,
- constr_pattern,
- evaluable_global_reference,
- inductive,
- ltac_constant,
- identifier,
- glob_tactic_expr,
- tlevel)
- Tacexpr.gen_tactic_expr
+type rule = prim_rule
-and atomic_tactic_expr =
- (constr,
- constr_pattern,
- evaluable_global_reference,
- inductive,
- ltac_constant,
- identifier,
- glob_tactic_expr,
- tlevel)
- Tacexpr.gen_atomic_tactic_expr
-
-and tactic_arg =
- (constr,
- constr_pattern,
- evaluable_global_reference,
- inductive,
- ltac_constant,
- identifier,
- glob_tactic_expr,
- tlevel)
- Tacexpr.gen_tactic_arg
+(** Ltac traces *)
type ltac_call_kind =
- | LtacNotationCall of string
+ | LtacMLCall of glob_tactic_expr
+ | LtacNotationCall of KerName.t
| LtacNameCall of ltac_constant
- | LtacAtomCall of glob_atomic_tactic_expr * atomic_tactic_expr option ref
- | LtacVarCall of identifier * glob_tactic_expr
- | LtacConstrInterp of glob_constr *
- (extended_patvar_map * (identifier * identifier option) list)
-
-type ltac_trace = (int * loc * ltac_call_kind) list
+ | LtacAtomCall of glob_atomic_tactic_expr
+ | LtacVarCall of Id.t * glob_tactic_expr
+ | LtacConstrInterp of glob_constr * Pretyping.ltac_var_map
-exception LtacLocated of (int * ltac_call_kind * ltac_trace * loc) * exn
+type ltac_trace = (Loc.t * ltac_call_kind) list
-let abstract_tactic_box = ref (ref None)
+let (ltac_trace_info : ltac_trace Exninfo.t) = Exninfo.make ()
diff --git a/proofs/proof_type.mli b/proofs/proof_type.mli
index 6fa8087e..e709be5b 100644
--- a/proofs/proof_type.mli
+++ b/proofs/proof_type.mli
@@ -1,44 +1,34 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Environ
open Evd
open Names
-open Libnames
open Term
-open Util
+open Context
open Tacexpr
open Glob_term
-open Genarg
open Nametab
-open Pattern
+open Misctypes
(** 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
- | FixRule of identifier * int * (identifier * int * constr) list * int
- | Cofix of identifier * (identifier * constr) list * int
+ | Cut of bool * bool * Id.t * types
+ | FixRule of Id.t * int * (Id.t * int * constr) list * int
+ | Cofix of Id.t * (Id.t * constr) list * int
| Refine of constr
- | Convert_concl of types * cast_kind
- | Convert_hyp of named_declaration
- | Thin of identifier list
- | ThinBody of identifier list
- | Move of bool * identifier * identifier move_location
- | Order of identifier list
- | Rename of identifier * identifier
- | Change_evars
+ | Thin of Id.t list
+ | Move of Id.t * Id.t move_location
+
+(** Nowadays, the only rules we'll consider are the primitive rules *)
+
+type rule = prim_rule
(** The type [goal sigma] is the type of subgoal. It has the following form
{v it = \{ evar_concl = [the conclusion of the subgoal]
@@ -65,70 +55,22 @@ type prim_rule =
in the type of evar] \} \} \} v}
*)
-(** {6 ... } *)
-(** Proof trees.
- [ref] = [None] if the goal has still to be proved,
- and [Some (r,l)] if the rule [r] was applied to the goal
- and gave [l] as subproofs to be completed.
- if [ref = (Some(Nested(Tactic t,p),l))] then [p] is the proof
- that the goal can be proven if the goals in [l] are solved. *)
-type proof_tree = {
- goal : goal;
- ref : (rule * proof_tree list) option }
-
-and rule =
- | Prim of prim_rule
- | Nested of compound_rule * proof_tree
- | Decl_proof of bool
- | Daimon
-
-and compound_rule=
- (** the boolean of Tactic tells if the default tactic is used *)
- | Tactic of tactic_expr * bool
+type goal = Goal.goal
-and tactic_expr =
- (constr,
- constr_pattern,
- evaluable_global_reference,
- inductive,
- ltac_constant,
- identifier,
- glob_tactic_expr,
- tlevel)
- Tacexpr.gen_tactic_expr
+type tactic = goal sigma -> goal list sigma
-and atomic_tactic_expr =
- (constr,
- constr_pattern,
- evaluable_global_reference,
- inductive,
- ltac_constant,
- identifier,
- glob_tactic_expr,
- tlevel)
- Tacexpr.gen_atomic_tactic_expr
+(** Ltac traces *)
-and tactic_arg =
- (constr,
- constr_pattern,
- evaluable_global_reference,
- inductive,
- ltac_constant,
- identifier,
- glob_tactic_expr,
- tlevel)
- Tacexpr.gen_tactic_arg
+(** TODO: Move those definitions somewhere sensible *)
type ltac_call_kind =
- | LtacNotationCall of string
+ | LtacMLCall of glob_tactic_expr
+ | LtacNotationCall of KerName.t
| LtacNameCall of ltac_constant
- | LtacAtomCall of glob_atomic_tactic_expr * atomic_tactic_expr option ref
- | LtacVarCall of identifier * glob_tactic_expr
- | LtacConstrInterp of glob_constr *
- (extended_patvar_map * (identifier * identifier option) list)
-
-type ltac_trace = (int * loc * ltac_call_kind) list
+ | LtacAtomCall of glob_atomic_tactic_expr
+ | LtacVarCall of Id.t * glob_tactic_expr
+ | LtacConstrInterp of glob_constr * Pretyping.ltac_var_map
-exception LtacLocated of (int * ltac_call_kind * ltac_trace * loc) * exn
+type ltac_trace = (Loc.t * ltac_call_kind) list
-val abstract_tactic_box : atomic_tactic_expr option ref ref
+val ltac_trace_info : ltac_trace Exninfo.t
diff --git a/proofs/proof_using.ml b/proofs/proof_using.ml
new file mode 100644
index 00000000..f66e9657
--- /dev/null
+++ b/proofs/proof_using.ml
@@ -0,0 +1,166 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open Environ
+open Util
+open Vernacexpr
+
+let to_string = function
+ | SsAll -> "All"
+ | SsType -> "Type"
+ | SsExpr(SsSet l)-> String.concat " " (List.map Id.to_string (List.map snd l))
+ | SsExpr e ->
+ let rec aux = function
+ | SsSet [] -> "( )"
+ | SsSet [_,x] -> Id.to_string x
+ | SsSet l ->
+ "(" ^ String.concat " " (List.map Id.to_string (List.map snd l)) ^ ")"
+ | SsCompl e -> "-" ^ aux e^""
+ | SsUnion(e1,e2) -> "("^aux e1 ^" + "^ aux e2^")"
+ | SsSubstr(e1,e2) -> "("^aux e1 ^" - "^ aux e2^")"
+ in aux e
+
+let known_names = Summary.ref [] ~name:"proofusing-nameset"
+
+let in_nameset =
+ let open Libobject in
+ declare_object { (default_object "proofusing-nameset") with
+ cache_function = (fun (_,x) -> known_names := x :: !known_names);
+ classify_function = (fun _ -> Dispose);
+ discharge_function = (fun _ -> None)
+ }
+
+let rec process_expr env e ty =
+ match e with
+ | SsAll ->
+ List.fold_right Id.Set.add (List.map pi1 (named_context env)) Id.Set.empty
+ | SsExpr e ->
+ let rec aux = function
+ | SsSet l -> set_of_list env (List.map snd l)
+ | SsUnion(e1,e2) -> Id.Set.union (aux e1) (aux e2)
+ | SsSubstr(e1,e2) -> Id.Set.diff (aux e1) (aux e2)
+ | SsCompl e -> Id.Set.diff (full_set env) (aux e)
+ in
+ aux e
+ | SsType ->
+ List.fold_left (fun acc ty ->
+ Id.Set.union (global_vars_set env ty) acc)
+ Id.Set.empty ty
+and set_of_list env = function
+ | [x] when CList.mem_assoc_f Id.equal x !known_names ->
+ process_expr env (CList.assoc_f Id.equal x !known_names) []
+ | l -> List.fold_right Id.Set.add l Id.Set.empty
+and full_set env = set_of_list env (List.map pi1 (named_context env))
+
+let process_expr env e ty =
+ let s = Id.Set.union (process_expr env SsType ty) (process_expr env e []) in
+ Id.Set.elements s
+
+let name_set id expr = Lib.add_anonymous_leaf (in_nameset (id,expr))
+
+let minimize_hyps env ids =
+ let rec aux ids =
+ let ids' =
+ Id.Set.fold (fun id alive ->
+ let impl_by_id =
+ Id.Set.remove id (really_needed env (Id.Set.singleton id)) in
+ if Id.Set.is_empty impl_by_id then alive
+ else Id.Set.diff alive impl_by_id)
+ ids ids in
+ if Id.Set.equal ids ids' then ids else aux ids'
+ in
+ aux ids
+
+let minimize_unused_hyps env ids =
+ let all_ids = List.map pi1 (named_context env) in
+ let deps_of =
+ let cache =
+ List.map (fun id -> id,really_needed env (Id.Set.singleton id)) all_ids in
+ fun id -> List.assoc id cache in
+ let inv_dep_of =
+ let cache_sum cache id stuff =
+ try Id.Map.add id (Id.Set.add stuff (Id.Map.find id cache)) cache
+ with Not_found -> Id.Map.add id (Id.Set.singleton stuff) cache in
+ let cache =
+ List.fold_left (fun cache id ->
+ Id.Set.fold (fun d cache -> cache_sum cache d id)
+ (Id.Set.remove id (deps_of id)) cache)
+ Id.Map.empty all_ids in
+ fun id -> try Id.Map.find id cache with Not_found -> Id.Set.empty in
+ let rec aux s =
+ let s' =
+ Id.Set.fold (fun id s ->
+ if Id.Set.subset (inv_dep_of id) s then Id.Set.diff s (inv_dep_of id)
+ else s)
+ s s in
+ if Id.Set.equal s s' then s else aux s' in
+ aux ids
+
+let suggest_Proof_using kn env vars ids_typ context_ids =
+ let module S = Id.Set in
+ let open Pp in
+ let used = S.union vars ids_typ in
+ let needed = minimize_hyps env used in
+ let all_needed = really_needed env needed in
+ let all = List.fold_right S.add context_ids S.empty in
+ let unneeded = minimize_unused_hyps env (S.diff all needed) in
+ let pr_set s =
+ let wrap ppcmds =
+ if S.cardinal s > 1 || S.equal s (S.singleton (Id.of_string "All"))
+ then str "(" ++ ppcmds ++ str ")"
+ else ppcmds in
+ wrap (prlist_with_sep (fun _ -> str" ") Id.print (S.elements s)) in
+ if !Flags.debug then begin
+ prerr_endline (string_of_ppcmds (str "All " ++ pr_set all));
+ prerr_endline (string_of_ppcmds (str "Type" ++ pr_set ids_typ));
+ prerr_endline (string_of_ppcmds (str "needed " ++ pr_set needed));
+ prerr_endline (string_of_ppcmds (str "unneeded " ++ pr_set unneeded));
+ end;
+ msg_info (
+ str"The proof of "++
+ Names.Constant.print kn ++ spc() ++ str "should start with:"++spc()++
+ str"Proof using " ++
+ if S.is_empty needed then str "."
+ else if S.subset needed ids_typ then str "Type."
+ else if S.equal all all_needed then str "All."
+ else
+ let s1 = string_of_ppcmds (str "-" ++ pr_set unneeded ++ str".") in
+ let s2 = string_of_ppcmds (pr_set needed ++ str".") in
+ if String.length s1 < String.length s2 then str s1 else str s2)
+
+let value = ref false
+
+let _ =
+ Goptions.declare_bool_option
+ { Goptions.optsync = true;
+ Goptions.optdepr = false;
+ Goptions.optname = "suggest Proof using";
+ Goptions.optkey = ["Suggest";"Proof";"Using"];
+ Goptions.optread = (fun () -> !value);
+ Goptions.optwrite = (fun b ->
+ value := b;
+ if b then Term_typing.set_suggest_proof_using suggest_Proof_using
+ else Term_typing.set_suggest_proof_using (fun _ _ _ _ _ -> ())
+ ) }
+
+let value = ref "_unset_"
+
+let _ =
+ Goptions.declare_string_option
+ { Goptions.optsync = true;
+ Goptions.optdepr = false;
+ Goptions.optname = "default value for Proof using";
+ Goptions.optkey = ["Default";"Proof";"Using"];
+ Goptions.optread = (fun () -> !value);
+ Goptions.optwrite = (fun b -> value := b;) }
+
+
+let get_default_proof_using () =
+ if !value = "_unset_" then None
+ else Some !value
diff --git a/proofs/proof_using.mli b/proofs/proof_using.mli
new file mode 100644
index 00000000..fb3497f1
--- /dev/null
+++ b/proofs/proof_using.mli
@@ -0,0 +1,26 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+
+(* [minimize_hyps e s1] gives [s2] s.t. [Id.Set.subset s2 s1] is [true]
+ * and [keep_hyps e s1] is equal to [keep_hyps e s2]. Inefficient. *)
+val minimize_hyps : Environ.env -> Names.Id.Set.t -> Names.Id.Set.t
+
+(* [minimize_unused_hyps e s1] gives [s2] s.t. [Id.Set.subset s2 s1] is [true]
+ * and s.t. calling [clear s1] would do the same as [clear s2]. Inefficient. *)
+val minimize_unused_hyps : Environ.env -> Names.Id.Set.t -> Names.Id.Set.t
+
+val process_expr :
+ Environ.env -> Vernacexpr.section_subset_descr -> Constr.types list ->
+ Names.Id.t list
+
+val name_set : Names.Id.t -> Vernacexpr.section_subset_descr -> unit
+
+val to_string : Vernacexpr.section_subset_descr -> string
+
+val get_default_proof_using : unit -> string option
diff --git a/proofs/proofs.mllib b/proofs/proofs.mllib
index 66001e77..32bf5576 100644
--- a/proofs/proofs.mllib
+++ b/proofs/proofs.mllib
@@ -1,12 +1,16 @@
+Miscprint
Goal
Evar_refiner
+Proof_using
+Proof_type
+Proof_errors
+Logic_monad
+Proofview_monad
+Logic
Proofview
Proof
Proof_global
-Tacexpr
-Proof_type
Redexpr
-Logic
Refiner
Tacmach
Pfedit
diff --git a/proofs/proofview.ml b/proofs/proofview.ml
index 17be1f7d..a25683bf 100644
--- a/proofs/proofview.ml
+++ b/proofs/proofview.ml
@@ -1,525 +1,1175 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Compat
-
-(* 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
- }
+(** This files defines the basic mechanism of proofs: the [proofview]
+ type is the state which tactics manipulate (a global state for
+ existential variables, together with the list of goals), and the type
+ ['a tactic] is the (abstract) type of tactics modifying the proof
+ state and returning a value of type ['a]. *)
+
+open Pp
+open Util
+open Proofview_monad
+
+(** Main state of tactics *)
+type proofview = Proofview_monad.proofview
+
+type entry = (Term.constr * Term.types) list
+
+(** Returns a stylised view of a proofview for use by, for instance,
+ ide-s. *)
+(* spiwack: the type of [proofview] will change as we push more
+ refined functions to ide-s. This would be better than spawning a
+ new nearly identical function everytime. Hence the generic name. *)
+(* In this version: returns the list of focused goals together with
+ the [evar_map] context. *)
let proofview p =
p.comb , p.solution
-(* Initialises a proofview, the argument is a list of environement,
- conclusion types, and optional names, creating that many initial goals. *)
-let init =
+let compact el { comb; solution } =
+ let nf = Evarutil.nf_evar solution in
+ let size = Evd.fold (fun _ _ i -> i+1) solution 0 in
+ let new_el = List.map (fun (t,ty) -> nf t, nf ty) el in
+ let pruned_solution = Evd.drop_all_defined solution in
+ let apply_subst_einfo _ ei =
+ Evd.({ ei with
+ evar_concl = nf ei.evar_concl;
+ evar_hyps = Environ.map_named_val nf ei.evar_hyps;
+ evar_candidates = Option.map (List.map nf) ei.evar_candidates }) in
+ let new_solution = Evd.raw_map_undefined apply_subst_einfo pruned_solution in
+ let new_size = Evd.fold (fun _ _ i -> i+1) new_solution 0 in
+ msg_info (Pp.str (Printf.sprintf "Evars: %d -> %d\n" size new_size));
+ new_el, { comb; solution = new_solution }
+
+
+(** {6 Starting and querying a proof view} *)
+
+type telescope =
+ | TNil of Evd.evar_map
+ | TCons of Environ.env * Evd.evar_map * Term.types * (Evd.evar_map -> Term.constr -> telescope)
+
+let dependent_init =
+ (* Goals are created with a store which marks them as unresolvable
+ for type classes. *)
+ let store = Typeclasses.set_resolvable Evd.Store.empty false in
+ (* Goals don't have a source location. *)
+ let src = (Loc.ghost,Evar_kinds.GoalEvar) in
+ (* Main routine *)
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 }
+ | TNil sigma -> [], { solution = sigma; comb = []; }
+ | TCons (env, sigma, typ, t) ->
+ let (sigma, econstr ) = Evarutil.new_evar env sigma ~src ~store typ in
+ let ret, { solution = sol; comb = comb } = aux (t sigma econstr) in
+ let (gl, _) = Term.destEvar econstr in
+ let entry = (econstr, typ) :: ret in
+ entry, { solution = sol; comb = gl :: comb; }
+ in
+ fun t ->
+ let entry, v = aux t in
+ (* The created goal are not to be shelved. *)
+ let solution = Evd.reset_future_goals v.solution in
+ entry, { v with solution }
+
+let init =
+ let rec aux sigma = function
+ | [] -> TNil sigma
+ | (env,g)::l -> TCons (env,sigma,g,(fun sigma _ -> aux sigma l))
in
- fun l -> let v = aux l in
- (* Marks all the goal unresolvable for typeclasses. *)
- { v with solution = Typeclasses.mark_unresolvables v.solution }
+ fun sigma l -> dependent_init (aux sigma l)
+
+let initial_goals initial = initial
-(* 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
-(* no handler: should not be allowed to reach toplevel *)
-
-(* [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
+let return { solution=defs } = defs
+
+let return_constr { solution = defs } c = Evarutil.nf_evar defs c
+
+let partial_proof entry pv = CList.map (return_constr pv) (CList.map fst entry)
+
+(** {6 Focusing commands} *)
+
+(** A [focus_context] represents the part of the proof view which has
+ been removed by a focusing action, it can be used to unfocus later
+ on. *)
+(* First component is a reverse list of the goals which come before
+ and second component is the list of the goals which go after (in
+ the expected order). *)
+type focus_context = Evar.t list * Evar.t list
+
+
+(** Returns a stylised view of a focus_context for use by, for
+ instance, ide-s. *)
+(* spiwack: the type of [focus_context] will change as we push more
+ refined functions to ide-s. This would be better than spawning a
+ new nearly identical function everytime. Hence the generic name. *)
+(* In this version: the goals in the context, as a "zipper" (the first
+ list is in reversed order). *)
let focus_context f = f
-(* 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 ]. *)
+(** 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 (left,sub_right) = CList.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"No such unproven subgoal")
+ try CList.chop (j-i+1) sub_right
+ with Failure _ -> raise CList.IndexOutOfRange
in
(sub, (left,right))
-(* Inverse operation to the previous one. *)
+(** Inverse operation to the previous one. *)
let unfocus_sublist (left,right) s =
- List.rev_append left (s@right)
+ CList.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). The first goal has index 1.
- It returns the focus proof, and a context for the focus trace. *)
+(** [focus i j] focuses a proofview on the goals from index [i] to
+ index [j] (inclusive, goals are indexed from [1]). I.e. goals
+ number [i] to [j] become the only focused goals of the returned
+ proofview. It returns the focused proofview, and a context for
+ the focus stack. *)
let focus i j sp =
let (new_comb, context) = focus_sublist i j sp.comb in
( { sp with comb = new_comb } , context )
-(* Unfocuses a proofview with respect to a context. *)
-let undefined defs l =
- Option.List.flatten (List.map (Goal.advance defs) l)
+
+(** [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. *)
+(* spiwack: [advance] is probably performance critical, and the good
+ behaviour of its definition may depend sensitively to the actual
+ definition of [Evd.find]. Currently, [Evd.find] starts looking for
+ a value in the heap of undefined variable, which is small. Hence in
+ the most common case, where [advance] is applied to an unsolved
+ goal ([advance] is used to figure if a side effect has modified the
+ goal) it terminates quickly. *)
+let rec advance sigma g =
+ let open Evd in
+ let evi = Evd.find sigma g in
+ match evi.evar_body with
+ | Evar_empty -> Some g
+ | Evar_defined v ->
+ if Option.default false (Store.get evi.evar_extra Evarutil.cleared) then
+ let (e,_) = Term.destEvar v in
+ advance sigma e
+ else
+ None
+
+(** [undefined defs l] is the list of goals in [l] which are still
+ unsolved (after advancing cleared goals). *)
+let undefined defs l = CList.map_filter (advance defs) l
+
+(** Unfocuses a proofview with respect to a context. *)
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).
+(** {6 The tactic monad} *)
+
+(** - Tactics are objects which apply a transformation to all the
+ subgoals of the current view at the same time. By opposition to
+ the old vision of applying it to a single goal. It allows tactics
+ such as [shelve_unifiable], tactics to reorder the focused goals,
+ 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). Another benefit is that it is possible to write tactics
+ that can be executed even if there are no focused goals.
+ - Tactics form 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 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") and [Proofview.tclTHEN]
+ (which is a specialized bind on unit-returning tactics).
+ - Tactics have support for full-backtracking. Tactics can be seen
+ having multiple success: if after returning the first success a
+ failure is encountered, the tactic can backtrack and use a second
+ success if available. The state is backtracked to its previous
+ value, except the non-logical state defined in the {!NonLogical}
+ module below.
*)
+(* spiwack: as far as I'm aware this doesn't really relate to
+ F. Kirchner and C. Muñoz. *)
-(* 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
- }
+module Proof = Logical
-(*** tacticals ***)
+(** type of tactics:
+ tactics can
+ - access the environment,
+ - report unsafe status, shelved goals and given up goals
+ - access and change the current [proofview]
+ - backtrack on previous changes of the proofview *)
+type +'a tactic = 'a Proof.t
-(* Unit of the tactic monad *)
-let tclUNIT a _ = { go = fun sk fk step -> sk a fk step }
+(** Applies a tactic to the current proofview. *)
+let apply env t sp =
+ let open Logic_monad in
+ let ans = Proof.repr (Proof.run t false (sp,env)) in
+ let ans = Logic_monad.NonLogical.run ans in
+ match ans with
+ | Nil (e, info) -> iraise (TacticFailure e, info)
+ | Cons ((r, (state, _), status, info), _) ->
+ r, state, status, Trace.to_tree info
-(* 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
-}
+(** {7 Monadic primitives} *)
-(* [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
-}
+(** Unit of the tactic monad. *)
+let tclUNIT = Proof.return
-(* [tclZERO e] always fails with error message [e]*)
-let tclZERO e env = { go = fun _ fk step -> fk e step }
+(** Bind operation of the tactic monad. *)
+let tclBIND = Proof.(>>=)
+(** Interpretes the ";" (semicolon) of Ltac. As a monadic operation,
+ it's a specialized "bind". *)
+let tclTHEN = Proof.(>>)
+
+(** [tclIGNORE t] has the same operational content as [t], but drops
+ the returned value. *)
+let tclIGNORE = Proof.ignore
+
+module Monad = Proof
-(* 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
+(** {7 Failure and backtracking} *)
+
+
+(** [tclZERO e] fails with exception [e]. It has no success. *)
+let tclZERO ?info e =
+ let info = match info with
+ | None -> Exninfo.null
+ | Some info -> info
+ in
+ Proof.zero (e, info)
+
+(** [tclOR t1 t2] behaves like [t1] as long as [t1] succeeds. Whenever
+ the successes of [t1] have been depleted and it failed with [e],
+ then it behaves as [t2 e]. In other words, [tclOR] inserts a
+ backtracking point. *)
+let tclOR = Proof.plus
+
+(** [tclORELSE t1 t2] is equal to [t1] if [t1] has at least one
+ success or [t2 e] if [t1] fails with [e]. It is analogous to
+ [try/with] handler of exception in that it is not a backtracking
+ point. *)
+let tclORELSE t1 t2 =
+ let open Logic_monad in
+ let open Proof in
+ split t1 >>= function
+ | Nil e -> t2 e
+ | Cons (a,t1') -> plus (return a) t1'
+
+(** [tclIFCATCH a s f] is a generalisation of {!tclORELSE}: if [a]
+ succeeds at least once then it behaves as [tclBIND a s] otherwise,
+ if [a] fails with [e], then it behaves as [f e]. *)
+let tclIFCATCH a s f =
+ let open Logic_monad in
+ let open Proof in
+ split a >>= function
+ | Nil e -> f e
+ | Cons (x,a') -> plus (s x) (fun e -> (a' e) >>= fun x' -> (s x'))
+
+(** [tclONCE t] behave like [t] except it has at most one success:
+ [tclONCE t] stops after the first success of [t]. If [t] fails
+ with [e], [tclONCE t] also fails with [e]. *)
+let tclONCE = Proof.once
+
+exception MoreThanOneSuccess
let _ = Errors.register_handler begin function
- | SizeMismatch -> Util.error "Incorrect number of goals."
+ | MoreThanOneSuccess -> Errors.error "This tactic has more than one success."
| _ -> raise Errors.Unhandled
end
-(* 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 when Errors.noncritical 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
+(** [tclEXACTLY_ONCE e t] succeeds as [t] if [t] has exactly one
+ success. Otherwise it fails. The tactic [t] is run until its first
+ success, then a failure with exception [e] is simulated. It [t]
+ yields another success, then [tclEXACTLY_ONCE e t] fails with
+ [MoreThanOneSuccess] (it is a user error). Otherwise,
+ [tclEXACTLY_ONCE e t] succeeds with the first success of
+ [t]. Notice that the choice of [e] is relevant, as the presence of
+ further successes may depend on [e] (see {!tclOR}). *)
+let tclEXACTLY_ONCE e t =
+ let open Logic_monad in
+ let open Proof in
+ split t >>= function
+ | Nil (e, info) -> tclZERO ~info e
+ | Cons (x,k) ->
+ Proof.split (k (e, Exninfo.null)) >>= function
+ | Nil _ -> tclUNIT x
+ | _ -> tclZERO MoreThanOneSuccess
-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)
+
+(** [tclCASE t] wraps the {!Proofview_monad.Logical.split} primitive. *)
+type 'a case =
+| Fail of iexn
+| Next of 'a * (iexn -> 'a tactic)
+let tclCASE t =
+ let open Logic_monad in
+ let map = function
+ | Nil e -> Fail e
+ | Cons (x, t) -> Next (x, t)
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
+ Proof.map map (Proof.split t)
+
+let tclBREAK = Proof.break
+
+
+
+(** {7 Focusing tactics} *)
+
+exception NoSuchGoals of int
+
+(* This hook returns a string to be appended to the usual message.
+ Primarily used to add a suggestion about the right bullet to use to
+ focus the next goal, if applicable. *)
+let nosuchgoals_hook:(int -> string option) ref = ref ((fun n -> None))
+let set_nosuchgoals_hook f = nosuchgoals_hook := f
+
+
+
+(* This uses the hook above *)
+let _ = Errors.register_handler begin function
+ | NoSuchGoals n ->
+ let suffix:string option = (!nosuchgoals_hook) n in
+ Errors.errorlabstrm ""
+ (str "No such " ++ str (String.plural n "goal") ++ str "."
+ ++ pr_opt str suffix)
+ | _ -> raise Errors.Unhandled
+end
+
+(** [tclFOCUS_gen nosuchgoal i j t] applies [t] in a context where
+ only the goals numbered [i] to [j] are focused (the rest of the goals
+ is restored at the end of the tactic). If the range [i]-[j] is not
+ valid, then it [tclFOCUS_gen nosuchgoal i j t] is [nosuchgoal]. *)
+let tclFOCUS_gen nosuchgoal i j t =
+ let open Proof in
+ Pv.get >>= fun initial ->
+ try
+ let (focused,context) = focus i j initial in
+ Pv.set focused >>
+ t >>= fun result ->
+ Pv.modify (fun next -> unfocus context next) >>
+ return result
+ with CList.IndexOutOfRange -> nosuchgoal
+
+let tclFOCUS i j t = tclFOCUS_gen (tclZERO (NoSuchGoals (j+1-i))) i j t
+let tclTRYFOCUS i j t = tclFOCUS_gen (tclUNIT ()) i j t
-(* 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
+(** Like {!tclFOCUS} but selects a single goal by name. *)
+let tclFOCUSID id t =
+ let open Proof in
+ Pv.get >>= fun initial ->
+ let rec aux n = function
+ | [] -> tclZERO (NoSuchGoals 1)
+ | g::l ->
+ if Names.Id.equal (Evd.evar_ident g initial.solution) id then
+ let (focused,context) = focus n n initial in
+ Pv.set focused >>
+ t >>= fun result ->
+ Pv.modify (fun next -> unfocus context next) >>
+ return result
+ else
+ aux (n+1) l in
+ aux 1 initial.comb
+
+
+
+(** {7 Dispatching on goals} *)
+
+exception SizeMismatch of int*int
+let _ = Errors.register_handler begin function
+ | SizeMismatch (i,_) ->
+ let open Pp in
+ let errmsg =
+ str"Incorrect number of goals" ++ spc() ++
+ str"(expected "++int i++str(String.plural i " tactic") ++ str")."
+ in
+ Errors.errorlabstrm "" errmsg
+ | _ -> raise Errors.Unhandled
+end
+
+(** A variant of [Monad.List.iter] where we iter over the focused list
+ of goals. The argument tactic is executed in a focus comprising
+ only of the current goal, a goal which has been solved by side
+ effect is skipped. The generated subgoals are concatenated in
+ order. *)
+let iter_goal i =
+ let open Proof in
+ Comb.get >>= fun initial ->
+ Proof.List.fold_left begin fun (subgoals as cur) goal ->
+ Solution.get >>= fun step ->
+ match advance step goal with
+ | None -> return cur
+ | Some goal ->
+ Comb.set [goal] >>
+ i goal >>
+ Proof.map (fun comb -> comb :: subgoals) Comb.get
+ end [] initial >>= fun subgoals ->
+ Solution.get >>= fun evd ->
+ Comb.set CList.(undefined evd (flatten (rev subgoals)))
+
+(** A variant of [Monad.List.fold_left2] where the first list is the
+ list of focused goals. The argument tactic is executed in a focus
+ comprising only of the current goal, a goal which has been solved
+ by side effect is skipped. The generated subgoals are concatenated
+ in order. *)
+let fold_left2_goal i s l =
+ let open Proof in
+ Pv.get >>= fun initial ->
+ let err =
+ return () >>= fun () -> (* Delay the computation of list lengths. *)
+ tclZERO (SizeMismatch (CList.length initial.comb,CList.length l))
+ in
+ Proof.List.fold_left2 err begin fun ((r,subgoals) as cur) goal a ->
+ Solution.get >>= fun step ->
+ match advance step goal with
+ | None -> return cur
+ | Some goal ->
+ Comb.set [goal] >>
+ i goal a r >>= fun r ->
+ Proof.map (fun comb -> (r, comb :: subgoals)) Comb.get
+ end (s,[]) initial.comb l >>= fun (r,subgoals) ->
+ Solution.get >>= fun evd ->
+ Comb.set CList.(undefined evd (flatten (rev subgoals))) >>
+ return r
+
+(** Dispatch tacticals are used to apply a different tactic to each
+ goal under focus. They come in two flavours: [tclDISPATCH] takes a
+ list of [unit tactic]-s and build a [unit tactic]. [tclDISPATCHL]
+ takes a list of ['a tactic] and returns an ['a list tactic].
+
+ They both work by applying each of the tactic in a focus
+ restricted to the corresponding goal (starting with the first
+ goal). In the case of [tclDISPATCHL], the tactic returns a list of
+ the same size as the argument list (of tactics), each element
+ being the result of the tactic executed in the corresponding goal.
+
+ When the length of the tactic list is not the number of goal,
+ raises [SizeMismatch (g,t)] where [g] is the number of available
+ goals, and [t] the number of tactics passed.
+
+ [tclDISPATCHGEN join tacs] generalises both functions as the
+ successive results of [tacs] are stored in reverse order in a
+ list, and [join] is used to convert the result into the expected
+ form. *)
+let tclDISPATCHGEN0 join tacs =
+ match tacs with
+ | [] ->
+ begin
+ let open Proof in
+ Comb.get >>= function
+ | [] -> tclUNIT (join [])
+ | comb -> tclZERO (SizeMismatch (CList.length comb,0))
+ end
+ | [tac] ->
+ begin
+ let open Proof in
+ Pv.get >>= function
+ | { comb=[goal] ; solution } ->
+ begin match advance solution goal with
+ | None -> tclUNIT (join [])
+ | Some _ -> Proof.map (fun res -> join [res]) tac
+ end
+ | {comb} -> tclZERO (SizeMismatch(CList.length comb,1))
+ end
+ | _ ->
+ let iter _ t cur = Proof.map (fun y -> y :: cur) t in
+ let ans = fold_left2_goal iter [] tacs in
+ Proof.map join ans
+
+let tclDISPATCHGEN join tacs =
+ let branch t = InfoL.tag (Info.DBranch) t in
+ let tacs = CList.map branch tacs in
+ InfoL.tag (Info.Dispatch) (tclDISPATCHGEN0 join tacs)
+
+let tclDISPATCH tacs = tclDISPATCHGEN Pervasives.ignore tacs
+
+let tclDISPATCHL tacs = tclDISPATCHGEN CList.rev tacs
+
+
+(** [extend_to_list startxs rx endxs l] builds a list
+ [startxs@[rx,...,rx]@endxs] of the same length as [l]. Raises
+ [SizeMismatch] if [startxs@endxs] is already longer than [l]. *)
+let extend_to_list startxs rx endxs l =
+ (* spiwack: I use [l] essentially as a natural number *)
+ let rec duplicate acc = function
+ | [] -> acc
+ | _::rest -> duplicate (rx::acc) rest
in
- purify begin
- tclDISPATCHGEN Goal.null Goal.plus tacs
+ let rec tail to_match rest =
+ match rest, to_match with
+ | [] , _::_ -> raise (SizeMismatch(0,0)) (* placeholder *)
+ | _::rest , _::to_match -> tail to_match rest
+ | _ , [] -> duplicate endxs rest
+ in
+ let rec copy pref rest =
+ match rest,pref with
+ | [] , _::_ -> raise (SizeMismatch(0,0)) (* placeholder *)
+ | _::rest, a::pref -> a::(copy pref rest)
+ | _ , [] -> tail endxs rest
+ in
+ copy startxs l
+
+(** [tclEXTEND b r e] is a variant of {!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). *)
+let tclEXTEND tacs1 rtac tacs2 =
+ let open Proof in
+ Comb.get >>= fun comb ->
+ try
+ let tacs = extend_to_list tacs1 rtac tacs2 comb in
+ tclDISPATCH tacs
+ with SizeMismatch _ ->
+ tclZERO (SizeMismatch(
+ CList.length comb,
+ (CList.length tacs1)+(CList.length tacs2)))
+(* spiwack: failure occurs only when the number of goals is too
+ small. Hence we can assume that [rtac] is replicated 0 times for
+ any error message. *)
+
+(** [tclEXTEND [] tac []]. *)
+let tclINDEPENDENT tac =
+ let open Proof in
+ Pv.get >>= fun initial ->
+ match initial.comb with
+ | [] -> tclUNIT ()
+ | [_] -> tac
+ | _ ->
+ let tac = InfoL.tag (Info.DBranch) tac in
+ InfoL.tag (Info.Dispatch) (iter_goal (fun _ -> tac))
+
+
+
+(** {7 Goal manipulation} *)
+
+(** Shelves all the goals under focus. *)
+let shelve =
+ let open Proof in
+ Comb.get >>= fun initial ->
+ Comb.set [] >>
+ InfoL.leaf (Info.Tactic (fun () -> Pp.str"shelve")) >>
+ Shelf.put initial
+
+
+(** [contained_in_info e evi] checks whether the evar [e] appears in
+ the hypotheses, the conclusion or the body of the evar_info
+ [evi]. Note: since we want to use it on goals, the body is actually
+ supposed to be empty. *)
+let contained_in_info sigma e evi =
+ Evar.Set.mem e (Evd.evars_of_filtered_evar_info (Evarutil.nf_evar_info sigma evi))
+
+(** [depends_on sigma src tgt] checks whether the goal [src] appears
+ as an existential variable in the definition of the goal [tgt] in
+ [sigma]. *)
+let depends_on sigma src tgt =
+ let evi = Evd.find sigma tgt in
+ contained_in_info sigma src evi
+
+(** [unifiable sigma g l] checks whether [g] appears in another
+ subgoal of [l]. The list [l] may contain [g], but it does not
+ affect the result. *)
+let unifiable sigma g l =
+ CList.exists (fun tgt -> not (Evar.equal g tgt) && depends_on sigma g tgt) l
+
+(** [partition_unifiable sigma l] partitions [l] into a pair [(u,n)]
+ where [u] is composed of the unifiable goals, i.e. the goals on
+ whose definition other goals of [l] depend, and [n] are the
+ non-unifiable goals. *)
+let partition_unifiable sigma l =
+ CList.partition (fun g -> unifiable sigma g l) l
+
+(** Shelves the unifiable goals under focus, i.e. the goals which
+ appear in other goals under focus (the unfocused goals are not
+ considered). *)
+let shelve_unifiable =
+ let open Proof in
+ Pv.get >>= fun initial ->
+ let (u,n) = partition_unifiable initial.solution initial.comb in
+ Comb.set n >>
+ InfoL.leaf (Info.Tactic (fun () -> Pp.str"shelve_unifiable")) >>
+ Shelf.put u
+
+(** [guard_no_unifiable] fails with error [UnresolvedBindings] if some
+ goals are unifiable (see {!shelve_unifiable}) in the current focus. *)
+let guard_no_unifiable =
+ let open Proof in
+ Pv.get >>= fun initial ->
+ let (u,n) = partition_unifiable initial.solution initial.comb in
+ match u with
+ | [] -> tclUNIT ()
+ | gls ->
+ let l = CList.map (fun g -> Evd.dependent_evar_ident g initial.solution) gls in
+ let l = CList.map (fun id -> Names.Name id) l in
+ tclZERO (Logic.RefinerError (Logic.UnresolvedBindings l))
+
+(** [unshelve l p] adds all the goals in [l] at the end of the focused
+ goals of p *)
+let unshelve l p =
+ (* advance the goals in case of clear *)
+ let l = undefined p.solution l in
+ { p with comb = p.comb@l }
+
+
+(** [goodmod p m] computes the representative of [p] modulo [m] in the
+ interval [[0,m-1]].*)
+let goodmod p m =
+ let p' = p mod m in
+ (* if [n] is negative [n mod l] is negative of absolute value less
+ than [l], so [(n mod l)+l] is the representative of [n] in the
+ interval [[0,l-1]].*)
+ if p' < 0 then p'+m else p'
+
+let cycle n =
+ let open Proof in
+ InfoL.leaf (Info.Tactic (fun () -> Pp.(str"cycle"++spc()++int n))) >>
+ Comb.modify begin fun initial ->
+ let l = CList.length initial in
+ let n' = goodmod n l in
+ let (front,rear) = CList.chop n' initial in
+ rear@front
end
-let rec tclGOALBINDU s k =
- purify begin
- tclBIND (list_of_sensitive s k) begin fun tacs ->
- tclDISPATCHGEN () unitK tacs
- end
+let swap i j =
+ let open Proof in
+ InfoL.leaf (Info.Tactic (fun () -> Pp.(str"swap"++spc()++int i++spc()++int j))) >>
+ Comb.modify begin fun initial ->
+ let l = CList.length initial in
+ let i = if i>0 then i-1 else i and j = if j>0 then j-1 else j in
+ let i = goodmod i l and j = goodmod j l in
+ CList.map_i begin fun k x ->
+ match k with
+ | k when Int.equal k i -> CList.nth initial j
+ | k when Int.equal k j -> CList.nth initial i
+ | _ -> x
+ end 0 initial
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
- | Loc.Exc_located(_,e) -> catchable_exception e
- | Util.UserError _
- | Type_errors.TypeError _ | PretypeError (_,_,TypingError _)
- | 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,[])
+let revgoals =
+ let open Proof in
+ InfoL.leaf (Info.Tactic (fun () -> Pp.str"revgoals")) >>
+ Comb.modify CList.rev
+
+let numgoals =
+ let open Proof in
+ Comb.get >>= fun comb ->
+ return (CList.length comb)
+
+
+
+(** {7 Access primitives} *)
+
+let tclEVARMAP = Solution.get
+
+let tclENV = Env.get
+
+
+
+(** {7 Put-like primitives} *)
+
+
+let emit_side_effects eff x =
+ { x with solution = Evd.emit_side_effects eff x.solution }
+
+let tclEFFECTS eff =
+ let open Proof in
+ return () >>= fun () -> (* The Global.env should be taken at exec time *)
+ Env.set (Global.env ()) >>
+ Pv.modify (fun initial -> emit_side_effects eff initial)
+
+let mark_as_unsafe = Status.put false
+
+(** Gives up on the goal under focus. Reports an unsafe status. Proofs
+ with given up goals cannot be closed. *)
+let give_up =
+ let open Proof in
+ Comb.get >>= fun initial ->
+ Comb.set [] >>
+ mark_as_unsafe >>
+ InfoL.leaf (Info.Tactic (fun () -> Pp.str"give_up")) >>
+ Giveup.put initial
+
+
+
+(** {7 Control primitives} *)
+
+(** Equality function on goals *)
+let goal_equal evars1 gl1 evars2 gl2 =
+ let evi1 = Evd.find evars1 gl1 in
+ let evi2 = Evd.find evars2 gl2 in
+ Evd.eq_evar_info evars2 evi1 evi2
+
+let tclPROGRESS t =
+ let open Proof in
+ Pv.get >>= fun initial ->
+ t >>= fun res ->
+ Pv.get >>= fun final ->
+ let test =
+ Evd.progress_evar_map initial.solution final.solution &&
+ not (Util.List.for_all2eq (fun i f -> goal_equal initial.solution i final.solution f) initial.comb final.comb)
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
+ if test then
+ tclUNIT res
+ else
+ tclZERO (Errors.UserError ("Proofview.tclPROGRESS" , Pp.str"Failed to progress."))
+
+exception Timeout
+let _ = Errors.register_handler begin function
+ | Timeout -> Errors.errorlabstrm "Proofview.tclTIMEOUT" (Pp.str"Tactic timeout!")
+ | _ -> Pervasives.raise Errors.Unhandled
+end
+
+let tclTIMEOUT n t =
+ let open Proof in
+ (* spiwack: as one of the monad is a continuation passing monad, it
+ doesn't force the computation to be threaded inside the underlying
+ (IO) monad. Hence I force it myself by asking for the evaluation of
+ a dummy value first, lest [timeout] be called when everything has
+ already been computed. *)
+ let t = Proof.lift (Logic_monad.NonLogical.return ()) >> t in
+ Proof.get >>= fun initial ->
+ Proof.current >>= fun envvar ->
+ Proof.lift begin
+ Logic_monad.NonLogical.catch
+ begin
+ let open Logic_monad.NonLogical in
+ timeout n (Proof.repr (Proof.run t envvar initial)) >>= fun r ->
+ match r with
+ | Logic_monad.Nil e -> return (Util.Inr e)
+ | Logic_monad.Cons (r, _) -> return (Util.Inl r)
+ end
+ begin let open Logic_monad.NonLogical in function (e, info) ->
+ match e with
+ | Logic_monad.Timeout -> return (Util.Inr (Timeout, info))
+ | Logic_monad.TacticFailure e ->
+ return (Util.Inr (e, info))
+ | e -> Logic_monad.NonLogical.raise ~info e
+ end
+ end >>= function
+ | Util.Inl (res,s,m,i) ->
+ Proof.set s >>
+ Proof.put m >>
+ Proof.update (fun _ -> i) >>
+ return res
+ | Util.Inr (e, info) -> tclZERO ~info e
+
+let tclTIME s t =
+ let pr_time t1 t2 n msg =
+ let msg =
+ if n = 0 then
+ str msg
+ else
+ str (msg ^ " after ") ++ int n ++ str (String.plural n " backtracking")
+ in
+ msg_info(str "Tactic call" ++ pr_opt str s ++ str " ran for " ++
+ System.fmt_time_difference t1 t2 ++ str " " ++ surround msg) in
+ let rec aux n t =
+ let open Proof in
+ tclUNIT () >>= fun () ->
+ let tstart = System.get_time() in
+ Proof.split t >>= let open Logic_monad in function
+ | Nil (e, info) ->
+ begin
+ let tend = System.get_time() in
+ pr_time tstart tend n "failure";
+ tclZERO ~info e
+ end
+ | Cons (x,k) ->
+ let tend = System.get_time() in
+ pr_time tstart tend n "success";
+ tclOR (tclUNIT x) (fun e -> aux (n+1) (k e))
+ in aux 0 t
+
-(*** Commands ***)
+(** {7 Unsafe primitives} *)
-let in_proofview p k =
- k p.solution
+module Unsafe = struct
+
+ let tclEVARS evd =
+ Pv.modify (fun ps -> { ps with solution = evd })
+
+ let tclNEWGOALS gls =
+ Pv.modify begin fun step ->
+ let gls = undefined step.solution gls in
+ { step with comb = step.comb @ gls }
+ end
+
+ let tclGETGOALS = Comb.get
+
+ let tclSETGOALS = Comb.set
+
+ let tclEVARSADVANCE evd =
+ Pv.modify (fun ps -> { solution = evd; comb = undefined evd ps.comb })
+
+ let tclEVARUNIVCONTEXT ctx =
+ Pv.modify (fun ps -> { ps with solution = Evd.set_universe_context ps.solution ctx })
+
+ let reset_future_goals p =
+ { p with solution = Evd.reset_future_goals p.solution }
+
+ let mark_as_goal_evm evd content =
+ let info = Evd.find evd content in
+ let info =
+ { info with Evd.evar_source = match info.Evd.evar_source with
+ | _, (Evar_kinds.VarInstance _ | Evar_kinds.GoalEvar) as x -> x
+ | loc,_ -> loc,Evar_kinds.GoalEvar }
+ in
+ let info = Typeclasses.mark_unresolvable info in
+ Evd.add evd content info
+
+ let mark_as_goal p gl =
+ { p with solution = mark_as_goal_evm p.solution gl }
+
+end
+
+
+
+(** {7 Notations} *)
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 (>>=) = tclBIND
let (<*>) = tclTHEN
- let (<+>) = tclOR
+ let (<+>) t1 t2 = tclOR t1 (fun _ -> t2)
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 )
+open Notations
+
+
+
+(** {6 Goal-dependent tactics} *)
+
+(* To avoid shadowing by the local [Goal] module *)
+module GoalV82 = Goal.V82
+
+let catchable_exception = function
+ | Logic_monad.Exception _ -> false
+ | e -> Errors.noncritical e
+
+
+module Goal = struct
+
+ type 'a t = {
+ env : Environ.env;
+ sigma : Evd.evar_map;
+ concl : Term.constr ;
+ self : Evar.t ; (* for compatibility with old-style definitions *)
+ }
+
+ let assume (gl : 'a t) = (gl :> [ `NF ] t)
+
+ let env { env=env } = env
+ let sigma { sigma=sigma } = sigma
+ let hyps { env=env } = Environ.named_context env
+ let concl { concl=concl } = concl
+ let extra { sigma=sigma; self=self } = Goal.V82.extra sigma self
+
+ let raw_concl { concl=concl } = concl
+
+
+ let gmake_with info env sigma goal =
+ { env = Environ.reset_with_named_context (Evd.evar_filtered_hyps info) env ;
+ sigma = sigma ;
+ concl = Evd.evar_concl info ;
+ self = goal }
+
+ let nf_gmake env sigma goal =
+ let info = Evarutil.nf_evar_info sigma (Evd.find sigma goal) in
+ let sigma = Evd.add sigma goal info in
+ gmake_with info env sigma goal , sigma
+
+ let nf_enter f =
+ InfoL.tag (Info.Dispatch) begin
+ iter_goal begin fun goal ->
+ Env.get >>= fun env ->
+ tclEVARMAP >>= fun sigma ->
+ try
+ let (gl, sigma) = nf_gmake env sigma goal in
+ tclTHEN (Unsafe.tclEVARS sigma) (InfoL.tag (Info.DBranch) (f gl))
+ with e when catchable_exception e ->
+ let (e, info) = Errors.push e in
+ tclZERO ~info e
+ end
+ end
+
+ let normalize { self } =
+ Env.get >>= fun env ->
+ tclEVARMAP >>= fun sigma ->
+ let (gl,sigma) = nf_gmake env sigma self in
+ tclTHEN (Unsafe.tclEVARS sigma) (tclUNIT gl)
+
+ let gmake env sigma goal =
+ let info = Evd.find sigma goal in
+ gmake_with info env sigma goal
+
+ let enter f =
+ let f gl = InfoL.tag (Info.DBranch) (f gl) in
+ InfoL.tag (Info.Dispatch) begin
+ iter_goal begin fun goal ->
+ Env.get >>= fun env ->
+ tclEVARMAP >>= fun sigma ->
+ try f (gmake env sigma goal)
+ with e when catchable_exception e ->
+ let (e, info) = Errors.push e in
+ tclZERO ~info e
+ end
+ end
+
+ let goals =
+ Env.get >>= fun env ->
+ Pv.get >>= fun step ->
+ let sigma = step.solution in
+ let map goal =
+ match advance sigma goal with
+ | None -> None (** ppedrot: Is this check really necessary? *)
+ | Some goal ->
+ let gl =
+ tclEVARMAP >>= fun sigma ->
+ tclUNIT (gmake env sigma goal)
+ in
+ Some gl
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
+ tclUNIT (CList.map_filter map step.comb)
+
+ (* compatibility *)
+ let goal { self=self } = self
+
+end
+
+
+
+(** {6 The refine tactic} *)
+
+module Refine =
+struct
+
+ let typecheck_evar ev env sigma =
+ let info = Evd.find sigma ev in
+ let evdref = ref sigma in
+ let env = Environ.reset_with_named_context (Evd.evar_hyps info) env in
+ let _ = Typing.sort_of env evdref (Evd.evar_concl info) in
+ !evdref
+
+ let typecheck_proof c concl env sigma =
+ let evdref = ref sigma in
+ let () = Typing.check env evdref c concl in
+ !evdref
+
+ let (pr_constrv,pr_constr) =
+ Hook.make ~default:(fun _env _sigma _c -> Pp.str"<constr>") ()
+
+ let refine ?(unsafe = true) f = Goal.enter begin fun gl ->
+ let sigma = Goal.sigma gl in
+ let env = Goal.env gl in
+ let concl = Goal.concl gl in
+ (** Save the [future_goals] state to restore them after the
+ refinement. *)
+ let prev_future_goals = Evd.future_goals sigma in
+ let prev_principal_goal = Evd.principal_future_goal sigma in
+ (** Create the refinement term *)
+ let (sigma, c) = f (Evd.reset_future_goals sigma) in
+ let evs = Evd.future_goals sigma in
+ let evkmain = Evd.principal_future_goal sigma in
+ (** Check that the introduced evars are well-typed *)
+ let fold accu ev = typecheck_evar ev env accu in
+ let sigma = if unsafe then sigma else CList.fold_left fold sigma evs in
+ (** Check that the refined term is typesafe *)
+ let sigma = if unsafe then sigma else typecheck_proof c concl env sigma in
+ (** Check that the goal itself does not appear in the refined term *)
+ let _ =
+ if not (Evarutil.occur_evar_upto sigma gl.Goal.self c) then ()
+ else Pretype_errors.error_occur_check env sigma gl.Goal.self c
in
- let (goalss,evd) = Goal.list_map tac initgoals initevd in
- let sgs = List.flatten goalss in
- sk () fk { defs = evd ; goals = sgs }
-}
+ (** Proceed to the refinement *)
+ let sigma = match evkmain with
+ | None -> Evd.define gl.Goal.self c sigma
+ | Some evk ->
+ let id = Evd.evar_ident gl.Goal.self sigma in
+ Evd.rename evk id (Evd.define gl.Goal.self c sigma)
+ in
+ (** Restore the [future goals] state. *)
+ let sigma = Evd.restore_future_goals sigma prev_future_goals prev_principal_goal in
+ (** Select the goals *)
+ let comb = undefined sigma (CList.rev evs) in
+ let sigma = CList.fold_left Unsafe.mark_as_goal_evm sigma comb in
+ let open Proof in
+ InfoL.leaf (Info.Tactic (fun () -> Pp.(str"refine"++spc()++ Hook.get pr_constrv env sigma c))) >>
+ Pv.set { solution = sigma; comb; }
+ end
+
+ (** Useful definitions *)
+
+ let with_type env evd c t =
+ let my_type = Retyping.get_type_of env evd c in
+ let j = Environ.make_judge c my_type in
+ let (evd,j') =
+ Coercion.inh_conv_coerce_to true (Loc.ghost) env evd j t
+ in
+ evd , j'.Environ.uj_val
+
+ let refine_casted ?unsafe f = Goal.enter begin fun gl ->
+ let concl = Goal.concl gl in
+ let env = Goal.env gl in
+ let f h = let (h, c) = f h in with_type env h c concl in
+ refine ?unsafe f
+ end
+end
+
+
+
+(** {6 Trace} *)
+
+module Trace = struct
+
+ let record_info_trace = InfoL.record_trace
+ let log m = InfoL.leaf (Info.Msg m)
+ let name_tactic m t = InfoL.tag (Info.Tactic m) t
+
+ let pr_info ?(lvl=0) info =
+ assert (lvl >= 0);
+ Info.(print (collapse lvl info))
+
+end
+
+
+
+(** {6 Non-logical state} *)
+
+module NonLogical = Logic_monad.NonLogical
+
+let tclLIFT = Proof.lift
+
+let tclCHECKINTERRUPT =
+ tclLIFT (NonLogical.make Control.check_for_interrupt)
+
+
+
+
+
+(*** Compatibility layer with <= 8.2 tactics ***)
+module V82 = struct
+ type tac = Evar.t Evd.sigma -> Evar.t list Evd.sigma
+
+ let tactic tac =
+ (* spiwack: we ignore the dependencies between goals here,
+ expectingly preserving the semantics of <= 8.2 tactics *)
+ (* spiwack: convenience notations, waiting for ocaml 3.12 *)
+ let open Proof in
+ Pv.get >>= fun ps ->
+ try
+ let tac gl evd =
+ let glsigma =
+ tac { Evd.it = gl ; 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) =
+ Evd.Monad.List.map (fun g s -> GoalV82.nf_evar s g) ps.comb ps.solution
+ in
+ let (goalss,evd) = Evd.Monad.List.map tac initgoals initevd in
+ let sgs = CList.flatten goalss in
+ let sgs = undefined evd sgs in
+ InfoL.leaf (Info.Tactic (fun () -> Pp.str"<unknown>")) >>
+ Pv.set { solution = evd; comb = sgs; }
+ with e when catchable_exception e ->
+ let (e, info) = Errors.push e in
+ tclZERO ~info e
+
+
+ (* normalises the evars in the goals, and stores the result in
+ solution. *)
+ let nf_evar_goals =
+ Pv.modify begin fun ps ->
+ let map g s = GoalV82.nf_evar s g in
+ let (goals,evd) = Evd.Monad.List.map map ps.comb ps.solution in
+ { solution = evd; comb = goals; }
+ end
+
let has_unresolved_evar pv =
Evd.has_undefined pv.solution
(* Main function in the implementation of Grab Existential Variables.*)
let grab pv =
- let goals =
- List.map begin fun (e,_) ->
- Goal.build e
- end (Evd.undefined_list pv.solution)
- in
+ let undef = Evd.undefined_map pv.solution in
+ let goals = CList.rev_map fst (Evar.Map.bindings undef) in
{ pv with comb = goals }
(* 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 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 top_goals initial { solution=solution; } =
+ let goals = CList.map (fun (t,_) -> fst (Term.destEvar t)) initial in
+ { Evd.it = goals ; sigma=solution; }
- let top_evars { initial=initial } =
+ let top_evars initial =
let evars_of_initial (c,_) =
- Util.Intset.elements (Evarutil.evars_of_term c)
+ Evar.Set.elements (Evd.evars_of_term c)
in
- List.flatten (List.map evars_of_initial initial)
+ CList.flatten (CList.map evars_of_initial initial)
let instantiate_evar n com pv =
- let (evk,_) =
+ 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"
+ let evl = Evar.Map.bindings evl in
+ if (n <= 0) then
+ Errors.error "incorrect existential variable index"
+ else if CList.length evl < n then
+ Errors.error "not so many uninstantiated existential variables"
else
- List.nth evl (n-1)
+ CList.nth evl (n-1)
in
{ pv with
solution = Evar_refiner.instantiate_pf_com evk com pv.solution }
- let purify = purify
+ let of_tactic t gls =
+ try
+ let init = { solution = gls.Evd.sigma ; comb = [gls.Evd.it] } in
+ let (_,final,_,_) = apply (GoalV82.env gls.Evd.sigma gls.Evd.it) t init in
+ { Evd.sigma = final.solution ; it = final.comb }
+ with Logic_monad.TacticFailure e as src ->
+ let (_, info) = Errors.push src in
+ iraise (e, info)
+
+ let put_status = Status.put
+
+ let catchable_exception = catchable_exception
+
+ let wrap_exceptions f =
+ try f ()
+ with e when catchable_exception e ->
+ let (e, info) = Errors.push e in tclZERO ~info e
+
end
diff --git a/proofs/proofview.mli b/proofs/proofview.mli
index a87383c8..ec255f6a 100644
--- a/proofs/proofview.mli
+++ b/proofs/proofview.mli
@@ -1,32 +1,25 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* 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. *)
+(** This files defines the basic mechanism of proofs: the [proofview]
+ type is the state which tactics manipulate (a global state for
+ existential variables, together with the list of goals), and the type
+ ['a tactic] is the (abstract) type of tactics modifying the proof
+ state and returning a value of type ['a]. *)
+open Util
open Term
-type proofview
+(** Main state of tactics *)
+type proofview
-
-(* Returns a stylised view of a proofview for use by, for instance,
- ide-s. *)
+(** Returns a stylised view of a proofview for use by, for instance,
+ ide-s. *)
(* spiwack: the type of [proofview] will change as we push more
refined functions to ide-s. This would be better than spawning a
new nearly identical function everytime. Hence the generic name. *)
@@ -34,183 +27,518 @@ type proofview
the [evar_map] context. *)
val proofview : proofview -> Goal.goal list * Evd.evar_map
-(* 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. *)
+(** {6 Starting and querying a proof view} *)
+
+(** Abstract representation of the initial goals of a proof. *)
+type entry
+
+(** Optimize memory consumption *)
+val compact : entry -> proofview -> entry * proofview
+
+(** Initialises a proofview, the main argument is a list of
+ environements (including a [named_context] which are used as
+ hypotheses) pair with conclusion types, creating accordingly many
+ initial goals. Because a proof does not necessarily starts in an
+ empty [evar_map] (indeed a proof can be triggered by an incomplete
+ pretyping), [init] takes an additional argument to represent the
+ initial [evar_map]. *)
+val init : Evd.evar_map -> (Environ.env * Term.types) list -> entry * proofview
+
+(** A [telescope] is a list of environment and conclusion like in
+ {!init}, except that each element may depend on the previous
+ goals. The telescope passes the goals in the form of a
+ [Term.constr] which represents the goal as an [evar]. The
+ [evar_map] is threaded in state passing style. *)
+type telescope =
+ | TNil of Evd.evar_map
+ | TCons of Environ.env * Evd.evar_map * Term.types * (Evd.evar_map -> Term.constr -> telescope)
+
+(** Like {!init}, but goals are allowed to be dependent on one
+ another. Dependencies between goals is represented with the type
+ [telescope] instead of [list]. Note that the first [evar_map] of
+ the telescope plays the role of the [evar_map] argument in
+ [init]. *)
+val dependent_init : telescope -> entry * proofview
+
+(** [finished pv] is [true] if and only if [pv] is complete. That is,
+ if it has an empty list of focused goals. There could still be
+ unsolved subgoaled, but they would then be out of focus. *)
val finished : proofview -> bool
-(* Returns the current value of the proofview partial proofs. *)
-val return : proofview -> (constr*types) list
+(** Returns the current [evar] state. *)
+val return : proofview -> Evd.evar_map
+
+val partial_proof : entry -> proofview -> constr list
+val initial_goals : entry -> (constr * types) list
-(*** Focusing operations ***)
-(* [IndexOutOfRange] occurs in case of malformed indices
- with respect to list lengths. *)
-exception IndexOutOfRange
+(** {6 Focusing commands} *)
-(* Type of the object which allow to unfocus a view.*)
+(** A [focus_context] represents the part of the proof view which has
+ been removed by a focusing action, it can be used to unfocus later
+ on. *)
type focus_context
-(* Returns a stylised view of a focus_context for use by, for
- instance, ide-s. *)
+(** Returns a stylised view of a focus_context for use by, for
+ instance, ide-s. *)
(* spiwack: the type of [focus_context] will change as we push more
refined functions to ide-s. This would be better than spawning a
new nearly identical function everytime. Hence the generic name. *)
-(* In this version: returns the number of goals that are held *)
+(* In this version: the goals in the context, as a "zipper" (the first
+ list is in reversed order). *)
val focus_context : focus_context -> Goal.goal list * Goal.goal list
-(* [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. *)
+(** [focus i j] focuses a proofview on the goals from index [i] to
+ index [j] (inclusive, goals are indexed from [1]). I.e. goals
+ number [i] to [j] become the only focused goals of the returned
+ proofview. It returns the focused proofview, and a context for
+ the focus stack. *)
val focus : int -> int -> proofview -> proofview * focus_context
-(* Unfocuses a proofview with respect to a 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).
+
+(** {6 The tactic monad} *)
+
+(** - Tactics are objects which apply a transformation to all the
+ subgoals of the current view at the same time. By opposition to
+ the old vision of applying it to a single goal. It allows tactics
+ such as [shelve_unifiable], tactics to reorder the focused goals,
+ 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). Another benefit is that it is possible to write tactics
+ that can be executed even if there are no focused goals.
+ - Tactics form 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 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") and [Proofview.tclTHEN]
+ (which is a specialized bind on unit-returning tactics).
+ - Tactics have support for full-backtracking. Tactics can be seen
+ having multiple success: if after returning the first success a
+ failure is encountered, the tactic can backtrack and use a second
+ success if available. The state is backtracked to its previous
+ value, except the non-logical state defined in the {!NonLogical}
+ module below.
*)
+(** The abstract type of 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 *)
+(** Applies a tactic to the current proofview. Returns a tuple
+ [a,pv,(b,sh,gu)] where [a] is the return value of the tactic, [pv]
+ is the updated proofview, [b] a boolean which is [true] if the
+ tactic has not done any action considered unsafe (such as
+ admitting a lemma), [sh] is the list of goals which have been
+ shelved by the tactic, and [gu] the list of goals on which the
+ tactic has given up. In case of multiple success the first one is
+ selected. If there is no success, fails with
+ {!Logic_monad.TacticFailure}*)
+val apply : Environ.env -> 'a tactic -> proofview -> 'a
+ * proofview
+ * (bool*Goal.goal list*Goal.goal list)
+ * Proofview_monad.Info.tree
+
+(** {7 Monadic primitives} *)
+
+(** Unit of the tactic monad. *)
val tclUNIT : 'a -> 'a tactic
-
-(* Bind operation of the tactic monad *)
+
+(** Bind operation of the tactic monad. *)
val tclBIND : 'a tactic -> ('a -> 'b tactic) -> 'b tactic
-(* Interprets the ";" (semicolon) of Ltac.
- As a monadic operation, it's a specialized "bind"
- on unit-returning tactic (meaning "there is no value to bind") *)
+(** Interprets the ";" (semicolon) of Ltac. As a monadic operation,
+ it's a specialized "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. *)
+(** [tclIGNORE t] has the same operational content as [t], but drops
+ the returned value. *)
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
+(** Generic monadic combinators for tactics. *)
+module Monad : Monad.S with type +'a t = 'a tactic
+
+(** {7 Failure and backtracking} *)
+
+(** [tclZERO e] fails with exception [e]. It has no success. *)
+val tclZERO : ?info:Exninfo.info -> exn -> 'a tactic
+
+(** [tclOR t1 t2] behaves like [t1] as long as [t1] succeeds. Whenever
+ the successes of [t1] have been depleted and it failed with [e],
+ then it behaves as [t2 e]. In other words, [tclOR] inserts a
+ backtracking point. *)
+val tclOR : 'a tactic -> (iexn -> 'a tactic) -> 'a tactic
+
+(** [tclORELSE t1 t2] is equal to [t1] if [t1] has at least one
+ success or [t2 e] if [t1] fails with [e]. It is analogous to
+ [try/with] handler of exception in that it is not a backtracking
+ point. *)
+val tclORELSE : 'a tactic -> (iexn -> 'a tactic) -> 'a tactic
+
+(** [tclIFCATCH a s f] is a generalisation of {!tclORELSE}: if [a]
+ succeeds at least once then it behaves as [tclBIND a s] otherwise,
+ if [a] fails with [e], then it behaves as [f e]. *)
+val tclIFCATCH : 'a tactic -> ('a -> 'b tactic) -> (iexn -> 'b tactic) -> 'b tactic
+
+(** [tclONCE t] behave like [t] except it has at most one success:
+ [tclONCE t] stops after the first success of [t]. If [t] fails
+ with [e], [tclONCE t] also fails with [e]. *)
+val tclONCE : 'a tactic -> 'a tactic
+
+(** [tclEXACTLY_ONCE e t] succeeds as [t] if [t] has exactly one
+ success. Otherwise it fails. The tactic [t] is run until its first
+ success, then a failure with exception [e] is simulated. It [t]
+ yields another success, then [tclEXACTLY_ONCE e t] fails with
+ [MoreThanOneSuccess] (it is a user error). Otherwise,
+ [tclEXACTLY_ONCE e t] succeeds with the first success of
+ [t]. Notice that the choice of [e] is relevant, as the presence of
+ further successes may depend on [e] (see {!tclOR}). *)
+exception MoreThanOneSuccess
+val tclEXACTLY_ONCE : exn -> 'a tactic -> 'a tactic
+
+(** [tclCASE t] splits [t] into its first success and a
+ continuation. It is the most general primitive to control
+ backtracking. *)
+type 'a case =
+ | Fail of iexn
+ | Next of 'a * (iexn -> 'a tactic)
+val tclCASE : 'a tactic -> 'a case tactic
+
+(** [tclBREAK p t] is a generalization of [tclONCE t]. Instead of
+ stopping after the first success, it succeeds like [t] until a
+ failure with an exception [e] such that [p e = Some e'] is raised. At
+ which point it drops the remaining successes, failing with [e'].
+ [tclONCE t] is equivalent to [tclBREAK (fun e -> Some e) t]. *)
+val tclBREAK : (iexn -> iexn option) -> 'a tactic -> 'a tactic
+
+
+(** {7 Focusing tactics} *)
+
+(** [tclFOCUS i j t] applies [t] after focusing on the goals number
+ [i] to [j] (see {!focus}). The rest of the goals is restored after
+ the tactic action. If the specified range doesn't correspond to
+ existing goals, fails with [NoSuchGoals] (a user error). this
+ exception is catched at toplevel with a default message + a hook
+ message that can be customized by [set_nosuchgoals_hook] below.
+ This hook is used to add a suggestion about bullets when
+ applicable. *)
+exception NoSuchGoals of int
+val set_nosuchgoals_hook: (int -> string option) -> unit
-(* [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. *)
+(** [tclFOCUSID x t] applies [t] on a (single) focused goal like
+ {!tclFOCUS}. The goal is found by its name rather than its
+ number.*)
+val tclFOCUSID : Names.Id.t -> 'a tactic -> 'a tactic
+
+(** [tclTRYFOCUS i j t] behaves like {!tclFOCUS}, except that if the
+ specified range doesn't correspond to existing goals, behaves like
+ [tclUNIT ()] instead of failing. *)
+val tclTRYFOCUS : int -> int -> unit tactic -> unit tactic
+
+
+(** {7 Dispatching on goals} *)
+
+(** Dispatch tacticals are used to apply a different tactic to each
+ goal under focus. They come in two flavours: [tclDISPATCH] takes a
+ list of [unit tactic]-s and build a [unit tactic]. [tclDISPATCHL]
+ takes a list of ['a tactic] and returns an ['a list tactic].
+
+ They both work by applying each of the tactic in a focus
+ restricted to the corresponding goal (starting with the first
+ goal). In the case of [tclDISPATCHL], the tactic returns a list of
+ the same size as the argument list (of tactics), each element
+ being the result of the tactic executed in the corresponding goal.
+
+ When the length of the tactic list is not the number of goal,
+ raises [SizeMismatch (g,t)] where [g] is the number of available
+ goals, and [t] the number of tactics passed. *)
+exception SizeMismatch of int*int
val tclDISPATCH : unit tactic list -> unit tactic
-val tclDISPATCHS : 'a Goal.sensitive tactic list -> 'a Goal.sensitive tactic
+val tclDISPATCHL : 'a tactic list -> 'a list 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. *)
+(** [tclEXTEND b r e] is a variant of {!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
+(** [tclINDEPENDENT tac] runs [tac] on each goal successively, from
+ the first one to the last one. Backtracking in one goal is
+ independent of backtracking in another. It is equivalent to
+ [tclEXTEND [] tac []]. *)
+val tclINDEPENDENT : unit tactic -> unit tactic
+
+
+(** {7 Goal manipulation} *)
+
+(** Shelves all the goals under focus. The goals are placed on the
+ shelf for later use (or being solved by side-effects). *)
+val shelve : unit tactic
+
+(** Shelves the unifiable goals under focus, i.e. the goals which
+ appear in other goals under focus (the unfocused goals are not
+ considered). *)
+val shelve_unifiable : unit tactic
+
+(** [guard_no_unifiable] fails with error [UnresolvedBindings] if some
+ goals are unifiable (see {!shelve_unifiable}) in the current focus. *)
+val guard_no_unifiable : unit tactic
+
+(** [unshelve l p] adds all the goals in [l] at the end of the focused
+ goals of p *)
+val unshelve : Goal.goal list -> proofview -> proofview
+
+(** If [n] is positive, [cycle n] puts the [n] first goal last. If [n]
+ is negative, then it puts the [n] last goals first.*)
+val cycle : int -> unit tactic
+
+(** [swap i j] swaps the position of goals number [i] and [j]
+ (negative numbers can be used to address goals from the end. Goals
+ are indexed from [1]. For simplicity index [0] corresponds to goal
+ [1] as well, rather than raising an error. *)
+val swap : int -> int -> unit tactic
+
+(** [revgoals] reverses the list of focused goals. *)
+val revgoals : unit tactic
-(* [tclSENSITIVE] views goal-type tactics as a special kind of tactics.*)
-val tclSENSITIVE : Goal.subgoals Goal.sensitive -> unit tactic
+(** [numgoals] returns the number of goals under focus. *)
+val numgoals : int tactic
-(*** Commands ***)
+(** {7 Access primitives} *)
-val in_proofview : proofview -> (Evd.evar_map -> 'a) -> 'a
+(** [tclEVARMAP] doesn't affect the proof, it returns the current
+ [evar_map]. *)
+val tclEVARMAP : Evd.evar_map tactic
+(** [tclENV] doesn't affect the proof, it returns the current
+ environment. It is not the environment of a particular goal,
+ rather the "global" environment of the proof. The goal-wise
+ environment is obtained via {!Proofview.Goal.env}. *)
+val tclENV : Environ.env tactic
+(** {7 Put-like primitives} *)
+(** [tclEFFECTS eff] add the effects [eff] to the current state. *)
+val tclEFFECTS : Declareops.side_effects -> unit tactic
+
+(** [mark_as_unsafe] declares the current tactic is unsafe. *)
+val mark_as_unsafe : unit tactic
+
+(** Gives up on the goal under focus. Reports an unsafe status. Proofs
+ with given up goals cannot be closed. *)
+val give_up : unit tactic
+
+
+(** {7 Control primitives} *)
+
+(** [tclPROGRESS t] checks the state of the proof after [t]. It it is
+ identical to the state before, then [tclePROGRESS t] fails, otherwise
+ it succeeds like [t]. *)
+val tclPROGRESS : 'a tactic -> 'a tactic
+
+(** Checks for interrupts *)
+val tclCHECKINTERRUPT : unit tactic
+
+exception Timeout
+(** [tclTIMEOUT n t] can have only one success.
+ In case of timeout if fails with [tclZERO Timeout]. *)
+val tclTIMEOUT : int -> 'a tactic -> 'a tactic
+
+(** [tclTIME s t] displays time for each atomic call to t, using s as an
+ identifying annotation if present *)
+val tclTIME : string option -> 'a tactic -> 'a tactic
+
+(** {7 Unsafe primitives} *)
+
+(** The primitives in the [Unsafe] module should be avoided as much as
+ possible, since they can make the proof state inconsistent. They are
+ nevertheless helpful, in particular when interfacing the pretyping and
+ the proof engine. *)
+module Unsafe : sig
+
+ (** [tclEVARS sigma] replaces the current [evar_map] by [sigma]. If
+ [sigma] has new unresolved [evar]-s they will not appear as
+ goal. If goals have been solved in [sigma] they will still
+ appear as unsolved goals. *)
+ val tclEVARS : Evd.evar_map -> unit tactic
+
+ (** Like {!tclEVARS} but also checks whether goals have been solved. *)
+ val tclEVARSADVANCE : Evd.evar_map -> unit tactic
+
+ (** [tclNEWGOALS gls] adds the goals [gls] to the ones currently
+ being proved, appending them to the list of focused goals. If a
+ goal is already solved, it is not added. *)
+ val tclNEWGOALS : Goal.goal list -> unit tactic
+
+ (** [tclSETGOALS gls] sets goals [gls] as the goals being under focus. If a
+ goal is already solved, it is not set. *)
+ val tclSETGOALS : Goal.goal list -> unit tactic
+
+ (** [tclGETGOALS] returns the list of goals under focus. *)
+ val tclGETGOALS : Goal.goal list tactic
+
+ (** Sets the evar universe context. *)
+ val tclEVARUNIVCONTEXT : Evd.evar_universe_context -> unit tactic
+
+ (** Clears the future goals store in the proof view. *)
+ val reset_future_goals : proofview -> proofview
+
+ (** Give an evar the status of a goal (changes its source location
+ and makes it unresolvable for type classes. *)
+ val mark_as_goal : proofview -> Evar.t -> proofview
+end
+
+(** {7 Notations} *)
-(* 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 *)
+
+ (** {!tclBIND} *)
+ val (>>=) : 'a tactic -> ('a -> 'b tactic) -> 'b tactic
+ (** {!tclTHEN} *)
val (<*>) : unit tactic -> 'a tactic -> 'a tactic
- (* tclOR *)
+ (** {!tclOR}: [t1+t2] = [tclOR t1 (fun _ -> t2)]. *)
val (<+>) : 'a tactic -> 'a tactic -> 'a tactic
+
end
+
+(** {6 Goal-dependent tactics} *)
+
+module Goal : sig
+
+ (** The type of goals. The parameter type is a phantom argument indicating
+ whether the data contained in the goal has been normalized w.r.t. the
+ current sigma. If it is the case, it is flagged [ `NF ]. You may still
+ access the un-normalized data using {!assume} if you known you do not rely
+ on the assumption of being normalized, at your own risk. *)
+ type 'a t
+
+ (** Assume that you do not need the goal to be normalized. *)
+ val assume : 'a t -> [ `NF ] t
+
+ (** Normalises the argument goal. *)
+ val normalize : 'a t -> [ `NF ] t tactic
+
+ (** [concl], [hyps], [env] and [sigma] given a goal [gl] return
+ respectively the conclusion of [gl], the hypotheses of [gl], the
+ environment of [gl] (i.e. the global environment and the
+ hypotheses) and the current evar map. *)
+ val concl : [ `NF ] t -> Term.constr
+ val hyps : [ `NF ] t -> Context.named_context
+ val env : 'a t -> Environ.env
+ val sigma : 'a t -> Evd.evar_map
+ val extra : 'a t -> Evd.Store.t
+
+ (** Returns the goal's conclusion even if the goal is not
+ normalised. *)
+ val raw_concl : 'a t -> Term.constr
+
+ (** [nf_enter t] applies the goal-dependent tactic [t] in each goal
+ independently, in the manner of {!tclINDEPENDENT} except that
+ the current goal is also given as an argument to [t]. The goal
+ is normalised with respect to evars. *)
+ val nf_enter : ([ `NF ] t -> unit tactic) -> unit tactic
+
+ (** Like {!nf_enter}, but does not normalize the goal beforehand. *)
+ val enter : ([ `LZ ] t -> unit tactic) -> unit tactic
+
+ (** Recover the list of current goals under focus, without evar-normalization *)
+ val goals : [ `LZ ] t tactic list tactic
+
+ (** Compatibility: avoid if possible *)
+ val goal : [ `NF ] t -> Evar.t
+
+end
+
+
+(** {6 The refine tactic} *)
+
+module Refine : sig
+
+ (** Printer used to print the constr which refine refines. *)
+ val pr_constr :
+ (Environ.env -> Evd.evar_map -> Term.constr -> Pp.std_ppcmds) Hook.t
+
+ (** {7 Refinement primitives} *)
+
+ val refine : ?unsafe:bool -> (Evd.evar_map -> Evd.evar_map * Constr.t) -> unit tactic
+ (** In [refine ?unsafe t], [t] is a term with holes under some
+ [evar_map] context. The term [t] is used as a partial solution
+ for the current goal (refine is a goal-dependent tactic), the
+ new holes created by [t] become the new subgoals. Exception
+ raised during the interpretation of [t] are caught and result in
+ tactic failures. If [unsafe] is [true] (default) [t] is
+ type-checked beforehand. *)
+
+ (** {7 Helper functions} *)
+
+ val with_type : Environ.env -> Evd.evar_map ->
+ Term.constr -> Term.types -> Evd.evar_map * Term.constr
+ (** [with_type env sigma c t] ensures that [c] is of type [t]
+ inserting a coercion if needed. *)
+
+ val refine_casted : ?unsafe:bool -> (Evd.evar_map -> Evd.evar_map*Constr.t) -> unit tactic
+ (** Like {!refine} except the refined term is coerced to the conclusion of the
+ current goal. *)
+
+end
+
+
+(** {6 Trace} *)
+
+module Trace : sig
+
+ (** [record_info_trace t] behaves like [t] except the [info] trace
+ is stored. *)
+ val record_info_trace : 'a tactic -> 'a tactic
+
+ val log : Proofview_monad.lazy_msg -> unit tactic
+ val name_tactic : Proofview_monad.lazy_msg -> 'a tactic -> 'a tactic
+
+ val pr_info : ?lvl:int -> Proofview_monad.Info.tree -> Pp.std_ppcmds
+
+end
+
+
+(** {6 Non-logical state} *)
+
+(** The [NonLogical] module allows the execution of effects (including
+ I/O) in tactics (non-logical side-effects are not discarded at
+ failures). *)
+module NonLogical : module type of Logic_monad.NonLogical
+
+(** [tclLIFT c] is a tactic which behaves exactly as [c]. *)
+val tclLIFT : 'a NonLogical.t -> 'a tactic
+
+
+(**/**)
+
(*** Compatibility layer with <= 8.2 tactics ***)
module V82 : sig
- type tac = Goal.goal Evd.sigma -> Goal.goal list Evd.sigma
+ type tac = Evar.t Evd.sigma -> Evar.t list Evd.sigma
val tactic : tac -> unit tactic
+ (* normalises the evars in the goals, and stores the result in
+ solution. *)
+ val nf_evar_goals : unit tactic
+
val has_unresolved_evar : proofview -> bool
(* Main function in the implementation of Grab Existential Variables.
@@ -220,17 +548,29 @@ module V82 : sig
(* Returns the open goals of the proofview together with the evar_map to
interprete them. *)
- val goals : proofview -> Goal.goal list Evd.sigma
+ val goals : proofview -> Evar.t list Evd.sigma
- val top_goals : proofview -> Goal.goal list Evd.sigma
+ val top_goals : entry -> proofview -> Evar.t list Evd.sigma
(* returns the existential variable used to start the proof *)
- val top_evars : proofview -> Evd.evar list
+ val top_evars : entry -> Evd.evar list
(* Implements the Existential command *)
- val instantiate_evar : int -> Topconstr.constr_expr -> proofview -> proofview
+ val instantiate_evar : int -> Constrexpr.constr_expr -> proofview -> proofview
+
+ (* Caution: this function loses quite a bit of information. It
+ should be avoided as much as possible. It should work as
+ expected for a tactic obtained from {!V82.tactic} though. *)
+ val of_tactic : 'a tactic -> tac
+
+ (* marks as unsafe if the argument is [false] *)
+ val put_status : bool -> unit tactic
+
+ (* exception for which it is deemed to be safe to transmute into
+ tactic failure. *)
+ val catchable_exception : exn -> bool
- (* 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
+ (* transforms every Ocaml (catchable) exception into a failure in
+ the monad. *)
+ val wrap_exceptions : (unit -> 'a tactic) -> 'a tactic
end
diff --git a/proofs/proofview_monad.ml b/proofs/proofview_monad.ml
new file mode 100644
index 00000000..6e68cd2e
--- /dev/null
+++ b/proofs/proofview_monad.ml
@@ -0,0 +1,270 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** This file defines the datatypes used as internal states by the
+ tactic monad, and specialises the [Logic_monad] to these type. *)
+
+(** {6 Trees/forest for traces} *)
+
+module Trace = struct
+
+ (** The intent is that an ['a forest] is a list of messages of type
+ ['a]. But messages can stand for a list of more precise
+ messages, hence the structure is organised as a tree. *)
+ type 'a forest = 'a tree list
+ and 'a tree = Seq of 'a * 'a forest
+
+ (** To build a trace incrementally, we use an intermediary data
+ structure on which we can define an S-expression like language
+ (like a simplified xml except the closing tags do not carry a
+ name). Note that nodes are built from right to left in ['a
+ incr], the result is mirrored when returning so that in the
+ exposed interface, the forest is read from left to right.
+
+ Concretely, we want to add a new tree to a forest: and we are
+ building it by adding new trees to the left of its left-most
+ subtrees which is built the same way. *)
+ type 'a incr = { head:'a forest ; opened: 'a tree list }
+
+ (** S-expression like language as ['a incr] transformers. It is the
+ responsibility of the library builder not to use [close] when no
+ tag is open. *)
+ let empty_incr = { head=[] ; opened=[] }
+ let opn a { head ; opened } = { head ; opened = Seq(a,[])::opened }
+ let close { head ; opened } =
+ match opened with
+ | [a] -> { head = a::head ; opened=[] }
+ | a::Seq(b,f)::opened -> { head ; opened=Seq(b,a::f)::opened }
+ | [] -> assert false
+ let leaf a s = close (opn a s)
+
+ (** Returning a forest. It is the responsibility of the library
+ builder to close all the tags. *)
+ (* spiwack: I may want to close the tags instead, to deal with
+ interruptions. *)
+ let rec mirror f = List.rev_map mirror_tree f
+ and mirror_tree (Seq(a,f)) = Seq(a,mirror f)
+
+ let to_tree = function
+ | { head ; opened=[] } -> mirror head
+ | { head ; opened=_::_} -> assert false
+
+end
+
+
+
+(** {6 State types} *)
+
+(** We typically label nodes of [Trace.tree] with messages to
+ print. But we don't want to compute the result. *)
+type lazy_msg = unit -> Pp.std_ppcmds
+let pr_lazy_msg msg = msg ()
+
+(** Info trace. *)
+module Info = struct
+
+ (** The type of the tags for [info]. *)
+ type tag =
+ | Msg of lazy_msg (** A simple message *)
+ | Tactic of lazy_msg (** A tactic call *)
+ | Dispatch (** A call to [tclDISPATCH]/[tclEXTEND] *)
+ | DBranch (** A special marker to delimit individual branch of a dispatch. *)
+
+ type state = tag Trace.incr
+ type tree = tag Trace.forest
+
+
+
+ let pr_in_comments m = Pp.(str"(* "++pr_lazy_msg m++str" *)")
+
+ let unbranch = function
+ | Trace.Seq (DBranch,brs) -> brs
+ | _ -> assert false
+
+
+ let is_empty_branch = let open Trace in function
+ | Seq(DBranch,[]) -> true
+ | _ -> false
+
+ (** Dispatch with empty branches are (supposed to be) equivalent to
+ [idtac] which need not appear, so they are removed from the
+ trace. *)
+ let dispatch brs =
+ let open Trace in
+ if CList.for_all is_empty_branch brs then None
+ else Some (Seq(Dispatch,brs))
+
+ let constr = let open Trace in function
+ | Dispatch -> dispatch
+ | t -> fun br -> Some (Seq(t,br))
+
+ let rec compress_tree = let open Trace in function
+ | Seq(t,f) -> constr t (compress f)
+ and compress f =
+ CList.map_filter compress_tree f
+
+ let rec is_empty = let open Trace in function
+ | Seq(Dispatch,brs) -> List.for_all is_empty brs
+ | Seq(DBranch,br) -> List.for_all is_empty br
+ | _ -> false
+
+ (** [with_sep] is [true] when [Tactic m] must be printed with a
+ trailing semi-colon. *)
+ let rec pr_tree with_sep = let open Trace in function
+ | Seq (Msg m,[]) -> pr_in_comments m
+ | Seq (Tactic m,_) ->
+ let tail = if with_sep then Pp.str";" else Pp.mt () in
+ Pp.(pr_lazy_msg m ++ tail)
+ | Seq (Dispatch,brs) ->
+ let tail = if with_sep then Pp.str";" else Pp.mt () in
+ Pp.(pr_dispatch brs++tail)
+ | Seq (Msg _,_::_) | Seq (DBranch,_) -> assert false
+ and pr_dispatch brs =
+ let open Pp in
+ let brs = List.map unbranch brs in
+ match brs with
+ | [br] -> pr_forest br
+ | _ ->
+ let sep () = spc()++str"|"++spc() in
+ let branches = prlist_with_sep sep pr_forest brs in
+ str"[>"++spc()++branches++spc()++str"]"
+ and pr_forest = function
+ | [] -> Pp.mt ()
+ | [tr] -> pr_tree false tr
+ | tr::l -> Pp.(pr_tree true tr ++ pr_forest l)
+
+ let print f =
+ pr_forest (compress f)
+
+ let rec collapse_tree n t =
+ let open Trace in
+ match n , t with
+ | 0 , t -> [t]
+ | _ , (Seq(Tactic _,[]) as t) -> [t]
+ | n , Seq(Tactic _,f) -> collapse (pred n) f
+ | n , Seq(Dispatch,brs) -> [Seq(Dispatch, (collapse n brs))]
+ | n , Seq(DBranch,br) -> [Seq(DBranch, (collapse n br))]
+ | _ , (Seq(Msg _,_) as t) -> [t]
+ and collapse n f =
+ CList.map_append (collapse_tree n) f
+end
+
+
+(** Type of proof views: current [evar_map] together with the list of
+ focused goals. *)
+type proofview = { solution : Evd.evar_map; comb : Goal.goal list }
+
+
+(** {6 Instantiation of the logic monad} *)
+
+(** Parameters of the logic monads *)
+module P = struct
+
+ type s = proofview * Environ.env
+
+ (** Recording info trace (true) or not. *)
+ type e = bool
+
+ (** Status (safe/unsafe) * shelved goals * given up *)
+ type w = bool * Evar.t list * Evar.t list
+
+ let wunit = true , [] , []
+ let wprod (b1,s1,g1) (b2,s2,g2) = b1 && b2 , s1@s2 , g1@g2
+
+ type u = Info.state
+
+ let uunit = Trace.empty_incr
+
+end
+
+module Logical = Logic_monad.Logical(P)
+
+
+(** {6 Lenses to access to components of the states} *)
+
+module type State = sig
+ type t
+ val get : t Logical.t
+ val set : t -> unit Logical.t
+ val modify : (t->t) -> unit Logical.t
+end
+
+module type Writer = sig
+ type t
+ val put : t -> unit Logical.t
+end
+
+module Pv : State with type t := proofview = struct
+ let get = Logical.(map fst get)
+ let set p = Logical.modify (fun (_,e) -> (p,e))
+ let modify f= Logical.modify (fun (p,e) -> (f p,e))
+end
+
+module Solution : State with type t := Evd.evar_map = struct
+ let get = Logical.map (fun {solution} -> solution) Pv.get
+ let set s = Pv.modify (fun pv -> { pv with solution = s })
+ let modify f = Pv.modify (fun pv -> { pv with solution = f pv.solution })
+end
+
+module Comb : State with type t = Evar.t list = struct
+ (* spiwack: I don't know why I cannot substitute ([:=]) [t] with a type expression. *)
+ type t = Evar.t list
+ let get = Logical.map (fun {comb} -> comb) Pv.get
+ let set c = Pv.modify (fun pv -> { pv with comb = c })
+ let modify f = Pv.modify (fun pv -> { pv with comb = f pv.comb })
+end
+
+module Env : State with type t := Environ.env = struct
+ let get = Logical.(map snd get)
+ let set e = Logical.modify (fun (p,_) -> (p,e))
+ let modify f = Logical.modify (fun (p,e) -> (p,f e))
+end
+
+module Status : Writer with type t := bool = struct
+ let put s = Logical.put (s,[],[])
+end
+
+module Shelf : Writer with type t = Evar.t list = struct
+ (* spiwack: I don't know why I cannot substitute ([:=]) [t] with a type expression. *)
+ type t = Evar.t list
+ let put sh = Logical.put (true,sh,[])
+end
+
+module Giveup : Writer with type t = Evar.t list = struct
+ (* spiwack: I don't know why I cannot substitute ([:=]) [t] with a type expression. *)
+ type t = Evar.t list
+ let put gs = Logical.put (true,[],gs)
+end
+
+(** Lens and utilies pertaining to the info trace *)
+module InfoL = struct
+ let recording = Logical.current
+ let if_recording t =
+ let open Logical in
+ recording >>= fun r ->
+ if r then t else return ()
+
+ let record_trace t = Logical.local true t
+
+ let raw_update = Logical.update
+ let update f = if_recording (raw_update f)
+ let opn a = update (Trace.opn a)
+ let close = update Trace.close
+ let leaf a = update (Trace.leaf a)
+
+ let tag a t =
+ let open Logical in
+ recording >>= fun r ->
+ if r then begin
+ raw_update (Trace.opn a) >>
+ t >>= fun a ->
+ raw_update Trace.close >>
+ return a
+ end else
+ t
+end
diff --git a/proofs/proofview_monad.mli b/proofs/proofview_monad.mli
new file mode 100644
index 00000000..d2a2e55f
--- /dev/null
+++ b/proofs/proofview_monad.mli
@@ -0,0 +1,144 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** This file defines the datatypes used as internal states by the
+ tactic monad, and specialises the [Logic_monad] to these type. *)
+
+(** {6 Traces} *)
+
+module Trace : sig
+
+ (** The intent is that an ['a forest] is a list of messages of type
+ ['a]. But messages can stand for a list of more precise
+ messages, hence the structure is organised as a tree. *)
+ type 'a forest = 'a tree list
+ and 'a tree = Seq of 'a * 'a forest
+
+ (** To build a trace incrementally, we use an intermediary data
+ structure on which we can define an S-expression like language
+ (like a simplified xml except the closing tags do not carry a
+ name). *)
+ type 'a incr
+ val to_tree : 'a incr -> 'a forest
+
+ (** [open a] opens a tag with name [a]. *)
+ val opn : 'a -> 'a incr -> 'a incr
+
+ (** [close] closes the last open tag. It is the responsibility of
+ the user to close all the tags. *)
+ val close : 'a incr -> 'a incr
+
+ (** [leaf] creates an empty tag with name [a]. *)
+ val leaf : 'a -> 'a incr -> 'a incr
+
+end
+
+(** {6 State types} *)
+
+(** We typically label nodes of [Trace.tree] with messages to
+ print. But we don't want to compute the result. *)
+type lazy_msg = unit -> Pp.std_ppcmds
+
+(** Info trace. *)
+module Info : sig
+
+ (** The type of the tags for [info]. *)
+ type tag =
+ | Msg of lazy_msg (** A simple message *)
+ | Tactic of lazy_msg (** A tactic call *)
+ | Dispatch (** A call to [tclDISPATCH]/[tclEXTEND] *)
+ | DBranch (** A special marker to delimit individual branch of a dispatch. *)
+
+ type state = tag Trace.incr
+ type tree = tag Trace.forest
+
+ val print : tree -> Pp.std_ppcmds
+
+ (** [collapse n t] flattens the first [n] levels of [Tactic] in an
+ info trace, effectively forgetting about the [n] top level of
+ names (if there are fewer, the last name is kept). *)
+ val collapse : int -> tree -> tree
+
+end
+
+(** Type of proof views: current [evar_map] together with the list of
+ focused goals. *)
+type proofview = { solution : Evd.evar_map; comb : Goal.goal list }
+
+(** {6 Instantiation of the logic monad} *)
+
+module P : sig
+ type s = proofview * Environ.env
+
+ (** Status (safe/unsafe) * shelved goals * given up *)
+ type w = bool * Evar.t list * Evar.t list
+
+ val wunit : w
+ val wprod : w -> w -> w
+
+ (** Recording info trace (true) or not. *)
+ type e = bool
+
+ type u = Info.state
+
+ val uunit : u
+end
+
+module Logical : module type of Logic_monad.Logical(P)
+
+
+(** {6 Lenses to access to components of the states} *)
+
+module type State = sig
+ type t
+ val get : t Logical.t
+ val set : t -> unit Logical.t
+ val modify : (t->t) -> unit Logical.t
+end
+
+module type Writer = sig
+ type t
+ val put : t -> unit Logical.t
+end
+
+(** Lens to the [proofview]. *)
+module Pv : State with type t := proofview
+
+(** Lens to the [evar_map] of the proofview. *)
+module Solution : State with type t := Evd.evar_map
+
+(** Lens to the list of focused goals. *)
+module Comb : State with type t = Evar.t list
+
+(** Lens to the global environment. *)
+module Env : State with type t := Environ.env
+
+(** Lens to the tactic status ([true] if safe, [false] if unsafe) *)
+module Status : Writer with type t := bool
+
+(** Lens to the list of goals which have been shelved during the
+ execution of the tactic. *)
+module Shelf : Writer with type t = Evar.t list
+
+(** Lens to the list of goals which were given up during the execution
+ of the tactic. *)
+module Giveup : Writer with type t = Evar.t list
+
+(** Lens and utilies pertaining to the info trace *)
+module InfoL : sig
+ (** [record_trace t] behaves like [t] and compute its [info] trace. *)
+ val record_trace : 'a Logical.t -> 'a Logical.t
+
+ val update : (Info.state -> Info.state) -> unit Logical.t
+ val opn : Info.tag -> unit Logical.t
+ val close : unit Logical.t
+ val leaf : Info.tag -> unit Logical.t
+
+ (** [tag a t] opens tag [a] runs [t] then closes the tag. *)
+ val tag : Info.tag -> 'a Logical.t -> 'a Logical.t
+end
diff --git a/proofs/redexpr.ml b/proofs/redexpr.ml
index 3c2431a1..18588867 100644
--- a/proofs/redexpr.ml
+++ b/proofs/redexpr.ml
@@ -1,25 +1,26 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Pp
+open Errors
open Util
open Names
open Term
open Declarations
-open Libnames
-open Glob_term
+open Globnames
+open Genredexpr
open Pattern
open Reductionops
open Tacred
open Closure
open RedFlags
open Libobject
-open Summary
+open Misctypes
(* call by value normalisation function using the virtual machine *)
let cbv_vm env sigma c =
@@ -28,13 +29,35 @@ let cbv_vm env sigma c =
error "vm_compute does not support existential variables.";
Vnorm.cbv_vm env c ctyp
+let cbv_native env sigma c =
+ let ctyp = Retyping.get_type_of env sigma c in
+ let evars = Nativenorm.evars_of_evar_map sigma in
+ Nativenorm.native_norm env evars c ctyp
+
+let whd_cbn flags env sigma t =
+ let (state,_) =
+ (whd_state_gen true flags env sigma (t,Reductionops.Stack.empty))
+ in Reductionops.Stack.zip ~refold:true state
+
+let strong_cbn flags =
+ strong (whd_cbn flags)
+
+let simplIsCbn = ref (false)
+let _ = Goptions.declare_bool_option {
+ Goptions.optsync = true; Goptions.optdepr = false;
+ Goptions.optname =
+ "Plug the simpl tactic to the new cbn mechanism";
+ Goptions.optkey = ["SimplIsCbn"];
+ Goptions.optread = (fun () -> !simplIsCbn);
+ Goptions.optwrite = (fun a -> simplIsCbn:=a);
+}
let set_strategy_one ref l =
let k =
match ref with
| EvalConstRef sp -> ConstKey sp
| EvalVarRef id -> VarKey id in
- Conv_oracle.set_strategy k l;
+ Global.set_strategy k l;
match k,l with
ConstKey sp, Conv_oracle.Opaque ->
Csymtable.set_opaque_const sp
@@ -44,7 +67,7 @@ let set_strategy_one ref l =
| OpaqueDef _ ->
errorlabstrm "set_transparent_const"
(str "Cannot make" ++ spc () ++
- Nametab.pr_global_env Idset.empty (ConstRef sp) ++
+ Nametab.pr_global_env Id.Set.empty (ConstRef sp) ++
spc () ++ str "transparent because it was declared opaque.");
| _ -> Csymtable.set_transparent_const sp)
| _ -> ()
@@ -56,9 +79,9 @@ let cache_strategy (_,str) =
let subst_strategy (subs,(local,obj)) =
local,
- list_smartmap
+ List.smartmap
(fun (k,ql as entry) ->
- let ql' = list_smartmap (Mod_subst.subst_evaluable_reference subs) ql in
+ let ql' = List.smartmap (Mod_subst.subst_evaluable_reference subs) ql in
if ql==ql' then entry else (k,ql'))
obj
@@ -71,8 +94,8 @@ let map_strategy f l =
match f q with
Some q' -> q' :: ql
| None -> ql) ql [] in
- if ql'=[] then str else (lev,ql')::str) l [] in
- if l'=[] then None else Some (false,l')
+ if List.is_empty ql' then str else (lev,ql')::str) l [] in
+ if List.is_empty l' then None else Some (false,l')
let classify_strategy (local,_ as obj) =
if local then Dispose else Substitute obj
@@ -103,12 +126,6 @@ let inStrategy : strategy_obj -> obj =
let set_strategy local str =
Lib.add_anonymous_leaf (inStrategy (local,str))
-let _ = declare_summary "Transparent constants and variables"
- { freeze_function = Conv_oracle.freeze;
- unfreeze_function = Conv_oracle.unfreeze;
- init_function = Conv_oracle.init }
-
-
(* Generic reduction: reduction functions used in reduction tactics *)
type red_expr =
@@ -118,7 +135,7 @@ let make_flag_constant = function
| EvalVarRef id -> fVAR id
| EvalConstRef sp -> fCONST sp
-let make_flag f =
+let make_flag env f =
let red = no_red in
let red = if f.rBeta then red_add red fBETA else red in
let red = if f.rIota then red_add red fIOTA else red in
@@ -126,7 +143,8 @@ let make_flag f =
let red =
if f.rDelta then (* All but rConst *)
let red = red_add red fDELTA in
- let red = red_add_transparent red (Conv_oracle.get_transp_state()) in
+ let red = red_add_transparent red
+ (Conv_oracle.get_transp_state (Environ.oracle env)) in
List.fold_right
(fun v red -> red_sub red (make_flag_constant v))
f.rConst red
@@ -141,81 +159,90 @@ let is_reference = function PRef _ | PVar _ -> true | _ -> false
(* table of custom reductino fonctions, not synchronized,
filled via ML calls to [declare_reduction] *)
-let reduction_tab = ref Stringmap.empty
+let reduction_tab = ref String.Map.empty
(* table of custom reduction expressions, synchronized,
filled by command Declare Reduction *)
-let red_expr_tab = ref Stringmap.empty
+let red_expr_tab = Summary.ref String.Map.empty ~name:"Declare Reduction"
let declare_reduction s f =
- if Stringmap.mem s !reduction_tab || Stringmap.mem s !red_expr_tab
+ if String.Map.mem s !reduction_tab || String.Map.mem s !red_expr_tab
then error ("There is already a reduction expression of name "^s)
- else reduction_tab := Stringmap.add s f !reduction_tab
+ else reduction_tab := String.Map.add s f !reduction_tab
let check_custom = function
| ExtraRedExpr s ->
- if not (Stringmap.mem s !reduction_tab || Stringmap.mem s !red_expr_tab)
+ if not (String.Map.mem s !reduction_tab || String.Map.mem s !red_expr_tab)
then error ("Reference to undefined reduction expression "^s)
|_ -> ()
let decl_red_expr s e =
- if Stringmap.mem s !reduction_tab || Stringmap.mem s !red_expr_tab
+ if String.Map.mem s !reduction_tab || String.Map.mem s !red_expr_tab
then error ("There is already a reduction expression of name "^s)
else begin
check_custom e;
- red_expr_tab := Stringmap.add s e !red_expr_tab
+ red_expr_tab := String.Map.add s e !red_expr_tab
end
let out_arg = function
- | ArgVar _ -> anomaly "Unevaluated or_var variable"
+ | ArgVar _ -> anomaly (Pp.str "Unevaluated or_var variable")
| ArgArg x -> x
-let out_with_occurrences ((b,l),c) =
- ((b,List.map out_arg l), c)
+let out_with_occurrences (occs,c) =
+ (Locusops.occurrences_map (List.map out_arg) occs, c)
+
+let e_red f env evm c = evm, f env evm c
+
+let head_style = false (* Turn to true to have a semantics where simpl
+ only reduce at the head when an evaluable reference is given, e.g.
+ 2+n would just reduce to S(1+n) instead of S(S(n)) *)
+
+let contextualize f g = function
+ | Some (occs,c) ->
+ let l = Locusops.occurrences_map (List.map out_arg) occs in
+ let b,c,h = match c with
+ | Inl r -> true,PRef (global_of_evaluable_reference r),f
+ | Inr c -> false,c,f in
+ e_red (contextually b (l,c) (fun _ -> h))
+ | None -> e_red g
-let rec reduction_of_red_expr = function
+let reduction_of_red_expr env =
+ let make_flag = make_flag env in
+ let rec reduction_of_red_expr = function
| Red internal ->
- if internal then (try_red_product,DEFAULTcast)
- else (red_product,DEFAULTcast)
- | Hnf -> (hnf_constr,DEFAULTcast)
- | Simpl (Some (_,c as lp)) ->
- (contextually (is_reference c) (out_with_occurrences lp)
- (fun _ -> simpl),DEFAULTcast)
- | Simpl None -> (simpl,DEFAULTcast)
- | Cbv f -> (cbv_norm_flags (make_flag f),DEFAULTcast)
- | Lazy f -> (clos_norm_flags (make_flag f),DEFAULTcast)
- | Unfold ubinds -> (unfoldn (List.map out_with_occurrences ubinds),DEFAULTcast)
- | Fold cl -> (fold_commands cl,DEFAULTcast)
+ if internal then (e_red try_red_product,DEFAULTcast)
+ else (e_red red_product,DEFAULTcast)
+ | Hnf -> (e_red hnf_constr,DEFAULTcast)
+ | Simpl (f,o) ->
+ let whd_am = if !simplIsCbn then whd_cbn (make_flag f) else whd_simpl in
+ let am = if !simplIsCbn then strong_cbn (make_flag f) else simpl in
+ let () =
+ if not (!simplIsCbn || List.is_empty f.rConst) then
+ Pp.msg_warning (Pp.strbrk "The legacy simpl does not deal with delta flags.") in
+ (contextualize (if head_style then whd_am else am) am o,DEFAULTcast)
+ | Cbv f -> (e_red (cbv_norm_flags (make_flag f)),DEFAULTcast)
+ | Cbn f ->
+ (e_red (strong_cbn (make_flag f)), DEFAULTcast)
+ | Lazy f -> (e_red (clos_norm_flags (make_flag f)),DEFAULTcast)
+ | Unfold ubinds -> (e_red (unfoldn (List.map out_with_occurrences ubinds)),DEFAULTcast)
+ | Fold cl -> (e_red (fold_commands cl),DEFAULTcast)
| Pattern lp -> (pattern_occs (List.map out_with_occurrences lp),DEFAULTcast)
| ExtraRedExpr s ->
- (try (Stringmap.find s !reduction_tab,DEFAULTcast)
+ (try (e_red (String.Map.find s !reduction_tab),DEFAULTcast)
with Not_found ->
- (try reduction_of_red_expr (Stringmap.find s !red_expr_tab)
+ (try reduction_of_red_expr (String.Map.find s !red_expr_tab)
with Not_found ->
error("unknown user-defined reduction \""^s^"\"")))
- | CbvVm -> (cbv_vm ,VMcast)
+ | CbvVm o -> (contextualize cbv_vm cbv_vm o, VMcast)
+ | CbvNative o -> (contextualize cbv_native cbv_native o, VMcast)
+ in
+ reduction_of_red_expr
-
-let subst_flags subs flags =
- { flags with rConst = List.map subs flags.rConst }
-
-let subst_occs subs (occ,e) = (occ,subs e)
-
-let subst_gen_red_expr subs_a subs_b subs_c = function
- | Fold l -> Fold (List.map subs_a l)
- | Pattern occs_l -> Pattern (List.map (subst_occs subs_a) occs_l)
- | Simpl occs_o -> Simpl (Option.map (subst_occs subs_c) occs_o)
- | Unfold occs_l -> Unfold (List.map (subst_occs subs_b) occs_l)
- | Cbv flags -> Cbv (subst_flags subs_b flags)
- | Lazy flags -> Lazy (subst_flags subs_b flags)
- | e -> e
-
-let subst_red_expr subs e =
- subst_gen_red_expr
+let subst_red_expr subs =
+ Miscops.map_red_expr_gen
(Mod_subst.subst_mps subs)
(Mod_subst.subst_evaluable_reference subs)
- (Pattern.subst_pattern subs)
- e
+ (Patternops.subst_pattern subs)
let inReduction : bool * string * red_expr -> obj =
declare_object
@@ -229,8 +256,3 @@ let inReduction : bool * string * red_expr -> obj =
let declare_red_expr locality s expr =
Lib.add_anonymous_leaf (inReduction (locality,s,expr))
-
-let _ = declare_summary "Declare Reduction"
- { freeze_function = (fun () -> !red_expr_tab);
- unfreeze_function = ((:=) red_expr_tab);
- init_function = (fun () -> red_expr_tab := Stringmap.empty) }
diff --git a/proofs/redexpr.mli b/proofs/redexpr.mli
index 285f9cfd..b32cedf8 100644
--- a/proofs/redexpr.mli
+++ b/proofs/redexpr.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,18 +8,18 @@
open Names
open Term
-open Closure
open Pattern
-open Glob_term
+open Genredexpr
open Reductionops
-open Termops
+open Locus
type red_expr =
(constr, evaluable_global_reference, constr_pattern) red_expr_gen
-
+
val out_with_occurrences : 'a with_occurrences -> occurrences * 'a
-val reduction_of_red_expr : red_expr -> reduction_function * cast_kind
+val reduction_of_red_expr :
+ Environ.env -> red_expr -> e_reduction_function * cast_kind
(** [true] if we should use the vm to verify the reduction *)
diff --git a/proofs/refiner.ml b/proofs/refiner.ml
index ca82e882..974fa212 100644
--- a/proofs/refiner.ml
+++ b/proofs/refiner.ml
@@ -1,22 +1,16 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Compat
open Pp
+open Errors
open Util
-open Term
-open Termops
-open Sign
open Evd
-open Sign
open Environ
-open Reductionops
-open Type_errors
open Proof_type
open Logic
@@ -29,66 +23,41 @@ let project x = x.sigma
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))
-let abstract_operation syntax semantics =
- semantics
+let refiner pr goal_sigma =
+ let (sgl,sigma') = prim_refiner pr goal_sigma.sigma goal_sigma.it in
+ { it = sgl; sigma = sigma'; }
-let abstract_tactic_expr ?(dflt=false) te tacfun gls =
- abstract_operation (Tactic(te,dflt)) tacfun gls
-
-let abstract_tactic ?(dflt=false) te =
- !abstract_tactic_box := Some te;
- abstract_tactic_expr ~dflt (Tacexpr.TacAtom (dummy_loc,te))
-
-let abstract_extended_tactic ?(dflt=false) s args =
- abstract_tactic ~dflt (Tacexpr.TacExtend (dummy_loc, s, args))
-
-let refiner = function
- | 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'})
-
-
- | Nested (_,_) | Decl_proof _ ->
- failwith "Refiner: should not occur"
-
- (* Daimon is a canonical unfinished proof *)
-
- | Daimon ->
- fun gls ->
- {it=[];sigma=gls.sigma}
-
-
-let norm_evar_tac gl = refiner (Prim Change_evars) gl
+(* Profiling refiner *)
+let refiner =
+ if Flags.profile then
+ let refiner_key = Profile.declare_profile "refiner" in
+ Profile.profile2 refiner_key refiner
+ else refiner
(*********************)
(* Tacticals *)
(*********************)
-let unpackage glsig = (ref (glsig.sigma)),glsig.it
+let unpackage glsig = (ref (glsig.sigma)), glsig.it
-let repackage r v = {it=v;sigma = !r}
+let repackage r v = {it = v; sigma = !r; }
let apply_sig_tac r tac g =
- check_for_interrupt (); (* Breakpoint *)
+ Control.check_for_interrupt (); (* Breakpoint *)
let glsigma = tac (repackage r g) in
r := glsigma.sigma;
glsigma.it
(* [goal_goal_list : goal sigma -> goal list sigma] *)
-let goal_goal_list gls = {it=[gls.it];sigma=gls.sigma}
-
-(* forces propagation of evar constraints *)
-let tclNORMEVAR = norm_evar_tac
+let goal_goal_list gls = {it=[gls.it]; sigma=gls.sigma; }
(* identity tactic without any message *)
let tclIDTAC gls = goal_goal_list gls
(* the message printing identity tactic *)
let tclIDTAC_MESSAGE s gls =
- msg (hov 0 s); tclIDTAC gls
+ Pp.msg_info (hov 0 s); pp_flush (); tclIDTAC gls
(* General failure tactic *)
let tclFAIL_s s gls = errorlabstrm "Refiner.tclFAIL_s" (str s)
@@ -102,8 +71,8 @@ let tclFAIL lvl s g = raise (FailError (lvl,lazy s))
let tclFAIL_lazy lvl s g = raise (FailError (lvl,s))
let start_tac gls =
- let (sigr,g) = unpackage gls in
- (sigr,[g])
+ let sigr, g = unpackage gls in
+ (sigr, [g])
let finish_tac (sigr,gl) = repackage sigr gl
@@ -115,7 +84,7 @@ let thens3parts_tac tacfi tac tacli (sigr,gs) =
let ng = List.length gs in
if ng<nf+nl then errorlabstrm "Refiner.thensn_tac" (str "Not enough subgoals.");
let gll =
- (list_map_i (fun i ->
+ (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)
@@ -123,33 +92,14 @@ let thens3parts_tac tacfi tac tacli (sigr,gs) =
(* Apply [taci.(i)] on the first n subgoals and [tac] on the others *)
let thensf_tac taci tac = thens3parts_tac taci tac [||]
-(* Apply [taci.(i)] on the last n subgoals and [tac] on the others *)
-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) =
let gll =
- list_map_i (fun i -> apply_sig_tac sigr (tac i)) 1 gs in
+ 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
-let non_existent_goal n =
- errorlabstrm ("No such goal: "^(string_of_int n))
- (str"Trying to apply a tactic to a non existent goal")
-
-(* 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 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.")
- else if k >= 1 & k <= nsg then
- thensf_tac
- (Array.init k (fun i -> if i+1 = k then tac else tclIDTAC)) tclIDTAC
- subgoals
- else non_existent_goal k
-
(* [tclTHENS3PARTS tac1 [|t1 ; ... ; tn|] tac2 [|t'1 ; ... ; t'm|] gls]
applies the tactic [tac1] to [gls] then, applies [t1], ..., [tn] to
the first [n] resulting subgoals, [t'1], ..., [t'm] to the last [m]
@@ -236,26 +186,55 @@ let tclNOTSAMEGOAL (tac : tactic) goal =
(str"Tactic generated a subgoal identical to the original goal.")
else rslt
-let catch_failerror e =
- if catchable_exception e then check_for_interrupt ()
+(* Execute tac and show the names of hypothesis create by tac in
+ the "as" format. The resulting goals are printed *after* the
+ as-expression, which forces pg to some gymnastic. TODO: Have
+ something similar (better?) in the xml protocol. *)
+let tclSHOWHYPS (tac : tactic) (goal: Goal.goal Evd.sigma)
+ :Proof_type.goal list Evd.sigma =
+ let oldhyps:Context.named_context = pf_hyps goal in
+ let rslt:Proof_type.goal list Evd.sigma = tac goal in
+ let { it = gls; sigma = sigma; } = rslt in
+ let hyps:Context.named_context list =
+ List.map (fun gl -> pf_hyps { it = gl; sigma=sigma; }) gls in
+ let newhyps =
+ List.map
+ (fun hypl -> List.subtract Context.eq_named_declaration hypl oldhyps)
+ hyps
+ in
+ let emacs_str s =
+ if !Flags.print_emacs then s else "" in
+ let s =
+ let frst = ref true in
+ List.fold_left
+ (fun acc lh -> acc ^ (if !frst then (frst:=false;"") else " | ")
+ ^ (List.fold_left
+ (fun acc (nm,_,_) -> (Names.Id.to_string nm) ^ " " ^ acc)
+ "" lh))
+ "" newhyps in
+ pp (str (emacs_str "<infoH>")
+ ++ (hov 0 (str s))
+ ++ (str (emacs_str "</infoH>")) ++ fnl());
+ rslt;;
+
+
+let catch_failerror (e, info) =
+ if catchable_exception e then Control.check_for_interrupt ()
else match e with
- | FailError (0,_) | Loc.Exc_located(_, FailError (0,_))
- | Loc.Exc_located(_, LtacLocated (_,FailError (0,_))) ->
- check_for_interrupt ()
- | FailError (lvl,s) -> raise (FailError (lvl - 1, s))
- | Loc.Exc_located(s,FailError (lvl,s')) ->
- raise (Loc.Exc_located(s,FailError (lvl - 1, s')))
- | Loc.Exc_located(s,LtacLocated (s'',FailError (lvl,s'))) ->
- raise
- (Loc.Exc_located(s,LtacLocated (s'',FailError (lvl - 1,s'))))
- | e -> raise e
+ | FailError (0,_) ->
+ Control.check_for_interrupt ()
+ | FailError (lvl,s) ->
+ iraise (FailError (lvl - 1, s), info)
+ | e -> iraise (e, info)
+ (** FIXME: do we need to add a [Errors.push] here? *)
(* ORELSE0 t1 t2 tries to apply t1 and if it fails, applies t2 *)
let tclORELSE0 t1 t2 g =
try
t1 g
with (* Breakpoint *)
- | e when Errors.noncritical e -> catch_failerror e; t2 g
+ | e when Errors.noncritical e ->
+ let e = Errors.push e in catch_failerror e; t2 g
(* ORELSE t1 t2 tries to apply t1 and if it fails or does not progress,
then applies t2 *)
@@ -267,11 +246,12 @@ let tclORELSE t1 t2 = tclORELSE0 (tclPROGRESS t1) t2
let tclORELSE_THEN t1 t2then t2else gls =
match
try Some(tclPROGRESS t1 gls)
- with e when Errors.noncritical e -> catch_failerror e; None
+ with e when Errors.noncritical e ->
+ let e = Errors.push e in catch_failerror e; None
with
| None -> t2else gls
| Some sgl ->
- let (sigr,gl) = unpackage sgl in
+ let sigr, gl = unpackage sgl in
finish_tac (then_tac t2then (sigr,gl))
(* TRY f tries to apply f, and if it fails, leave the goal unchanged *)
@@ -292,13 +272,17 @@ let ite_gen tcal tac_if continue tac_else gl=
success:=true;result in
let tac_else0 e gl=
if !success then
- raise e
+ iraise e
else
- tac_else gl in
- try
- tcal tac_if0 continue gl
- with (* Breakpoint *)
- | e when Errors.noncritical e -> catch_failerror e; tac_else0 e gl
+ try
+ tac_else gl
+ with
+ e' when Errors.noncritical e' -> iraise e in
+ try
+ tcal tac_if0 continue gl
+ with (* Breakpoint *)
+ | e when Errors.noncritical e ->
+ let e = Errors.push e in catch_failerror e; tac_else0 e gl
(* Try the first tactic and, if it succeeds, continue with
the second one, and if it fails, use the third one *)
@@ -327,32 +311,11 @@ let tclDO n t =
let rec dorec k =
if k < 0 then errorlabstrm "Refiner.tclDO"
(str"Wrong argument : Do needs a positive integer.");
- if k = 0 then tclIDTAC
- else if k = 1 then t else (tclTHEN t (dorec (k-1)))
+ if Int.equal k 0 then tclIDTAC
+ else if Int.equal k 1 then t else (tclTHEN t (dorec (k-1)))
in
dorec n
-(* Fails if a tactic hasn't finished after a certain amount of time *)
-
-exception TacTimeout
-
-let tclTIMEOUT n t g =
- let timeout_handler _ = raise TacTimeout in
- let psh = Sys.signal Sys.sigalrm (Sys.Signal_handle timeout_handler) in
- ignore (Unix.alarm n);
- let restore_timeout () =
- ignore (Unix.alarm 0);
- Sys.set_signal Sys.sigalrm psh
- in
- try
- let res = t g in
- restore_timeout ();
- res
- with
- | TacTimeout | Loc.Exc_located(_,TacTimeout) ->
- restore_timeout ();
- errorlabstrm "Refiner.tclTIMEOUT" (str"Timeout!")
- | reraise -> restore_timeout (); raise reraise
(* Beware: call by need of CAML, g is needed *)
let rec tclREPEAT t g =
@@ -362,81 +325,20 @@ let tclAT_LEAST_ONCE t = (tclTHEN t (tclREPEAT t))
(* Repeat on the first subgoal (no failure if no more subgoal) *)
let rec tclREPEAT_MAIN t g =
- (tclORELSE (tclTHEN_i t (fun i -> if i = 1 then (tclREPEAT_MAIN t) else
+ (tclORELSE (tclTHEN_i t (fun i -> if Int.equal i 1 then (tclREPEAT_MAIN t) else
tclIDTAC)) tclIDTAC) g
-(*s Tactics handling a list of goals. *)
-
-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
-
-(* first_goal : goal list sigma -> goal sigma *)
-
-let first_goal gls =
- let gl = gls.it and sig_0 = gls.sigma in
- if gl = [] then error "first_goal";
- { it = List.hd gl; sigma = sig_0 }
-
-(* goal_goal_list : goal sigma -> goal list sigma *)
-
-let goal_goal_list gls =
- let gl = gls.it and sig_0 = gls.sigma in { it = [gl]; sigma = sig_0 }
-
-(* tactic -> tactic_list : Apply a tactic to the first goal in the list *)
-
-let apply_tac_list tac glls =
- let (sigr,lg) = unpackage glls in
- match lg with
- | (g1::rest) ->
- 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 = tacl1 glls in
- let glls2 = tacl2 glls1 in
- glls2
-
-(* Transform a tactic_list into a tactic *)
-
-let tactic_list_tactic tac gls =
- let glres = tac (goal_goal_list gls) in
- glres
-
(* Change evars *)
let tclEVARS sigma gls = tclIDTAC {gls with sigma=sigma}
-(* Pretty-printers. *)
+let tclEVARUNIVCONTEXT ctx gls = tclIDTAC {gls with sigma= Evd.set_universe_context gls.sigma ctx}
-let pp_info = ref (fun _ _ _ -> assert false)
-let set_info_printer f = pp_info := f
+(* Push universe context *)
+let tclPUSHCONTEXT rigid ctx tac gl =
+ tclTHEN (tclEVARS (Evd.merge_context_set rigid (project gl) ctx)) tac gl
-(* Check that holes in arguments have been resolved *)
+let tclPUSHEVARUNIVCONTEXT ctx gl =
+ tclEVARS (Evd.merge_universe_context (project gl) ctx) gl
-let check_evars env sigma extsigma gl =
- let origsigma = gl.sigma in
- let rest =
- Evd.fold_undefined (fun evk evi acc ->
- if Evd.is_undefined extsigma evk & not (Evd.mem origsigma evk) then
- evi::acc
- else
- acc)
- sigma []
- in
- if rest <> [] then
- let evi = List.hd rest in
- let (loc,k) = evi.evar_source in
- let evi = Evarutil.nf_evar_info sigma evi in
- Pretype_errors.error_unsolvable_implicit loc env sigma evi k None
-
-let tclWITHHOLES accept_unresolved_holes tac sigma c gl =
- if sigma == project gl then tac c gl
- else
- let res = tclTHEN (tclEVARS sigma) (tac c) gl in
- if not accept_unresolved_holes then
- check_evars (pf_env gl) (res).sigma sigma gl;
- res
+let tclPUSHCONSTRAINTS cst gl =
+ tclEVARS (Evd.add_constraints (project gl) cst) gl
diff --git a/proofs/refiner.mli b/proofs/refiner.mli
index 5db43eaf..a81555ff 100644
--- a/proofs/refiner.mli
+++ b/proofs/refiner.mli
@@ -1,17 +1,14 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
-open Sign
+open Context
open Evd
open Proof_type
-open Tacexpr
-open Logic
(** The refiner (handles primitive rules and high-level tactics). *)
@@ -26,31 +23,22 @@ val repackage : evar_map ref -> 'a -> 'a sigma
val apply_sig_tac :
evar_map ref -> (goal sigma -> goal list sigma) -> goal -> goal list
-(** {6 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
-val abstract_extended_tactic :
- ?dflt:bool -> string -> typed_generic_argument list -> tactic -> tactic
-
val refiner : rule -> tactic
(** {6 Tacticals. } *)
-(** [tclNORMEVAR] forces propagation of evar constraints *)
-val tclNORMEVAR : tactic
-
(** [tclIDTAC] is the identity tactic without message printing*)
val tclIDTAC : tactic
val tclIDTAC_MESSAGE : Pp.std_ppcmds -> tactic
(** [tclEVARS sigma] changes the current evar map *)
val tclEVARS : evar_map -> tactic
+val tclEVARUNIVCONTEXT : Evd.evar_universe_context -> tactic
+
+val tclPUSHCONTEXT : Evd.rigid -> Univ.universe_context_set -> tactic -> tactic
+val tclPUSHEVARUNIVCONTEXT : Evd.evar_universe_context -> tactic
+
+val tclPUSHCONSTRAINTS : Univ.constraints -> tactic
(** [tclTHEN tac1 tac2 gls] applies the tactic [tac1] to [gls] and applies
[tac2] to every resulting subgoals *)
@@ -83,10 +71,6 @@ val tclTHENSV : tactic -> tactic array -> tactic
(** Same with a list of tactics *)
val tclTHENS : tactic -> tactic list -> tactic
-(** [tclTHENST] is renamed [tclTHENSFIRSTn]
-val tclTHENST : tactic -> tactic array -> tactic -> tactic
-*)
-
(** [tclTHENS3PARTS tac1 [|t1 ; ... ; tn|] tac2 [|t'1 ; ... ; t'm|] gls]
applies the tactic [tac1] to [gls] then, applies [t1], ..., [tn] to
the first [n] resulting subgoals, [t'1], ..., [t'm] to the last [m]
@@ -118,7 +102,7 @@ exception FailError of int * Pp.std_ppcmds Lazy.t
(** Takes an exception and either raise it at the next
level or do nothing. *)
-val catch_failerror : exn -> unit
+val catch_failerror : Exninfo.iexn -> unit
val tclORELSE0 : tactic -> tactic -> tactic
val tclORELSE : tactic -> tactic -> tactic
@@ -133,9 +117,9 @@ 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 tclTIMEOUT : int -> tactic -> tactic
val tclWEAK_PROGRESS : tactic -> tactic
val tclPROGRESS : tactic -> tactic
+val tclSHOWHYPS : tactic -> tactic
val tclNOTSAMEGOAL : tactic -> tactic
(** [tclIFTHENELSE tac1 tac2 tac3 gls] first applies [tac1] to [gls] then,
@@ -150,21 +134,3 @@ val tclIFTHENSVELSE : tactic -> tactic array -> tactic ->tactic
Equivalent to [(tac1;try tac2)||tac2] *)
val tclIFTHENTRYELSEMUST : tactic -> tactic -> tactic
-
-(** {6 Tactics handling a list of goals. } *)
-
-type tactic_list = goal list sigma -> goal list sigma
-
-val tclFIRSTLIST : tactic_list list -> tactic_list
-val tclIDTAC_list : tactic_list
-val first_goal : 'a list sigma -> 'a sigma
-val apply_tac_list : tactic -> tactic_list
-val then_tactic_list : tactic_list -> tactic_list -> tactic_list
-val tactic_list_tactic : tactic_list -> tactic
-val goal_goal_list : 'a sigma -> 'a list sigma
-
-(** [tclWITHHOLES solve_holes tac (sigma,c)] applies [tac] to [c] which
- may have unresolved holes; if [solve_holes] these holes must be
- resolved after application of the tactic; [sigma] must be an
- extension of the sigma of the goal *)
-val tclWITHHOLES : bool -> ('a -> tactic) -> evar_map -> 'a -> tactic
diff --git a/proofs/tacexpr.ml b/proofs/tacexpr.ml
deleted file mode 100644
index 536c966c..00000000
--- a/proofs/tacexpr.ml
+++ /dev/null
@@ -1,345 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-open Names
-open Topconstr
-open Libnames
-open Nametab
-open Glob_term
-open Util
-open Genarg
-open Pattern
-open Decl_kinds
-
-type 'a or_metaid = AI of 'a | MetaId of loc * string
-
-type direction_flag = bool (* true = Left-to-right false = right-to-right *)
-type lazy_flag = bool (* true = lazy false = eager *)
-type evars_flag = bool (* true = pose evars false = fail on evars *)
-type rec_flag = bool (* true = recursive false = not recursive *)
-type advanced_flag = bool (* true = advanced false = basic *)
-type split_flag = bool (* true = exists false = split *)
-type hidden_flag = bool (* true = internal use false = user-level *)
-type letin_flag = bool (* true = use local def false = use Leibniz *)
-
-type debug = Debug | Info | Off (* for trivial / auto / eauto ... *)
-
-type glob_red_flag =
- | FBeta
- | FIota
- | FZeta
- | FConst of reference or_by_notation list
- | FDeltaBut of reference or_by_notation list
-
-let make_red_flag =
- let rec add_flag red = function
- | [] -> red
- | FBeta :: lf -> add_flag { red with rBeta = true } lf
- | FIota :: lf -> add_flag { red with rIota = true } lf
- | FZeta :: lf -> add_flag { red with rZeta = true } lf
- | FConst l :: lf ->
- if red.rDelta then
- error
- "Cannot set both constants to unfold and constants not to unfold";
- add_flag { red with rConst = list_union red.rConst l } lf
- | FDeltaBut l :: lf ->
- if red.rConst <> [] & not red.rDelta then
- error
- "Cannot set both constants to unfold and constants not to unfold";
- add_flag
- { red with rConst = list_union red.rConst l; rDelta = true }
- lf
- in
- add_flag
- {rBeta = false; rIota = false; rZeta = false; rDelta = false; rConst = []}
-
-type 'a raw_hyp_location = 'a with_occurrences * Termops.hyp_location_flag
-
-type 'id move_location =
- | MoveAfter of 'id
- | MoveBefore of 'id
- | MoveToEnd of bool
-
-let no_move = MoveToEnd true
-
-open Pp
-
-let pr_move_location pr_id = function
- | MoveAfter id -> brk(1,1) ++ str "after " ++ pr_id id
- | MoveBefore id -> brk(1,1) ++ str "before " ++ pr_id id
- | MoveToEnd toleft -> str (if toleft then " at bottom" else " at top")
-
-type 'a induction_arg =
- | ElimOnConstr of 'a
- | ElimOnIdent of identifier located
- | ElimOnAnonHyp of int
-
-type inversion_kind =
- | SimpleInversion
- | FullInversion
- | FullInversionClear
-
-type ('c,'id) inversion_strength =
- | NonDepInversion of
- inversion_kind * 'id list * intro_pattern_expr located option
- | DepInversion of
- inversion_kind * 'c option * intro_pattern_expr located option
- | InversionUsing of 'c * 'id list
-
-type ('a,'b) location = HypLocation of 'a | ConclLocation of 'b
-
-type 'id message_token =
- | MsgString of string
- | MsgInt of int
- | MsgIdent of 'id
-
-(* onhyps:
- [None] means *on every hypothesis*
- [Some l] means on hypothesis belonging to l *)
-type 'id gclause =
- { onhyps : 'id raw_hyp_location list option;
- concl_occs : occurrences_expr }
-
-let nowhere = {onhyps=Some[]; concl_occs=no_occurrences_expr}
-
-type 'constr induction_clause =
- 'constr with_bindings induction_arg *
- (intro_pattern_expr located option (* eqn:... *)
- * intro_pattern_expr located option) (* as ... *)
-
-type ('constr,'id) induction_clause_list =
- 'constr induction_clause list
- * 'constr with_bindings option (* using ... *)
- * 'id gclause option (* in ... *)
-
-type multi =
- | Precisely of int
- | UpTo of int
- | RepeatStar
- | RepeatPlus
-
-(* Type of patterns *)
-type 'a match_pattern =
- | Term of 'a
- | Subterm of bool * identifier option * 'a
-
-(* Type of hypotheses for a Match Context rule *)
-type 'a match_context_hyps =
- | Hyp of name located * 'a match_pattern
- | Def of name located * 'a match_pattern * 'a match_pattern
-
-(* Type of a Match rule for Match Context and Match *)
-type ('a,'t) match_rule =
- | Pat of 'a match_context_hyps list * 'a match_pattern * 't
- | All of 't
-
-type ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_atomic_tactic_expr =
- (* Basic tactics *)
- | TacIntroPattern of intro_pattern_expr located list
- | TacIntrosUntil of quantified_hypothesis
- | TacIntroMove of identifier option * 'id move_location
- | TacAssumption
- | TacExact of 'constr
- | TacExactNoCheck of 'constr
- | TacVmCastNoCheck of 'constr
- | TacApply of advanced_flag * evars_flag * 'constr with_bindings list *
- ('id * intro_pattern_expr located option) option
- | TacElim of evars_flag * 'constr with_bindings *
- 'constr with_bindings option
- | TacElimType of 'constr
- | TacCase of evars_flag * 'constr with_bindings
- | TacCaseType of 'constr
- | TacFix of identifier option * int
- | TacMutualFix of hidden_flag * identifier * int * (identifier * int *
- 'constr) list
- | TacCofix of identifier option
- | TacMutualCofix of hidden_flag * identifier * (identifier * 'constr) list
- | TacCut of 'constr
- | TacAssert of 'tac option * intro_pattern_expr located option * 'constr
- | TacGeneralize of ('constr with_occurrences * name) list
- | TacGeneralizeDep of 'constr
- | TacLetTac of name * 'constr * 'id gclause * letin_flag *
- intro_pattern_expr located option
-
- (* Derived basic tactics *)
- | TacSimpleInductionDestruct of rec_flag * quantified_hypothesis
- | TacInductionDestruct of rec_flag * evars_flag * ('constr,'id) induction_clause_list
- | TacDoubleInduction of quantified_hypothesis * quantified_hypothesis
- | TacDecomposeAnd of 'constr
- | TacDecomposeOr of 'constr
- | TacDecompose of 'ind list * 'constr
- | TacSpecialize of int option * 'constr with_bindings
- | TacLApply of 'constr
-
- (* Automation tactics *)
- | TacTrivial of debug * 'constr list * string list option
- | TacAuto of debug * int or_var option * 'constr list * string list option
-
- (* Context management *)
- | TacClear of bool * 'id list
- | TacClearBody of 'id list
- | TacMove of bool * 'id * 'id move_location
- | TacRename of ('id *'id) list
- | TacRevert of 'id list
-
- (* Constructors *)
- | TacLeft of evars_flag * 'constr bindings
- | TacRight of evars_flag * 'constr bindings
- | TacSplit of evars_flag * split_flag * 'constr bindings list
- | TacAnyConstructor of evars_flag * 'tac option
- | TacConstructor of evars_flag * int or_var * 'constr bindings
-
- (* Conversion *)
- | TacReduce of ('constr,'cst,'pat) red_expr_gen * 'id gclause
- | TacChange of 'pat option * 'constr * 'id gclause
-
- (* Equivalence relations *)
- | TacReflexivity
- | TacSymmetry of 'id gclause
- | TacTransitivity of 'constr option
-
- (* Equality and inversion *)
- | TacRewrite of
- evars_flag * (bool * multi * 'constr with_bindings) list * 'id gclause * 'tac option
- | TacInversion of ('constr,'id) inversion_strength * quantified_hypothesis
-
- (* For ML extensions *)
- | TacExtend of loc * string * 'lev generic_argument list
-
- (* For syntax extensions *)
- | TacAlias of loc * string *
- (identifier * 'lev generic_argument) list
- * (dir_path * glob_tactic_expr)
-
-and ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr =
- | TacAtom of loc * ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_atomic_tactic_expr
- | TacThen of ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr *
- ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr array *
- ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr *
- ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr array
- | TacThens of ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr *
- ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr list
- | TacFirst of ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr list
- | TacComplete of ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr
- | TacSolve of ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr list
- | TacTry of ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr
- | TacOrelse of ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr * ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr
- | TacDo of int or_var * ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr
- | TacTimeout of int or_var * ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr
- | TacRepeat of ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr
- | TacProgress of ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr
- | TacAbstract of ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr * identifier option
- | TacId of 'id message_token list
- | TacFail of int or_var * 'id message_token list
- | TacInfo of ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr
- | TacLetIn of rec_flag * (identifier located * ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_arg) list * ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr
- | TacMatch of lazy_flag * ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr * ('pat,('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr) match_rule list
- | TacMatchGoal of lazy_flag * direction_flag * ('pat,('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr) match_rule list
- | TacFun of ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_fun_ast
- | TacArg of ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_arg located
-
-and ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_fun_ast =
- identifier option list * ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_expr
-
- (* These are the possible arguments of a tactic definition *)
-and ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_arg =
- | TacDynamic of loc * Dyn.t
- | TacVoid
- | MetaIdArg of loc * bool * string
- | ConstrMayEval of ('constr,'cst,'pat) may_eval
- | IntroPattern of intro_pattern_expr located
- | Reference of 'ref
- | Integer of int
- | TacCall of loc *
- 'ref * ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_arg list
- | TacExternal of loc * string * string *
- ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_arg list
- | TacFreshId of string or_var list
- | Tacexp of 'tac
-
-(* Globalized tactics *)
-and glob_tactic_expr =
- (glob_constr_and_expr,
- glob_constr_and_expr * constr_pattern,
- evaluable_global_reference and_short_name or_var,
- inductive or_var,
- ltac_constant located or_var,
- identifier located,
- glob_tactic_expr,
- glevel) gen_tactic_expr
-
-type raw_tactic_expr =
- (constr_expr,
- constr_pattern_expr,
- reference or_by_notation,
- reference or_by_notation,
- reference,
- identifier located or_metaid,
- raw_tactic_expr,
- rlevel) gen_tactic_expr
-
-type raw_atomic_tactic_expr =
- (constr_expr, (* constr *)
- constr_pattern_expr, (* pattern *)
- reference or_by_notation, (* evaluable reference *)
- reference or_by_notation, (* inductive *)
- reference, (* ltac reference *)
- identifier located or_metaid, (* identifier *)
- raw_tactic_expr,
- rlevel) gen_atomic_tactic_expr
-
-type raw_tactic_arg =
- (constr_expr,
- constr_pattern_expr,
- reference or_by_notation,
- reference or_by_notation,
- reference,
- identifier located or_metaid,
- raw_tactic_expr,
- rlevel) gen_tactic_arg
-
-type raw_generic_argument = rlevel generic_argument
-
-type raw_red_expr =
- (constr_expr, reference or_by_notation, constr_expr) red_expr_gen
-
-type glob_atomic_tactic_expr =
- (glob_constr_and_expr,
- glob_constr_and_expr * constr_pattern,
- evaluable_global_reference and_short_name or_var,
- inductive or_var,
- ltac_constant located or_var,
- identifier located,
- glob_tactic_expr,
- glevel) gen_atomic_tactic_expr
-
-type glob_tactic_arg =
- (glob_constr_and_expr,
- glob_constr_and_expr * constr_pattern,
- evaluable_global_reference and_short_name or_var,
- inductive or_var,
- ltac_constant located or_var,
- identifier located,
- glob_tactic_expr,
- glevel) gen_tactic_arg
-
-type glob_generic_argument = glevel generic_argument
-
-type glob_red_expr =
- (glob_constr_and_expr, evaluable_global_reference or_var, constr_pattern)
- red_expr_gen
-
-type typed_generic_argument = tlevel generic_argument
-
-type 'a raw_abstract_argument_type = ('a,rlevel) abstract_argument_type
-
-type 'a glob_abstract_argument_type = ('a,glevel) abstract_argument_type
-
-type 'a typed_abstract_argument_type = ('a,tlevel) abstract_argument_type
-
-type declaration_hook = locality -> global_reference -> unit
diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml
index 8346e4c2..fa0d0362 100644
--- a/proofs/tacmach.ml
+++ b/proofs/tacmach.ml
@@ -1,17 +1,13 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Pp
open Util
-open Names
open Namegen
-open Sign
-open Term
open Termops
open Environ
open Reductionops
@@ -22,9 +18,8 @@ open Tacred
open Proof_type
open Logic
open Refiner
-open Tacexpr
-let re_sig it gc = { it = it; sigma = gc }
+let re_sig it gc = { it = it; sigma = gc; }
(**************************************************************)
(* Operations for handling terms under a local typing context *)
@@ -53,9 +48,9 @@ let pf_last_hyp gl = List.hd (pf_hyps gl)
let pf_get_hyp gls id =
try
- Sign.lookup_named id (pf_hyps gls)
+ Context.lookup_named id (pf_hyps gls)
with Not_found ->
- error ("No such hypothesis: " ^ (string_of_id id))
+ raise (RefinerError (NoSuchHyp id))
let pf_get_hyp_typ gls id =
let (_,_,ty)= (pf_get_hyp gls id) in
@@ -74,18 +69,17 @@ let pf_get_new_ids ids gls =
let pf_global gls id = Constrintern.construct_reference (pf_hyps gls) id
-let pf_parse_const gls = compose (pf_global gls) id_of_string
-
let pf_reduction_of_red_expr gls re c =
- (fst (reduction_of_red_expr re)) (pf_env gls) (project gls) c
+ (fst (reduction_of_red_expr (pf_env gls) re)) (pf_env gls) (project gls) c
let pf_apply f gls = f (pf_env gls) (project gls)
+let pf_eapply f gls x =
+ on_sig gls (fun evm -> f (pf_env gls) evm x)
let pf_reduce = pf_apply
+let pf_e_reduce = pf_apply
let pf_whd_betadeltaiota = pf_reduce whd_betadeltaiota
-let pf_whd_betadeltaiota_stack = pf_reduce whd_betadeltaiota_stack
let pf_hnf_constr = pf_reduce hnf_constr
-let pf_red_product = pf_reduce red_product
let pf_nf = pf_reduce simpl
let pf_nf_betaiota = pf_reduce (fun _ -> nf_betaiota)
let pf_compute = pf_reduce compute
@@ -93,38 +87,17 @@ let pf_unfoldn ubinds = pf_reduce (unfoldn ubinds)
let pf_type_of = pf_reduce type_of
let pf_get_type_of = pf_reduce Retyping.get_type_of
-let pf_conv_x = pf_reduce is_conv
-let pf_conv_x_leq = pf_reduce is_conv_leq
-let pf_const_value = pf_reduce (fun env _ -> constant_value env)
+let pf_conv_x gl = pf_reduce test_conversion gl Reduction.CONV
+let pf_conv_x_leq gl = pf_reduce test_conversion gl Reduction.CUMUL
+let pf_const_value = pf_reduce (fun env _ -> constant_value_in env)
+
let pf_reduce_to_quantified_ind = pf_reduce reduce_to_quantified_ind
let pf_reduce_to_atomic_ind = pf_reduce reduce_to_atomic_ind
let pf_hnf_type_of gls = compose (pf_whd_betadeltaiota gls) (pf_get_type_of gls)
-let pf_check_type gls c1 c2 =
- ignore (pf_type_of gls (mkCast (c1, DEFAULTcast, c2)))
-
-let pf_is_matching = pf_apply Matching.is_matching_conv
-let pf_matches = pf_apply Matching.matches_conv
-
-(************************************)
-(* Tactics handling a list of goals *)
-(************************************)
-
-type transformation_tactic = proof_tree -> goal list
-
-type validation_list = proof_tree list -> proof_tree list
-
-type tactic_list = Refiner.tactic_list
-
-let first_goal = first_goal
-let goal_goal_list = goal_goal_list
-let apply_tac_list = apply_tac_list
-let then_tactic_list = then_tactic_list
-let tactic_list_tactic = tactic_list_tactic
-let tclFIRSTLIST = tclFIRSTLIST
-let tclIDTAC_list = tclIDTAC_list
-
+let pf_is_matching = pf_apply Constr_matching.is_matching_conv
+let pf_matches = pf_apply Constr_matching.matches_conv
(********************************************)
(* Definition of the most primitive tactics *)
@@ -132,69 +105,39 @@ let tclIDTAC_list = tclIDTAC_list
let refiner = refiner
-(* This does not check that the variable name is not here *)
-let introduction_no_check id =
- refiner (Prim (Intro id))
-
let internal_cut_no_check replace id t gl =
- refiner (Prim (Cut (true,replace,id,t))) gl
+ refiner (Cut (true,replace,id,t)) gl
let internal_cut_rev_no_check replace id t gl =
- refiner (Prim (Cut (false,replace,id,t))) gl
+ refiner (Cut (false,replace,id,t)) gl
let refine_no_check c gl =
- refiner (Prim (Refine c)) gl
-
-let convert_concl_no_check c sty gl =
- refiner (Prim (Convert_concl (c,sty))) gl
-
-let convert_hyp_no_check d gl =
- refiner (Prim (Convert_hyp d)) gl
+ refiner (Refine c) gl
(* This does not check dependencies *)
let thin_no_check ids gl =
- if ids = [] then tclIDTAC gl else refiner (Prim (Thin ids)) gl
-
-(* This does not check dependencies *)
-let thin_body_no_check ids gl =
- if ids = [] then tclIDTAC gl else refiner (Prim (ThinBody ids)) gl
+ if List.is_empty ids then tclIDTAC gl else refiner (Thin ids) gl
-let move_hyp_no_check with_dep id1 id2 gl =
- refiner (Prim (Move (with_dep,id1,id2))) gl
-
-let order_hyps idl gl =
- refiner (Prim (Order idl)) gl
-
-let rec rename_hyp_no_check l gl = match l with
- | [] -> tclIDTAC gl
- | (id1,id2)::l ->
- tclTHEN (refiner (Prim (Rename (id1,id2))))
- (rename_hyp_no_check l) gl
+let move_hyp_no_check id1 id2 gl =
+ refiner (Move (id1,id2)) gl
let mutual_fix f n others j gl =
- with_check (refiner (Prim (FixRule (f,n,others,j)))) gl
+ with_check (refiner (FixRule (f,n,others,j))) gl
let mutual_cofix f others j gl =
- with_check (refiner (Prim (Cofix (f,others,j)))) gl
+ with_check (refiner (Cofix (f,others,j))) gl
(* Versions with consistency checks *)
-let introduction id = with_check (introduction_no_check id)
let internal_cut b d t = with_check (internal_cut_no_check b d t)
let internal_cut_rev b d t = with_check (internal_cut_rev_no_check b d t)
let refine c = with_check (refine_no_check c)
-let convert_concl d sty = with_check (convert_concl_no_check d sty)
-let convert_hyp d = with_check (convert_hyp_no_check d)
let thin c = with_check (thin_no_check c)
-let thin_body c = with_check (thin_body_no_check c)
-let move_hyp b id id' = with_check (move_hyp_no_check b id id')
-let rename_hyp l = with_check (rename_hyp_no_check l)
+let move_hyp id id' = with_check (move_hyp_no_check id id')
(* Pretty-printers *)
open Pp
-open Tacexpr
-open Glob_term
let db_pr_goal sigma g =
let env = Goal.V82.env sigma g in
@@ -209,5 +152,85 @@ let pr_gls gls =
let pr_glls glls =
hov 0 (pr_evar_map (Some 2) (sig_sig glls) ++ fnl () ++
- prlist_with_sep pr_fnl (db_pr_goal (project glls)) (sig_it glls))
+ prlist_with_sep fnl (db_pr_goal (project glls)) (sig_it glls))
+
+(* Variants of [Tacmach] functions built with the new proof engine *)
+module New = struct
+
+ let pf_apply f gl =
+ f (Proofview.Goal.env gl) (Proofview.Goal.sigma gl)
+
+ let of_old f gl =
+ f { Evd.it = Proofview.Goal.goal gl ; sigma = Proofview.Goal.sigma gl }
+
+ let pf_global id gl =
+ (** We only check for the existence of an [id] in [hyps] *)
+ let gl = Proofview.Goal.assume gl in
+ let hyps = Proofview.Goal.hyps gl in
+ Constrintern.construct_reference hyps id
+
+ let pf_env = Proofview.Goal.env
+ let pf_concl = Proofview.Goal.concl
+
+ let pf_type_of gl t =
+ pf_apply type_of gl t
+
+ let pf_conv_x gl t1 t2 = pf_apply is_conv gl t1 t2
+
+ let pf_ids_of_hyps gl =
+ (** We only get the identifiers in [hyps] *)
+ let gl = Proofview.Goal.assume gl in
+ let hyps = Proofview.Goal.hyps gl in
+ ids_of_named_context hyps
+
+ let pf_get_new_id id gl =
+ let ids = pf_ids_of_hyps gl in
+ next_ident_away id ids
+
+ let pf_get_hyp id gl =
+ let hyps = Proofview.Goal.hyps gl in
+ let sign =
+ try Context.lookup_named id hyps
+ with Not_found -> raise (RefinerError (NoSuchHyp id))
+ in
+ sign
+
+ let pf_get_hyp_typ id gl =
+ let (_,_,ty) = pf_get_hyp id gl in
+ ty
+
+ let pf_hyps_types gl =
+ let env = Proofview.Goal.env gl in
+ let sign = Environ.named_context env in
+ List.map (fun (id,_,x) -> (id, x)) sign
+
+ let pf_last_hyp gl =
+ let hyps = Proofview.Goal.hyps gl in
+ List.hd hyps
+
+ let pf_nf_concl (gl : [ `LZ ] Proofview.Goal.t) =
+ (** We normalize the conclusion just after *)
+ let gl = Proofview.Goal.assume gl in
+ let concl = Proofview.Goal.concl gl in
+ let sigma = Proofview.Goal.sigma gl in
+ nf_evar sigma concl
+
+ let pf_whd_betadeltaiota gl t = pf_apply whd_betadeltaiota gl t
+
+ let pf_get_type_of gl t = pf_apply Retyping.get_type_of gl t
+
+ let pf_reduce_to_quantified_ind gl t =
+ pf_apply reduce_to_quantified_ind gl t
+
+ let pf_hnf_constr gl t = pf_apply hnf_constr gl t
+ let pf_hnf_type_of gl t =
+ pf_whd_betadeltaiota gl (pf_get_type_of gl t)
+
+ let pf_matches gl pat t = pf_apply Constr_matching.matches_conv gl pat t
+
+ let pf_whd_betadeltaiota gl t = pf_apply whd_betadeltaiota gl t
+ let pf_compute gl t = pf_apply compute gl t
+
+ let pf_nf_evar gl t = nf_evar (Proofview.Goal.sigma gl) t
+end
diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli
index 39ecbd3b..f7fc6b54 100644
--- a/proofs/tacmach.mli
+++ b/proofs/tacmach.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,16 +8,14 @@
open Names
open Term
-open Sign
+open Context
open Environ
open Evd
-open Reduction
open Proof_type
-open Refiner
open Redexpr
-open Tacexpr
-open Glob_term
open Pattern
+open Locus
+open Misctypes
(** Operations for handling terms under a local typing context. *)
@@ -32,49 +30,50 @@ 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) -> goal -> (goal list)
+ evar_map ref -> (goal sigma -> (goal list) sigma) -> goal -> (goal list)
val pf_concl : goal sigma -> types
val pf_env : goal sigma -> env
val pf_hyps : goal sigma -> named_context
-(*i val pf_untyped_hyps : goal sigma -> (identifier * constr) list i*)
-val pf_hyps_types : goal sigma -> (identifier * types) list
-val pf_nth_hyp_id : goal sigma -> int -> identifier
+(*i val pf_untyped_hyps : goal sigma -> (Id.t * constr) list i*)
+val pf_hyps_types : goal sigma -> (Id.t * types) list
+val pf_nth_hyp_id : goal sigma -> int -> Id.t
val pf_last_hyp : goal sigma -> named_declaration
-val pf_ids_of_hyps : goal sigma -> identifier list
-val pf_global : goal sigma -> identifier -> constr
-val pf_parse_const : goal sigma -> string -> constr
+val pf_ids_of_hyps : goal sigma -> Id.t list
+val pf_global : goal sigma -> Id.t -> constr
val pf_type_of : goal sigma -> constr -> types
-val pf_check_type : goal sigma -> constr -> types -> unit
val pf_hnf_type_of : goal sigma -> constr -> types
-val pf_get_hyp : goal sigma -> identifier -> named_declaration
-val pf_get_hyp_typ : goal sigma -> identifier -> types
+val pf_get_hyp : goal sigma -> Id.t -> named_declaration
+val pf_get_hyp_typ : goal sigma -> Id.t -> types
-val pf_get_new_id : identifier -> goal sigma -> identifier
-val pf_get_new_ids : identifier list -> goal sigma -> identifier list
+val pf_get_new_id : Id.t -> goal sigma -> Id.t
+val pf_get_new_ids : Id.t list -> goal sigma -> Id.t list
-val pf_reduction_of_red_expr : goal sigma -> red_expr -> constr -> constr
+val pf_reduction_of_red_expr : goal sigma -> red_expr -> constr -> evar_map * constr
val pf_apply : (env -> evar_map -> 'a) -> goal sigma -> 'a
+val pf_eapply : (env -> evar_map -> 'a -> evar_map * 'b) ->
+ goal sigma -> 'a -> goal sigma * 'b
val pf_reduce :
(env -> evar_map -> constr -> constr) ->
- goal sigma -> constr -> constr
+ goal sigma -> constr -> constr
+val pf_e_reduce :
+ (env -> evar_map -> constr -> evar_map * constr) ->
+ goal sigma -> constr -> evar_map * constr
val pf_whd_betadeltaiota : goal sigma -> constr -> constr
-val pf_whd_betadeltaiota_stack : goal sigma -> constr -> constr * constr list
val pf_hnf_constr : goal sigma -> constr -> constr
-val pf_red_product : goal sigma -> constr -> constr
val pf_nf : goal sigma -> constr -> constr
val pf_nf_betaiota : goal sigma -> constr -> constr
-val pf_reduce_to_quantified_ind : goal sigma -> types -> inductive * types
-val pf_reduce_to_atomic_ind : goal sigma -> types -> inductive * types
+val pf_reduce_to_quantified_ind : goal sigma -> types -> pinductive * types
+val pf_reduce_to_atomic_ind : goal sigma -> types -> pinductive * types
val pf_compute : goal sigma -> constr -> constr
-val pf_unfoldn : (Termops.occurrences * evaluable_global_reference) list
- -> goal sigma -> constr -> constr
+val pf_unfoldn : (occurrences * evaluable_global_reference) list
+ -> goal sigma -> constr -> constr
-val pf_const_value : goal sigma -> constant -> constr
+val pf_const_value : goal sigma -> pconstant -> constr
val pf_conv_x : goal sigma -> constr -> constr -> bool
val pf_conv_x_leq : goal sigma -> constr -> constr -> bool
@@ -85,49 +84,56 @@ val pf_is_matching : goal sigma -> constr_pattern -> constr -> bool
(** {6 The most primitive tactics. } *)
val refiner : rule -> tactic
-val introduction_no_check : identifier -> tactic
-val internal_cut_no_check : bool -> identifier -> types -> tactic
-val internal_cut_rev_no_check : bool -> identifier -> types -> tactic
+val internal_cut_no_check : bool -> Id.t -> types -> tactic
val refine_no_check : constr -> tactic
-val convert_concl_no_check : types -> cast_kind -> tactic
-val convert_hyp_no_check : named_declaration -> tactic
-val thin_no_check : identifier list -> tactic
-val thin_body_no_check : identifier list -> tactic
-val move_hyp_no_check :
- bool -> identifier -> identifier move_location -> tactic
-val rename_hyp_no_check : (identifier*identifier) list -> tactic
-val order_hyps : identifier list -> tactic
+val thin_no_check : Id.t list -> tactic
val mutual_fix :
- identifier -> int -> (identifier * int * constr) list -> int -> tactic
-val mutual_cofix : identifier -> (identifier * constr) list -> int -> tactic
+ Id.t -> int -> (Id.t * int * constr) list -> int -> tactic
+val mutual_cofix : Id.t -> (Id.t * constr) list -> int -> tactic
(** {6 The most primitive tactics with consistency and type checking } *)
-val introduction : identifier -> tactic
-val internal_cut : bool -> identifier -> types -> tactic
-val internal_cut_rev : bool -> identifier -> types -> tactic
+val internal_cut : bool -> Id.t -> types -> tactic
+val internal_cut_rev : bool -> Id.t -> types -> tactic
val refine : constr -> tactic
-val convert_concl : types -> cast_kind -> tactic
-val convert_hyp : named_declaration -> tactic
-val thin : identifier list -> tactic
-val thin_body : identifier list -> tactic
-val move_hyp : bool -> identifier -> identifier move_location -> tactic
-val rename_hyp : (identifier*identifier) list -> tactic
+val thin : Id.t list -> tactic
+val move_hyp : Id.t -> Id.t move_location -> tactic
-(** {6 Tactics handling a list of goals. } *)
+(** {6 Pretty-printing functions (debug only). } *)
+val pr_gls : goal sigma -> Pp.std_ppcmds
+val pr_glls : goal list sigma -> Pp.std_ppcmds
-type validation_list = proof_tree list -> proof_tree list
+(* Variants of [Tacmach] functions built with the new proof engine *)
+module New : sig
+ val pf_apply : (env -> evar_map -> 'a) -> 'b Proofview.Goal.t -> 'a
+ val pf_global : identifier -> 'a Proofview.Goal.t -> constr
+ val of_old : (Proof_type.goal Evd.sigma -> 'a) -> [ `NF ] Proofview.Goal.t -> 'a
-type tactic_list = Refiner.tactic_list
+ val pf_env : 'a Proofview.Goal.t -> Environ.env
+ val pf_concl : [ `NF ] Proofview.Goal.t -> types
-val first_goal : 'a list sigma -> 'a sigma
-val goal_goal_list : 'a sigma -> 'a list sigma
-val apply_tac_list : tactic -> tactic_list
-val then_tactic_list : tactic_list -> tactic_list -> tactic_list
-val tactic_list_tactic : tactic_list -> tactic
-val tclFIRSTLIST : tactic_list list -> tactic_list
-val tclIDTAC_list : tactic_list
+ val pf_type_of : 'a Proofview.Goal.t -> Term.constr -> Term.types
+ val pf_conv_x : 'a Proofview.Goal.t -> Term.constr -> Term.constr -> bool
-(** {6 Pretty-printing functions (debug only). } *)
-val pr_gls : goal sigma -> Pp.std_ppcmds
-val pr_glls : goal list sigma -> Pp.std_ppcmds
+ val pf_get_new_id : identifier -> [ `NF ] Proofview.Goal.t -> identifier
+ val pf_ids_of_hyps : 'a Proofview.Goal.t -> identifier list
+ val pf_hyps_types : 'a Proofview.Goal.t -> (identifier * types) list
+
+ val pf_get_hyp : identifier -> [ `NF ] Proofview.Goal.t -> named_declaration
+ val pf_get_hyp_typ : identifier -> [ `NF ] Proofview.Goal.t -> types
+ val pf_last_hyp : [ `NF ] Proofview.Goal.t -> named_declaration
+
+ val pf_nf_concl : [ `LZ ] Proofview.Goal.t -> types
+ val pf_reduce_to_quantified_ind : 'a Proofview.Goal.t -> types -> pinductive * types
+
+ val pf_hnf_constr : 'a Proofview.Goal.t -> constr -> types
+ val pf_hnf_type_of : 'a Proofview.Goal.t -> constr -> types
+
+ val pf_whd_betadeltaiota : 'a Proofview.Goal.t -> constr -> constr
+ val pf_compute : 'a Proofview.Goal.t -> constr -> constr
+
+ val pf_matches : 'a Proofview.Goal.t -> constr_pattern -> constr -> patvar_map
+
+ val pf_nf_evar : 'a Proofview.Goal.t -> constr -> constr
+
+end
diff --git a/proofs/tactic_debug.ml b/proofs/tactic_debug.ml
index a7df80f6..3cc81daf 100644
--- a/proofs/tactic_debug.ml
+++ b/proofs/tactic_debug.ml
@@ -1,23 +1,21 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Util
open Names
-open Constrextern
open Pp
open Tacexpr
open Termops
-let prtac = ref (fun _ -> assert false)
-let set_tactic_printer f = prtac := f
-let prmatchpatt = ref (fun _ _ -> assert false)
-let set_match_pattern_printer f = prmatchpatt := f
-let prmatchrl = ref (fun _ -> assert false)
-let set_match_rule_printer f = prmatchrl := f
+let (prtac, tactic_printer) = Hook.make ()
+let (prmatchpatt, match_pattern_printer) = Hook.make ()
+let (prmatchrl, match_rule_printer) = Hook.make ()
+
(* This module intends to be a beginning of debugger for tactic expressions.
Currently, it is quite simple and we can hope to have, in the future, a more
@@ -33,23 +31,29 @@ let explain_logic_error = ref (fun e -> mt())
let explain_logic_error_no_anomaly = ref (fun e -> mt())
+let msg_tac_debug s = Proofview.NonLogical.print (s++fnl())
+
(* Prints the goal *)
-let db_pr_goal g =
- let env = Refiner.pf_env g in
+let db_pr_goal gl =
+ let env = Proofview.Goal.env gl in
+ let concl = Proofview.Goal.concl gl 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 () ++
+ let pc = print_constr_env env concl in
+ str" " ++ hv 0 (penv ++ fnl () ++
str "============================" ++ fnl () ++
str" " ++ pc) ++ fnl ()
-let db_pr_goal g =
- msgnl (str "Goal:" ++ fnl () ++ db_pr_goal g)
+let db_pr_goal =
+ Proofview.Goal.nf_enter begin fun gl ->
+ let pg = db_pr_goal gl in
+ Proofview.tclLIFT (msg_tac_debug (str "Goal:" ++ fnl () ++ pg))
+ end
(* Prints the commands *)
let help () =
- msgnl (str "Commands: <Enter> = Continue" ++ fnl() ++
+ msg_tac_debug (str "Commands: <Enter> = Continue" ++ fnl() ++
str " h/? = Help" ++ fnl() ++
str " r <num> = Run <num> times" ++ fnl() ++
str " r <string> = Run up to next idtac <string>" ++ fnl() ++
@@ -57,168 +61,257 @@ let help () =
str " x = Exit")
(* Prints the goal and the command to be executed *)
-let goal_com g tac =
- begin
- db_pr_goal g;
- msg (str "Going to execute:" ++ fnl () ++ !prtac tac ++ fnl ())
- end
-
-let skipped = ref 0
-let skip = ref 0
-let breakpoint = ref None
+let goal_com tac =
+ Proofview.tclTHEN
+ db_pr_goal
+ (Proofview.tclLIFT (msg_tac_debug (str "Going to execute:" ++ fnl () ++ Hook.get prtac tac)))
+
+(* [run (new_ref _)] gives us a ref shared among [NonLogical.t]
+ expressions. It avoids parametrizing everything over a
+ reference. *)
+let skipped = Proofview.NonLogical.run (Proofview.NonLogical.ref 0)
+let skip = Proofview.NonLogical.run (Proofview.NonLogical.ref 0)
+let breakpoint = Proofview.NonLogical.run (Proofview.NonLogical.ref None)
let rec drop_spaces inst i =
- if String.length inst > i && inst.[i] = ' ' then drop_spaces inst (i+1)
+ if String.length inst > i && inst.[i] == ' ' then drop_spaces inst (i+1)
else i
let possibly_unquote s =
- if String.length s >= 2 & s.[0] = '"' & s.[String.length s - 1] = '"' then
+ if String.length s >= 2 && s.[0] == '"' && s.[String.length s - 1] == '"' then
String.sub s 1 (String.length s - 2)
else
s
(* (Re-)initialize debugger *)
-let db_initialize () =
- skip:=0;skipped:=0;breakpoint:=None
+let db_initialize =
+ let open Proofview.NonLogical in
+ (skip:=0) >> (skipped:=0) >> (breakpoint:=None)
+
+let int_of_string s =
+ try Proofview.NonLogical.return (int_of_string s)
+ with e -> Proofview.NonLogical.raise e
+
+let string_get s i =
+ try Proofview.NonLogical.return (String.get s i)
+ with e -> Proofview.NonLogical.raise e
(* Gives the number of steps or next breakpoint of a run command *)
let run_com inst =
- if (String.get inst 0)='r' then
+ let open Proofview.NonLogical in
+ string_get inst 0 >>= fun first_char ->
+ if first_char ='r' then
let i = drop_spaces inst 1 in
if String.length inst > i then
let s = String.sub inst i (String.length inst - i) in
if inst.[0] >= '0' && inst.[0] <= '9' then
- let num = int_of_string s in
- if num<0 then raise (Invalid_argument "run_com");
- skip:=num;skipped:=0
+ int_of_string s >>= fun num ->
+ (if num<0 then invalid_arg "run_com" else return ()) >>
+ (skip:=num) >> (skipped:=0)
else
breakpoint:=Some (possibly_unquote s)
else
- raise (Invalid_argument "run_com")
+ invalid_arg "run_com"
else
- raise (Invalid_argument "run_com")
+ invalid_arg "run_com"
(* Prints the run counter *)
let run ini =
+ let open Proofview.NonLogical in
if not ini then
- begin
- for i=1 to 2 do
- print_char (Char.chr 8);print_char (Char.chr 13)
- done;
- msg (str "Executed expressions: " ++ int !skipped ++ fnl() ++ fnl())
- end;
- incr skipped
+ begin
+ Proofview.NonLogical.print (str"\b\r\b\r") >>
+ !skipped >>= fun skipped ->
+ msg_tac_debug (str "Executed expressions: " ++ int skipped ++ fnl())
+ end >>
+ !skipped >>= fun x ->
+ skipped := x+1
+ else
+ return ()
(* Prints the prompt *)
let rec prompt level =
+ (* spiwack: avoid overriding by the open below *)
+ let runtrue = run true in
begin
- msg (fnl () ++ str "TcDebug (" ++ int level ++ str ") > ");
- flush stdout;
- let exit () = skip:=0;skipped:=0;raise Sys.Break in
- let inst = try read_line () with End_of_file -> exit () in
+ let open Proofview.NonLogical in
+ Proofview.NonLogical.print (fnl () ++ str "TcDebug (" ++ int level ++ str ") > ") >>
+ let exit = (skip:=0) >> (skipped:=0) >> raise Sys.Break in
+ Proofview.NonLogical.catch Proofview.NonLogical.read_line
+ begin function (e, info) -> match e with
+ | End_of_file -> exit
+ | e -> raise ~info e
+ end
+ >>= fun inst ->
match inst with
- | "" -> DebugOn (level+1)
- | "s" -> DebugOff
- | "x" -> print_char (Char.chr 8); exit ()
+ | "" -> return (DebugOn (level+1))
+ | "s" -> return (DebugOff)
+ | "x" -> Proofview.NonLogical.print_char '\b' >> exit
| "h"| "?" ->
begin
- help ();
+ help () >>
prompt level
end
| _ ->
- (try run_com inst;run true;DebugOn (level+1)
- with Failure _ | Invalid_argument _ -> prompt level)
+ Proofview.NonLogical.catch (run_com inst >> runtrue >> return (DebugOn (level+1)))
+ begin function (e, info) -> match e with
+ | Failure _ | Invalid_argument _ -> prompt level
+ | e -> raise ~info e
+ end
end
(* Prints the state and waits for an instruction *)
-let debug_prompt lev g tac f =
+(* spiwack: the only reason why we need to take the continuation [f]
+ as an argument rather than returning the new level directly seems to
+ be that [f] is wrapped in with "explain_logic_error". I don't think
+ it serves any purpose in the current design, so we could just drop
+ that. *)
+let debug_prompt lev tac f =
+ (* spiwack: avoid overriding by the open below *)
+ let runfalse = run false in
+ let open Proofview.NonLogical in
+ let (>=) = Proofview.tclBIND in
(* What to print and to do next *)
let newlevel =
- if !skip = 0 then
- if !breakpoint = None then (goal_com g tac; prompt lev)
- else (run false; DebugOn (lev+1))
- else (decr skip; run false; if !skip=0 then skipped:=0; DebugOn (lev+1)) in
+ Proofview.tclLIFT !skip >= fun initial_skip ->
+ if Int.equal initial_skip 0 then
+ Proofview.tclLIFT !breakpoint >= fun breakpoint ->
+ if Option.is_empty breakpoint then Proofview.tclTHEN (goal_com tac) (Proofview.tclLIFT (prompt lev))
+ else Proofview.tclLIFT(runfalse >> return (DebugOn (lev+1)))
+ else Proofview.tclLIFT begin
+ (!skip >>= fun s -> skip:=s-1) >>
+ runfalse >>
+ !skip >>= fun new_skip ->
+ (if Int.equal new_skip 0 then skipped:=0 else return ()) >>
+ return (DebugOn (lev+1))
+ end in
+ newlevel >= fun newlevel ->
(* What to execute *)
- try f newlevel
- with reraise ->
- skip:=0; skipped:=0;
- if Logic.catchable_exception reraise then
- ppnl (str "Level " ++ int lev ++ str ": " ++ !explain_logic_error reraise);
- raise reraise
+ Proofview.tclOR
+ (f newlevel)
+ begin fun (reraise, info) ->
+ Proofview.tclTHEN
+ (Proofview.tclLIFT begin
+ (skip:=0) >> (skipped:=0) >>
+ if Logic.catchable_exception reraise then
+ msg_tac_debug (str "Level " ++ int lev ++ str ": " ++ Pervasives.(!) explain_logic_error reraise)
+ else return ()
+ end)
+ (Proofview.tclZERO ~info reraise)
+ end
+
+let is_debug db =
+ let open Proofview.NonLogical in
+ !breakpoint >>= fun breakpoint ->
+ match db, breakpoint with
+ | DebugOff, _ -> return false
+ | _, Some _ -> return false
+ | _ ->
+ !skip >>= fun skip ->
+ return (Int.equal skip 0)
(* Prints a constr *)
let db_constr debug env c =
- if debug <> DebugOff & !skip = 0 & !breakpoint = None then
- msgnl (str "Evaluated term: " ++ print_constr_env env c)
+ let open Proofview.NonLogical in
+ is_debug debug >>= fun db ->
+ if db then
+ msg_tac_debug (str "Evaluated term: " ++ print_constr_env env c)
+ else return ()
(* Prints the pattern rule *)
let db_pattern_rule debug num r =
- if debug <> DebugOff & !skip = 0 & !breakpoint = None then
+ let open Proofview.NonLogical in
+ is_debug debug >>= fun db ->
+ if db then
begin
- msgnl (str "Pattern rule " ++ int num ++ str ":");
- msgnl (str "|" ++ spc () ++ !prmatchrl r)
+ msg_tac_debug (str "Pattern rule " ++ int num ++ str ":" ++ fnl () ++
+ str "|" ++ spc () ++ Hook.get prmatchrl r)
end
+ else return ()
(* Prints the hypothesis pattern identifier if it exists *)
let hyp_bound = function
| Anonymous -> " (unbound)"
- | Name id -> " (bound to "^(Names.string_of_id id)^")"
+ | Name id -> " (bound to "^(Names.Id.to_string id)^")"
(* Prints a matched hypothesis *)
let db_matched_hyp debug env (id,_,c) ido =
- if debug <> DebugOff & !skip = 0 & !breakpoint = None then
- msgnl (str "Hypothesis " ++
- str ((Names.string_of_id id)^(hyp_bound ido)^
+ let open Proofview.NonLogical in
+ is_debug debug >>= fun db ->
+ if db then
+ msg_tac_debug (str "Hypothesis " ++
+ str ((Names.Id.to_string id)^(hyp_bound ido)^
" has been matched: ") ++ print_constr_env env c)
+ else return ()
(* Prints the matched conclusion *)
let db_matched_concl debug env c =
- if debug <> DebugOff & !skip = 0 & !breakpoint = None then
- msgnl (str "Conclusion has been matched: " ++ print_constr_env env c)
+ let open Proofview.NonLogical in
+ is_debug debug >>= fun db ->
+ if db then
+ msg_tac_debug (str "Conclusion has been matched: " ++ print_constr_env env c)
+ else return ()
(* Prints a success message when the goal has been matched *)
let db_mc_pattern_success debug =
- if debug <> DebugOff & !skip = 0 & !breakpoint = None then
- msgnl (str "The goal has been successfully matched!" ++ fnl() ++
+ let open Proofview.NonLogical in
+ is_debug debug >>= fun db ->
+ if db then
+ msg_tac_debug (str "The goal has been successfully matched!" ++ fnl() ++
str "Let us execute the right-hand side part..." ++ fnl())
+ else return ()
(* Prints a failure message for an hypothesis pattern *)
-let db_hyp_pattern_failure debug env (na,hyp) =
- if debug <> DebugOff & !skip = 0 & !breakpoint = None then
- msgnl (str ("The pattern hypothesis"^(hyp_bound na)^
+let db_hyp_pattern_failure debug env sigma (na,hyp) =
+ let open Proofview.NonLogical in
+ is_debug debug >>= fun db ->
+ if db then
+ msg_tac_debug (str ("The pattern hypothesis"^(hyp_bound na)^
" cannot match: ") ++
- !prmatchpatt env hyp)
+ Hook.get prmatchpatt env sigma hyp)
+ else return ()
(* Prints a matching failure message for a rule *)
let db_matching_failure debug =
- if debug <> DebugOff & !skip = 0 & !breakpoint = None then
- msgnl (str "This rule has failed due to matching errors!" ++ fnl() ++
+ let open Proofview.NonLogical in
+ is_debug debug >>= fun db ->
+ if db then
+ msg_tac_debug (str "This rule has failed due to matching errors!" ++ fnl() ++
str "Let us try the next one...")
+ else return ()
(* Prints an evaluation failure message for a rule *)
let db_eval_failure debug s =
- if debug <> DebugOff & !skip = 0 & !breakpoint = None then
+ let open Proofview.NonLogical in
+ is_debug debug >>= fun db ->
+ if db then
let s = str "message \"" ++ s ++ str "\"" in
- msgnl
+ msg_tac_debug
(str "This rule has failed due to \"Fail\" tactic (" ++
s ++ str ", level 0)!" ++ fnl() ++ str "Let us try the next one...")
+ else return ()
(* Prints a logic failure message for a rule *)
let db_logic_failure debug err =
- if debug <> DebugOff & !skip = 0 & !breakpoint = None then
+ let open Proofview.NonLogical in
+ is_debug debug >>= fun db ->
+ if db then
begin
- msgnl (!explain_logic_error err);
- msgnl (str "This rule has failed due to a logic error!" ++ fnl() ++
+ msg_tac_debug (Pervasives.(!) explain_logic_error err) >>
+ msg_tac_debug (str "This rule has failed due to a logic error!" ++ fnl() ++
str "Let us try the next one...")
end
+ else return ()
let is_breakpoint brkname s = match brkname, s with
- | Some s, MsgString s'::_ -> s = s'
+ | Some s, MsgString s'::_ -> String.equal s s'
| _ -> false
let db_breakpoint debug s =
+ let open Proofview.NonLogical in
+ !breakpoint >>= fun opt_breakpoint ->
match debug with
- | DebugOn lev when s <> [] & is_breakpoint !breakpoint s ->
+ | DebugOn lev when not (CList.is_empty s) && is_breakpoint opt_breakpoint s ->
breakpoint:=None
| _ ->
- ()
+ return ()
diff --git a/proofs/tactic_debug.mli b/proofs/tactic_debug.mli
index b4c05324..e4c0a23e 100644
--- a/proofs/tactic_debug.mli
+++ b/proofs/tactic_debug.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,22 +8,20 @@
open Environ
open Pattern
-open Evd
-open Proof_type
open Names
open Tacexpr
open Term
+open Evd
(** This module intends to be a beginning of debugger for tactic expressions.
Currently, it is quite simple and we can hope to have, in the future, a more
complete panel of commands dedicated to a proof assistant framework *)
-val set_tactic_printer : (glob_tactic_expr ->Pp.std_ppcmds) -> unit
-val set_match_pattern_printer :
- (env -> constr_pattern match_pattern -> Pp.std_ppcmds) -> unit
-val set_match_rule_printer :
- ((Genarg.glob_constr_and_expr * constr_pattern,glob_tactic_expr) match_rule -> Pp.std_ppcmds) ->
- unit
+val tactic_printer : (glob_tactic_expr -> Pp.std_ppcmds) Hook.t
+val match_pattern_printer :
+ (env -> evar_map -> constr_pattern match_pattern -> Pp.std_ppcmds) Hook.t
+val match_rule_printer :
+ ((Tacexpr.glob_constr_and_expr * constr_pattern,glob_tactic_expr) match_rule -> Pp.std_ppcmds) Hook.t
(** Debug information *)
type debug_info =
@@ -32,37 +30,37 @@ type debug_info =
(** Prints the state and waits *)
val debug_prompt :
- int -> goal sigma -> glob_tactic_expr -> (debug_info -> 'a) -> 'a
+ int -> glob_tactic_expr -> (debug_info -> 'a Proofview.tactic) -> 'a Proofview.tactic
(** Initializes debugger *)
-val db_initialize : unit -> unit
+val db_initialize : unit Proofview.NonLogical.t
(** Prints a constr *)
-val db_constr : debug_info -> env -> constr -> unit
+val db_constr : debug_info -> env -> constr -> unit Proofview.NonLogical.t
(** Prints the pattern rule *)
val db_pattern_rule :
- debug_info -> int -> (Genarg.glob_constr_and_expr * constr_pattern,glob_tactic_expr) match_rule -> unit
+ debug_info -> int -> (Tacexpr.glob_constr_and_expr * constr_pattern,glob_tactic_expr) match_rule -> unit Proofview.NonLogical.t
(** Prints a matched hypothesis *)
val db_matched_hyp :
- debug_info -> env -> identifier * constr option * constr -> name -> unit
+ debug_info -> env -> Id.t * constr option * constr -> Name.t -> unit Proofview.NonLogical.t
(** Prints the matched conclusion *)
-val db_matched_concl : debug_info -> env -> constr -> unit
+val db_matched_concl : debug_info -> env -> constr -> unit Proofview.NonLogical.t
(** Prints a success message when the goal has been matched *)
-val db_mc_pattern_success : debug_info -> unit
+val db_mc_pattern_success : debug_info -> unit Proofview.NonLogical.t
(** Prints a failure message for an hypothesis pattern *)
val db_hyp_pattern_failure :
- debug_info -> env -> name * constr_pattern match_pattern -> unit
+ debug_info -> env -> evar_map -> Name.t * constr_pattern match_pattern -> unit Proofview.NonLogical.t
(** Prints a matching failure message for a rule *)
-val db_matching_failure : debug_info -> unit
+val db_matching_failure : debug_info -> unit Proofview.NonLogical.t
(** Prints an evaluation failure message for a rule *)
-val db_eval_failure : debug_info -> Pp.std_ppcmds -> unit
+val db_eval_failure : debug_info -> Pp.std_ppcmds -> unit Proofview.NonLogical.t
(** An exception handler *)
val explain_logic_error: (exn -> Pp.std_ppcmds) ref
@@ -74,8 +72,8 @@ val explain_logic_error: (exn -> Pp.std_ppcmds) ref
val explain_logic_error_no_anomaly : (exn -> Pp.std_ppcmds) ref
(** Prints a logic failure message for a rule *)
-val db_logic_failure : debug_info -> exn -> unit
+val db_logic_failure : debug_info -> exn -> unit Proofview.NonLogical.t
(** Prints a logic failure message for a rule *)
val db_breakpoint : debug_info ->
- identifier Util.located message_token list -> unit
+ Id.t Loc.located message_token list -> unit Proofview.NonLogical.t