aboutsummaryrefslogtreecommitdiffhomepage
path: root/tactics
diff options
context:
space:
mode:
authorGravatar aspiwack <aspiwack@85f007b7-540e-0410-9357-904b9bb8a0f7>2013-11-02 15:34:01 +0000
committerGravatar aspiwack <aspiwack@85f007b7-540e-0410-9357-904b9bb8a0f7>2013-11-02 15:34:01 +0000
commit260965dcf60d793ba01110ace8945cf51ef6531f (patch)
treed07323383e16bb5a63492e2721cf0502ba931716 /tactics
parent328279514e65f47a689e2d23f132c43c86870c05 (diff)
Makes the new Proofview.tactic the basic type of Ltac.
On the compilation of Coq, we can see an increase of ~20% compile time on my completely non-scientific tests. Hopefully this can be fixed. There are a lot of low hanging fruits, but this is an iso-functionality commit. With a few exceptions which were not necessary for the compilation of the theories: - The declarative mode is not yet ported - The timeout tactical is currently deactivated because it needs some subtle I/O. The framework is ready to handle it, but I haven't done it yet. - For much the same reason, the ltac debugger is unplugged. It will be more difficult, but will eventually be back. A few comments: I occasionnally used a coercion from [unit Proofview.tactic] to the old [Prooftype.tactic]. It should work smoothely, but loses any backtracking information: the coerced tactics has at most one success. - It is used in autorewrite (it shouldn't be a problem there). Autorewrite's code is fairly old and tricky - It is used in eauto, mostly for "Hint Extern". It may be an issue as time goes as we might want to have various success in a "Hint Extern". But it would require a heavy port of eauto.ml4 - It is used in typeclass eauto, but with a little help from Matthieu, it should be easy to port the whole thing to the new tactic engine, actually simplifying the code. - It is used in fourier. I believe it to be inocuous. - It is used in firstorder and congruence. I think it's ok. Their code is somewhat intricate and I'm not sure they would be easy to actually port. - It is used heavily in Function. And honestly, I have no idea whether it can do harm or not. Updates: (11 June 2013) Pierre-Marie Pédrot contributed the rebase over his new stream based architecture for Ltac matching (r16533), which avoid painfully and expensively working around the exception-throwing control flow of the previous API. (11 October 2013) Rebasing over recent commits (somewhere in r16721-r16730) rendered a major bug in my implementation of Tacticals.New.tclREPEAT_MAIN apparent. It caused Field_theory.v to loop. The bug made rewrite !lemma, rewrite ?lemma and autorewrite incorrect (tclREPEAT_MAIN was essentially tclREPEAT, causing rewrites to be tried in the side-conditions of conditional rewrites as well). The new implementation makes Coq faster, but it is pretty much impossible to tell if it is significant at all. git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@16967 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'tactics')
-rw-r--r--tactics/auto.ml108
-rw-r--r--tactics/auto.mli28
-rw-r--r--tactics/autorewrite.ml55
-rw-r--r--tactics/autorewrite.mli8
-rw-r--r--tactics/class_tactics.ml33
-rw-r--r--tactics/class_tactics.mli6
-rw-r--r--tactics/contradiction.ml76
-rw-r--r--tactics/contradiction.mli4
-rw-r--r--tactics/eauto.ml438
-rw-r--r--tactics/elim.ml117
-rw-r--r--tactics/elim.mli24
-rw-r--r--tactics/eqdecide.ml4144
-rw-r--r--tactics/equality.ml432
-rw-r--r--tactics/equality.mli68
-rw-r--r--tactics/evar_tactics.ml11
-rw-r--r--tactics/evar_tactics.mli2
-rw-r--r--tactics/extratactics.ml4149
-rw-r--r--tactics/extratactics.mli8
-rw-r--r--tactics/g_class.ml48
-rw-r--r--tactics/g_rewrite.ml439
-rw-r--r--tactics/hiddentac.mli54
-rw-r--r--tactics/inv.ml185
-rw-r--r--tactics/inv.mli22
-rw-r--r--tactics/leminv.ml26
-rw-r--r--tactics/leminv.mli10
-rw-r--r--tactics/refine.ml105
-rw-r--r--tactics/refine.mli2
-rw-r--r--tactics/rewrite.ml131
-rw-r--r--tactics/rewrite.mli8
-rw-r--r--tactics/tacinterp.ml1363
-rw-r--r--tactics/tacinterp.mli22
-rw-r--r--tactics/tactic_option.mli2
-rw-r--r--tactics/tacticals.ml256
-rw-r--r--tactics/tacticals.mli83
-rw-r--r--tactics/tactics.ml911
-rw-r--r--tactics/tactics.mli167
-rw-r--r--tactics/tauto.ml440
37 files changed, 2817 insertions, 1928 deletions
diff --git a/tactics/auto.ml b/tactics/auto.ml
index cf8707a46..548fcf00b 100644
--- a/tactics/auto.ml
+++ b/tactics/auto.ml
@@ -40,6 +40,7 @@ open Tacexpr
open Mod_subst
open Misctypes
open Locus
+open Proofview.Notations
(****************************************************************************)
(* The Type of Constructions Autotactic Hints *)
@@ -1104,14 +1105,16 @@ si après Intros la conclusion matche le pattern.
let (forward_interp_tactic, extern_interp) = Hook.make ()
-let conclPattern concl pat tac gl =
+(* arnaud: potentiel bug avec ce try/with *)
+let conclPattern concl pat tac =
let constr_bindings =
match pat with
- | None -> Id.Map.empty
+ | None -> Proofview.tclUNIT Id.Map.empty
| Some pat ->
- try matches pat concl
- with PatternMatchingFailure -> error "conclPattern" in
- Hook.get forward_interp_tactic constr_bindings tac gl
+ try Proofview.tclUNIT (matches pat concl)
+ with PatternMatchingFailure -> Proofview.tclZERO (UserError ("conclPattern",str"conclPattern")) in
+ constr_bindings >= fun constr_bindings ->
+ Hook.get forward_interp_tactic constr_bindings tac
(***********************************************************)
(** A debugging / verbosity framework for trivial and auto *)
@@ -1192,6 +1195,9 @@ let tclLOG (dbg,depth,trace) pp tac =
raise reraise
end
+(* arnaud: todo replug debug auto *)
+let new_tclLOG (dbg,depth,trace) pp tac = tac
+
(** For info, from the linear trace information, we reconstitute the part
of the proof tree we're interested in. The last executed tactic
comes first in the trace (and it should be a successful one).
@@ -1238,6 +1244,9 @@ let tclTRY_dbg d tac =
if level == Info then msg_debug (pr_info_nop d);
tclIDTAC gl)
+(* arnaud: todo: rebrancher debug auto *)
+let new_tclTRY_dbg tac = Tacticals.New.tclTRY
+
(**************************************************************************)
(* The Trivial tactic *)
(**************************************************************************)
@@ -1261,20 +1270,23 @@ let exists_evaluable_reference env = function
| EvalConstRef _ -> true
| EvalVarRef v -> try ignore(lookup_named v env); true with Not_found -> false
-let dbg_intro dbg = tclLOG dbg (fun () -> str "intro") intro
+let dbg_intro dbg = new_tclLOG dbg (fun () -> str "intro") intro
let dbg_assumption dbg = tclLOG dbg (fun () -> str "assumption") assumption
-let rec trivial_fail_db dbg mod_delta db_list local_db gl =
+let rec trivial_fail_db dbg mod_delta db_list local_db =
let intro_tac =
- tclTHEN (dbg_intro dbg)
- (fun g'->
- let hintl = make_resolve_hyp (pf_env g') (project g') (pf_last_hyp g')
- in trivial_fail_db dbg mod_delta db_list (Hint_db.add_list hintl local_db) g')
+ Tacticals.New.tclTHEN (dbg_intro dbg)
+ ( Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
+ Tacmach.New.pf_last_hyp >>- fun hyp ->
+ let hintl = make_resolve_hyp env sigma hyp
+ in trivial_fail_db dbg mod_delta db_list (Hint_db.add_list hintl local_db))
in
- tclFIRST
- ((dbg_assumption dbg)::intro_tac::
- (List.map tclCOMPLETE
- (trivial_resolve dbg mod_delta db_list local_db (pf_concl gl)))) gl
+ Goal.concl >>- fun concl ->
+ Tacticals.New.tclFIRST
+ ((Proofview.V82.tactic (dbg_assumption dbg))::intro_tac::
+ (List.map Tacticals.New.tclCOMPLETE
+ (trivial_resolve dbg mod_delta db_list local_db concl)))
and my_find_search_nodelta db_list local_db hdc concl =
List.map (fun hint -> (None,hint))
@@ -1318,23 +1330,23 @@ and my_find_search_delta db_list local_db hdc concl =
and tac_of_hint dbg db_list local_db concl (flags, ({pat=p; code=t})) =
let tactic =
match t with
- | Res_pf (c,cl) -> unify_resolve_gen flags (c,cl)
- | ERes_pf _ -> (fun gl -> error "eres_pf")
- | Give_exact c -> exact_check c
+ | Res_pf (c,cl) -> Proofview.V82.tactic (unify_resolve_gen flags (c,cl))
+ | ERes_pf _ -> Proofview.V82.tactic (fun gl -> error "eres_pf")
+ | Give_exact c -> Proofview.V82.tactic (exact_check c)
| Res_pf_THEN_trivial_fail (c,cl) ->
- tclTHEN
- (unify_resolve_gen flags (c,cl))
+ Tacticals.New.tclTHEN
+ (Proofview.V82.tactic (unify_resolve_gen flags (c,cl)))
(* With "(debug) trivial", we shouldn't end here, and
with "debug auto" we don't display the details of inner trivial *)
(trivial_fail_db (no_dbg ()) (not (Option.is_empty flags)) db_list local_db)
| Unfold_nth c ->
- (fun gl ->
+ Proofview.V82.tactic (fun gl ->
if exists_evaluable_reference (pf_env gl) c then
tclPROGRESS (h_reduce (Unfold [AllOccurrences,c]) Locusops.onConcl) gl
else tclFAIL 0 (str"Unbound reference") gl)
| Extern tacast -> conclPattern concl p tacast
in
- tclLOG dbg (fun () -> pr_autotactic t) tactic
+ new_tclLOG dbg (fun () -> pr_autotactic t) tactic
and trivial_resolve dbg mod_delta db_list local_db cl =
try
@@ -1360,19 +1372,21 @@ let make_db_list dbnames =
in
List.map lookup dbnames
-let trivial ?(debug=Off) lems dbnames gl =
+let trivial ?(debug=Off) lems dbnames =
let db_list = make_db_list dbnames in
let d = mk_trivial_dbg debug in
- tclTRY_dbg d
- (trivial_fail_db d false db_list (make_local_hint_db false lems gl)) gl
+ Tacmach.New.of_old (make_local_hint_db false lems) >>- fun hints ->
+ new_tclTRY_dbg d
+ (trivial_fail_db d false db_list hints)
-let full_trivial ?(debug=Off) lems gl =
+let full_trivial ?(debug=Off) lems =
let dbs = !searchtable in
let dbs = String.Map.remove "v62" dbs in
let db_list = List.map snd (String.Map.bindings dbs) in
let d = mk_trivial_dbg debug in
- tclTRY_dbg d
- (trivial_fail_db d false db_list (make_local_hint_db false lems gl)) gl
+ Tacmach.New.of_old (make_local_hint_db false lems) >>- fun hints ->
+ new_tclTRY_dbg d
+ (trivial_fail_db d false db_list hints)
let gen_trivial ?(debug=Off) lems = function
| None -> full_trivial ~debug lems
@@ -1396,7 +1410,7 @@ let possible_resolve dbg mod_delta db_list local_db cl =
with Not_found -> []
let dbg_case dbg id =
- tclLOG dbg (fun () -> str "case " ++ pr_id id) (simplest_case (mkVar id))
+ new_tclLOG dbg (fun () -> str "case " ++ pr_id id) (simplest_case (mkVar id))
let extend_local_db gl decl db =
Hint_db.add_list (make_resolve_hyp (pf_env gl) (project gl) decl) db
@@ -1405,33 +1419,36 @@ let extend_local_db gl decl db =
with the hint db extended with the so-obtained hypothesis *)
let intro_register dbg kont db =
- tclTHEN (dbg_intro dbg)
- (onLastDecl (fun decl gl -> kont (extend_local_db gl decl db) gl))
+ Tacticals.New.tclTHEN (dbg_intro dbg)
+ (Tacmach.New.of_old extend_local_db >>- fun extend_local_db ->
+ Tacticals.New.onLastDecl (fun decl -> kont (extend_local_db decl db)))
(* n is the max depth of search *)
(* local_db contains the local Hypotheses *)
let search d n mod_delta db_list local_db =
let rec search d n local_db =
- if Int.equal n 0 then (fun gl -> error "BOUND 2") else
- tclORELSE0 (dbg_assumption d)
- (tclORELSE0 (intro_register d (search d n) local_db)
- (fun gl ->
+ if Int.equal n 0 then Proofview.tclZERO (Errors.UserError ("",str"BOUND 2")) else
+ Tacticals.New.tclORELSE0 (Proofview.V82.tactic (dbg_assumption d))
+ (Tacticals.New.tclORELSE0 (intro_register d (search d n) local_db)
+ ( Goal.concl >>- fun concl ->
let d' = incr_dbg d in
- tclFIRST
+ Tacticals.New.tclFIRST
(List.map
- (fun ntac -> tclTHEN ntac (search d' (n-1) local_db))
- (possible_resolve d mod_delta db_list local_db (pf_concl gl))) gl))
+ (fun ntac -> Tacticals.New.tclTHEN ntac (search d' (n-1) local_db))
+ (possible_resolve d mod_delta db_list local_db concl))))
in
search d n local_db
let default_search_depth = ref 5
-let delta_auto ?(debug=Off) mod_delta n lems dbnames gl =
+let delta_auto ?(debug=Off) mod_delta n lems dbnames =
+ Proofview.tclUNIT () >= fun () -> (* delay for the side effects *)
let db_list = make_db_list dbnames in
let d = mk_auto_dbg debug in
- tclTRY_dbg d
- (search d n mod_delta db_list (make_local_hint_db false lems gl)) gl
+ Tacmach.New.of_old (make_local_hint_db false lems) >>- fun hints ->
+ new_tclTRY_dbg d
+ (search d n mod_delta db_list hints)
let auto ?(debug=Off) n = delta_auto ~debug false n
@@ -1439,18 +1456,19 @@ let new_auto ?(debug=Off) n = delta_auto ~debug true n
let default_auto = auto !default_search_depth [] []
-let delta_full_auto ?(debug=Off) mod_delta n lems gl =
+let delta_full_auto ?(debug=Off) mod_delta n lems =
let dbs = !searchtable in
let dbs = String.Map.remove "v62" dbs in
let db_list = List.map snd (String.Map.bindings dbs) in
let d = mk_auto_dbg debug in
- tclTRY_dbg d
- (search d n mod_delta db_list (make_local_hint_db false lems gl)) gl
+ Tacmach.New.of_old (make_local_hint_db false lems) >>- fun hints ->
+ new_tclTRY_dbg d
+ (search d n mod_delta db_list hints)
let full_auto ?(debug=Off) n = delta_full_auto ~debug false n
let new_full_auto ?(debug=Off) n = delta_full_auto ~debug true n
-let default_full_auto gl = full_auto !default_search_depth [] gl
+let default_full_auto = full_auto !default_search_depth []
let gen_auto ?(debug=Off) n lems dbnames =
let n = match n with None -> !default_search_depth | Some n -> n in
diff --git a/tactics/auto.mli b/tactics/auto.mli
index b8860f097..7ce351f43 100644
--- a/tactics/auto.mli
+++ b/tactics/auto.mli
@@ -175,7 +175,7 @@ val make_extern :
-> hint_entry
val extern_interp :
- (patvar_map -> Tacexpr.glob_tactic_expr -> tactic) Hook.t
+ (patvar_map -> Tacexpr.glob_tactic_expr -> unit Proofview.tactic) Hook.t
val extern_intern_tac :
(patvar list -> Tacexpr.raw_tactic_expr -> Tacexpr.glob_tactic_expr) Hook.t
@@ -205,7 +205,7 @@ val unify_resolve : Unification.unify_flags -> (constr * clausenv) -> tactic
[Pattern.somatches], then replace [?1] [?2] metavars in tacast by the
right values to build a tactic *)
-val conclPattern : constr -> constr_pattern option -> Tacexpr.glob_tactic_expr -> tactic
+val conclPattern : constr -> constr_pattern option -> Tacexpr.glob_tactic_expr -> unit Proofview.tactic
(** The Auto tactic *)
@@ -215,47 +215,47 @@ val conclPattern : constr -> constr_pattern option -> Tacexpr.glob_tactic_expr -
val make_db_list : hint_db_name list -> hint_db list
val auto : ?debug:Tacexpr.debug ->
- int -> open_constr list -> hint_db_name list -> tactic
+ int -> open_constr list -> hint_db_name list -> unit Proofview.tactic
(** Auto with more delta. *)
val new_auto : ?debug:Tacexpr.debug ->
- int -> open_constr list -> hint_db_name list -> tactic
+ int -> open_constr list -> hint_db_name list -> unit Proofview.tactic
(** auto with default search depth and with the hint database "core" *)
-val default_auto : tactic
+val default_auto : unit Proofview.tactic
(** auto with all hint databases except the "v62" compatibility database *)
val full_auto : ?debug:Tacexpr.debug ->
- int -> open_constr list -> tactic
+ int -> open_constr list -> unit Proofview.tactic
(** auto with all hint databases except the "v62" compatibility database
and doing delta *)
val new_full_auto : ?debug:Tacexpr.debug ->
- int -> open_constr list -> tactic
+ int -> open_constr list -> unit Proofview.tactic
(** auto with default search depth and with all hint databases
except the "v62" compatibility database *)
-val default_full_auto : tactic
+val default_full_auto : unit Proofview.tactic
(** The generic form of auto (second arg [None] means all bases) *)
val gen_auto : ?debug:Tacexpr.debug ->
- int option -> open_constr list -> hint_db_name list option -> tactic
+ int option -> open_constr list -> hint_db_name list option -> unit Proofview.tactic
(** The hidden version of auto *)
val h_auto : ?debug:Tacexpr.debug ->
- int option -> open_constr list -> hint_db_name list option -> tactic
+ int option -> open_constr list -> hint_db_name list option -> unit Proofview.tactic
(** Trivial *)
val trivial : ?debug:Tacexpr.debug ->
- open_constr list -> hint_db_name list -> tactic
+ open_constr list -> hint_db_name list -> unit Proofview.tactic
val gen_trivial : ?debug:Tacexpr.debug ->
- open_constr list -> hint_db_name list option -> tactic
+ open_constr list -> hint_db_name list option -> unit Proofview.tactic
val full_trivial : ?debug:Tacexpr.debug ->
- open_constr list -> tactic
+ open_constr list -> unit Proofview.tactic
val h_trivial : ?debug:Tacexpr.debug ->
- open_constr list -> hint_db_name list option -> tactic
+ open_constr list -> hint_db_name list option -> unit Proofview.tactic
val pr_autotactic : 'a auto_tactic -> Pp.std_ppcmds
diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml
index c65c75f63..329198ccc 100644
--- a/tactics/autorewrite.ml
+++ b/tactics/autorewrite.ml
@@ -19,6 +19,7 @@ open Util
open Tacexpr
open Mod_subst
open Locus
+open Proofview.Notations
(* Rewriting rules *)
type rew_rule = { rew_lemma: constr;
@@ -92,29 +93,28 @@ type raw_rew_rule = Loc.t * constr * bool * raw_tactic_expr option
let one_base general_rewrite_maybe_in tac_main bas =
let lrul = find_rewrites bas in
let lrul = List.map (fun h ->
- let tac = match h.rew_tac with None -> tclIDTAC | Some t -> Tacinterp.eval_tactic t in
+ let tac = match h.rew_tac with None -> Proofview.tclUNIT () | Some t -> Tacinterp.eval_tactic t in
(h.rew_lemma,h.rew_l2r,tac)) lrul in
- tclREPEAT_MAIN (tclPROGRESS (List.fold_left (fun tac (csr,dir,tc) ->
- tclTHEN tac
- (tclREPEAT_MAIN
- (tclTHENFIRST (general_rewrite_maybe_in dir csr tc) tac_main)))
- tclIDTAC lrul))
+ Tacticals.New.tclREPEAT_MAIN (Proofview.tclPROGRESS (List.fold_left (fun tac (csr,dir,tc) ->
+ Tacticals.New.tclTHEN tac
+ (Tacticals.New.tclREPEAT_MAIN
+ (Tacticals.New.tclTHENFIRST (general_rewrite_maybe_in dir csr tc) tac_main)))
+ (Proofview.tclUNIT()) lrul))
(* The AutoRewrite tactic *)
let autorewrite ?(conds=Naive) tac_main lbas =
- tclREPEAT_MAIN (tclPROGRESS
+ Tacticals.New.tclREPEAT_MAIN (Proofview.tclPROGRESS
(List.fold_left (fun tac bas ->
- tclTHEN tac
+ Tacticals.New.tclTHEN tac
(one_base (fun dir c tac ->
let tac = (tac, conds) in
general_rewrite dir AllOccurrences true false ~tac c)
tac_main bas))
- tclIDTAC lbas))
+ (Proofview.tclUNIT()) lbas))
-let autorewrite_multi_in ?(conds=Naive) idl tac_main lbas : tactic =
- fun gl ->
+let autorewrite_multi_in ?(conds=Naive) idl tac_main lbas =
(* let's check at once if id exists (to raise the appropriate error) *)
- let _ = List.map (Tacmach.pf_get_hyp gl) idl in
+ Goal.sensitive_list_map Tacmach.New.pf_get_hyp idl >>- fun _ ->
let general_rewrite_in id =
let id = ref id in
let to_be_cleared = ref false in
@@ -125,7 +125,7 @@ let autorewrite_multi_in ?(conds=Naive) idl tac_main lbas : tactic =
| _ -> (* even the hypothesis id is missing *)
error ("No such hypothesis: " ^ (Id.to_string !id) ^".")
in
- let gl' = general_rewrite_in dir AllOccurrences true ~tac:(tac, conds) false !id cstr false gl in
+ let gl' = Proofview.V82.of_tactic (general_rewrite_in dir AllOccurrences true ~tac:(tac, conds) false !id cstr false) gl in
let gls = gl'.Evd.it in
match gls with
g::_ ->
@@ -149,11 +149,12 @@ let autorewrite_multi_in ?(conds=Naive) idl tac_main lbas : tactic =
| _ -> assert false) (* there must be at least an hypothesis *)
| _ -> assert false (* rewriting cannot complete a proof *)
in
- tclMAP (fun id ->
- tclREPEAT_MAIN (tclPROGRESS
+ let general_rewrite_in x y z w = Proofview.V82.tactic (general_rewrite_in x y z w) in
+ Tacticals.New.tclMAP (fun id ->
+ Tacticals.New.tclREPEAT_MAIN (Proofview.tclPROGRESS
(List.fold_left (fun tac bas ->
- tclTHEN tac (one_base (general_rewrite_in id) tac_main bas)) tclIDTAC lbas)))
- idl gl
+ Tacticals.New.tclTHEN tac (one_base (general_rewrite_in id) tac_main bas)) (Proofview.tclUNIT()) lbas)))
+ idl
let autorewrite_in ?(conds=Naive) id = autorewrite_multi_in ~conds [id]
@@ -164,37 +165,35 @@ let gen_auto_multi_rewrite conds tac_main lbas cl =
if cl.concl_occs != AllOccurrences &&
cl.concl_occs != NoOccurrences
then
- error "The \"at\" syntax isn't available yet for the autorewrite tactic."
+ Proofview.tclZERO (UserError("" , str"The \"at\" syntax isn't available yet for the autorewrite tactic."))
else
let compose_tac t1 t2 =
match cl.onhyps with
| Some [] -> t1
- | _ -> tclTHENFIRST t1 t2
+ | _ -> Tacticals.New.tclTHENFIRST t1 t2
in
compose_tac
- (if cl.concl_occs != NoOccurrences then autorewrite ~conds tac_main lbas else tclIDTAC)
+ (if cl.concl_occs != NoOccurrences then autorewrite ~conds tac_main lbas else Proofview.tclUNIT ())
(match cl.onhyps with
| Some l -> try_do_hyps (fun ((_,id),_) -> id) l
| None ->
- fun gl ->
(* try to rewrite in all hypothesis
(except maybe the rewritten one) *)
- let ids = Tacmach.pf_ids_of_hyps gl
- in try_do_hyps (fun id -> id) ids gl)
+ Tacmach.New.pf_ids_of_hyps >>- fun ids ->
+ try_do_hyps (fun id -> id) ids)
-let auto_multi_rewrite ?(conds=Naive) = gen_auto_multi_rewrite conds Refiner.tclIDTAC
+let auto_multi_rewrite ?(conds=Naive) = gen_auto_multi_rewrite conds (Proofview.tclUNIT())
-let auto_multi_rewrite_with ?(conds=Naive) tac_main lbas cl gl =
+let auto_multi_rewrite_with ?(conds=Naive) tac_main lbas cl =
let onconcl = match cl.Locus.concl_occs with NoOccurrences -> false | _ -> true in
match onconcl,cl.Locus.onhyps with
| false,Some [_] | true,Some [] | false,Some [] ->
(* autorewrite with .... in clause using tac n'est sur que
si clause represente soit le but soit UNE hypothese
*)
- gen_auto_multi_rewrite conds tac_main lbas cl gl
+ gen_auto_multi_rewrite conds tac_main lbas cl
| _ ->
- Errors.errorlabstrm "autorewrite"
- (strbrk "autorewrite .. in .. using can only be used either with a unique hypothesis or on the conclusion.")
+ Proofview.tclZERO (UserError ("autorewrite",strbrk "autorewrite .. in .. using can only be used either with a unique hypothesis or on the conclusion."))
(* Functions necessary to the library object declaration *)
let cache_hintrewrite (_,(rbase,lrl)) =
diff --git a/tactics/autorewrite.mli b/tactics/autorewrite.mli
index 9ddabc584..69ab9c55f 100644
--- a/tactics/autorewrite.mli
+++ b/tactics/autorewrite.mli
@@ -21,8 +21,8 @@ val add_rew_rules : string -> raw_rew_rule list -> unit
The optional conditions tell rewrite how to handle matching and side-condition solving.
Default is Naive: first match in the clause, don't look at the side-conditions to
tell if the rewrite succeeded. *)
-val autorewrite : ?conds:conditions -> tactic -> string list -> tactic
-val autorewrite_in : ?conds:conditions -> Names.Id.t -> tactic -> string list -> tactic
+val autorewrite : ?conds:conditions -> unit Proofview.tactic -> string list -> unit Proofview.tactic
+val autorewrite_in : ?conds:conditions -> Names.Id.t -> unit Proofview.tactic -> string list -> unit Proofview.tactic
(** Rewriting rules *)
type rew_rule = { rew_lemma: constr;
@@ -35,9 +35,9 @@ val find_rewrites : string -> rew_rule list
val find_matches : string -> constr -> rew_rule list
-val auto_multi_rewrite : ?conds:conditions -> string list -> Locus.clause -> tactic
+val auto_multi_rewrite : ?conds:conditions -> string list -> Locus.clause -> unit Proofview.tactic
-val auto_multi_rewrite_with : ?conds:conditions -> tactic -> string list -> Locus.clause -> tactic
+val auto_multi_rewrite_with : ?conds:conditions -> unit Proofview.tactic -> string list -> Locus.clause -> unit Proofview.tactic
val print_rewrite_hintdb : string -> Pp.std_ppcmds
diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml
index 87086cfae..71f8c2ba8 100644
--- a/tactics/class_tactics.ml
+++ b/tactics/class_tactics.ml
@@ -24,6 +24,7 @@ open Globnames
open Evd
open Locus
open Misctypes
+open Proofview.Notations
(** Hint database named "typeclass_instances", now created directly in Auto *)
@@ -81,14 +82,14 @@ let rec eq_constr_mod_evars x y =
| Evar (e1, l1), Evar (e2, l2) when not (Evar.equal e1 e2) -> true
| _, _ -> compare_constr eq_constr_mod_evars x y
-let progress_evars t gl =
- let concl = pf_concl gl in
- let check gl' =
- let newconcl = pf_concl gl' in
+let progress_evars t =
+ Goal.concl >>- fun concl ->
+ let check =
+ Goal.concl >>- fun newconcl ->
if eq_constr_mod_evars concl newconcl
- then tclFAIL 0 (str"No progress made (modulo evars)") gl'
- else tclIDTAC gl'
- in tclTHEN t check gl
+ then Tacticals.New.tclFAIL 0 (str"No progress made (modulo evars)")
+ else Proofview.tclUNIT ()
+ in t <*> check
let unify_e_resolve flags (c,clenv) gls =
let clenv' = connect_clenv gls clenv in
@@ -105,7 +106,7 @@ let clenv_of_prods nprods (c, clenv) gls =
else
let ty = pf_type_of gls c in
let diff = nb_prod ty - nprods in
- if diff >= 0 then
+ if Pervasives.(>=) diff 0 then
Some (mk_clenv_from_n gls (Some diff) (c,ty))
else None
@@ -125,7 +126,7 @@ let flags_of_state st =
let rec e_trivial_fail_db db_list local_db goal =
let tacl =
Eauto.registered_e_assumption ::
- (tclTHEN Tactics.intro
+ (tclTHEN (Proofview.V82.of_tactic Tactics.intro)
(function g'->
let d = pf_last_hyp g' in
let hintl = make_resolve_hyp (pf_env g') (project g') d in
@@ -164,7 +165,7 @@ and e_my_find_search db_list local_db hdc complete concl =
| Extern tacast ->
(* tclTHEN *)
(* (fun gl -> Refiner.tclEVARS (mark_unresolvables (project gl)) gl) *)
- (conclPattern concl p tacast)
+ Proofview.V82.of_tactic (conclPattern concl p tacast)
in
let tac = if complete then tclCOMPLETE tac else tac in
match t with
@@ -293,8 +294,8 @@ let lift_tactic tac (f : goal list sigma -> autoinfo -> autogoal list sigma) : '
| None -> fk () }
let intro_tac : atac =
- lift_tactic Tactics.intro
- (fun {it = gls; sigma = s;} info ->
+ lift_tactic (Proofview.V82.of_tactic Tactics.intro)
+ (fun {it = gls; sigma = s} info ->
let gls' =
List.map (fun g' ->
let env = Goal.V82.env s g' in
@@ -723,7 +724,9 @@ let typeclasses_eauto ?(only_classes=false) ?(st=full_transparent_state) dbs gl
eauto ?limit:!typeclasses_depth ~only_classes ~st dbs gl
with Not_found -> tclFAIL 0 (str" typeclasses eauto failed on: " ++ Printer.pr_goal gl) gl
-let _ = Classes.refine_ref := Refine.refine
+let _ = Classes.refine_ref := begin fun c ->
+ Refine.refine c
+end
(** Take the head of the arity of a constr.
Used in the partial application tactic. *)
@@ -741,8 +744,8 @@ let head_of_constr h c =
letin_tac None (Name h) c None Locusops.allHyps
let not_evar c = match kind_of_term c with
-| Evar _ -> tclFAIL 0 (str"Evar")
-| _ -> tclIDTAC
+| Evar _ -> Tacticals.New.tclFAIL 0 (str"Evar")
+| _ -> Proofview.tclUNIT ()
let is_ground c gl =
if Evarutil.is_ground_term (project gl) c then tclIDTAC gl
diff --git a/tactics/class_tactics.mli b/tactics/class_tactics.mli
index 0b74027c3..f64f708de 100644
--- a/tactics/class_tactics.mli
+++ b/tactics/class_tactics.mli
@@ -18,14 +18,14 @@ val get_typeclasses_debug : unit -> bool
val set_typeclasses_depth : int option -> unit
val get_typeclasses_depth : unit -> int option
-val progress_evars : tactic -> tactic
+val progress_evars : unit Proofview.tactic -> unit Proofview.tactic
val typeclasses_eauto : ?only_classes:bool -> ?st:transparent_state ->
Auto.hint_db_name list -> tactic
-val head_of_constr : Id.t -> Term.constr -> tactic
+val head_of_constr : Id.t -> Term.constr -> unit Proofview.tactic
-val not_evar : constr -> tactic
+val not_evar : constr -> unit Proofview.tactic
val is_ground : constr -> tactic
diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml
index fbd5a8cb0..565f30cfb 100644
--- a/tactics/contradiction.ml
+++ b/tactics/contradiction.ml
@@ -15,6 +15,7 @@ open Tactics
open Coqlib
open Reductionops
open Misctypes
+open Proofview.Notations
(* Absurd *)
@@ -27,63 +28,78 @@ let absurd c gls =
(tclTHEN (elim_type (build_coq_False ())) (cut c))
([(tclTHENS
(cut (mkApp(build_coq_not (),[|c|])))
- ([(tclTHEN intros
+ ([(tclTHEN (Proofview.V82.of_tactic intros)
((fun gl ->
let ida = pf_nth_hyp_id gl 1
and idna = pf_nth_hyp_id gl 2 in
exact_no_check (mkApp(mkVar idna,[|mkVar ida|])) gl)));
tclIDTAC]));
tclIDTAC])) { gls with Evd.sigma; }
+let absurd c = Proofview.V82.tactic (absurd c)
(* Contradiction *)
-let filter_hyp f tac gl =
+let filter_hyp f tac =
let rec seek = function
| [] -> raise Not_found
- | (id,_,t)::rest when f t -> tac id gl
+ | (id,_,t)::rest when f t -> tac id
| _::rest -> seek rest in
- seek (pf_hyps gl)
+ Goal.hyps >>- fun hyps ->
+ seek (Environ.named_context_of_val hyps)
-let contradiction_context gl =
- let env = pf_env gl in
- let sigma = project gl in
- let rec seek_neg l gl = match l with
- | [] -> error "No such contradiction"
+let contradiction_context =
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
+ let rec seek_neg l = match l with
+ | [] -> Proofview.tclZERO (UserError ("" , Pp.str"No such contradiction"))
| (id,_,typ)::rest ->
let typ = whd_betadeltaiota env sigma typ in
if is_empty_type typ then
- simplest_elim (mkVar id) gl
+ simplest_elim (mkVar id)
else match kind_of_term typ with
| Prod (na,t,u) when is_empty_type u ->
- (try
- filter_hyp (fun typ -> pf_conv_x_leq gl typ t)
- (fun id' -> simplest_elim (mkApp (mkVar id,[|mkVar id'|])))
- gl
- with Not_found -> seek_neg rest gl)
- | _ -> seek_neg rest gl in
- seek_neg (pf_hyps gl) gl
+ (Proofview.tclORELSE
+ begin
+ Tacmach.New.pf_apply is_conv_leq >>- fun is_conv_leq ->
+ filter_hyp (fun typ -> is_conv_leq typ t)
+ (fun id' -> simplest_elim (mkApp (mkVar id,[|mkVar id'|])))
+ end
+ begin function
+ | Not_found -> seek_neg rest
+ | e -> Proofview.tclZERO e
+ end)
+ | _ -> seek_neg rest in
+ Goal.hyps >>- fun hyps ->
+ let hyps = Environ.named_context_of_val hyps in
+ seek_neg hyps
let is_negation_of env sigma typ t =
match kind_of_term (whd_betadeltaiota env sigma t) with
| Prod (na,t,u) -> is_empty_type u && is_conv_leq env sigma typ t
| _ -> false
-let contradiction_term (c,lbind as cl) gl =
- let env = pf_env gl in
- let sigma = project gl in
- let typ = pf_type_of gl c in
+let contradiction_term (c,lbind as cl) =
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
+ Tacmach.New.pf_apply Typing.type_of >>- fun type_of ->
+ let typ = type_of c in
let _, ccl = splay_prod env sigma typ in
if is_empty_type ccl then
- tclTHEN (elim false cl None) (tclTRY assumption) gl
+ Tacticals.New.tclTHEN (elim false cl None) (Proofview.V82.tactic (tclTRY assumption))
else
- try
- if lbind = NoBindings then
- filter_hyp (is_negation_of env sigma typ)
- (fun id -> simplest_elim (mkApp (mkVar id,[|c|]))) gl
- else
- raise Not_found
- with Not_found -> error "Not a contradiction."
+ Proofview.tclORELSE
+ begin
+ if lbind = NoBindings then
+ filter_hyp (is_negation_of env sigma typ)
+ (fun id -> simplest_elim (mkApp (mkVar id,[|c|])))
+ else
+ Proofview.tclZERO Not_found
+ end
+ begin function
+ | Not_found -> Proofview.tclZERO (Errors.UserError ("",Pp.str"Not a contradiction."))
+ | e -> Proofview.tclZERO e
+ end
let contradiction = function
- | None -> tclTHEN intros contradiction_context
+ | None -> Tacticals.New.tclTHEN intros contradiction_context
| Some c -> contradiction_term c
diff --git a/tactics/contradiction.mli b/tactics/contradiction.mli
index 38d28b6b9..e3dabfe98 100644
--- a/tactics/contradiction.mli
+++ b/tactics/contradiction.mli
@@ -11,5 +11,5 @@ open Term
open Proof_type
open Misctypes
-val absurd : constr -> tactic
-val contradiction : constr with_bindings option -> tactic
+val absurd : constr -> unit Proofview.tactic
+val contradiction : constr with_bindings option -> unit Proofview.tactic
diff --git a/tactics/eauto.ml4 b/tactics/eauto.ml4
index 1756f89ce..369107f6c 100644
--- a/tactics/eauto.ml4
+++ b/tactics/eauto.ml4
@@ -42,11 +42,11 @@ let e_assumption gl =
tclFIRST (List.map assumption (pf_ids_of_hyps gl)) gl
TACTIC EXTEND eassumption
-| [ "eassumption" ] -> [ e_assumption ]
+| [ "eassumption" ] -> [ Proofview.V82.tactic e_assumption ]
END
TACTIC EXTEND eexact
-| [ "eexact" constr(c) ] -> [ e_give_exact c ]
+| [ "eexact" constr(c) ] -> [ Proofview.V82.tactic (e_give_exact c) ]
END
let registered_e_assumption gl =
@@ -58,7 +58,7 @@ let registered_e_assumption gl =
(************************************************************************)
let one_step l gl =
- [Tactics.intro]
+ [Proofview.V82.of_tactic Tactics.intro]
@ (List.map h_simplest_eapply (List.map mkVar (pf_ids_of_hyps gl)))
@ (List.map h_simplest_eapply l)
@ (List.map assumption (pf_ids_of_hyps gl))
@@ -80,7 +80,7 @@ let prolog_tac l n gl =
errorlabstrm "Prolog.prolog" (str "Prolog failed.")
TACTIC EXTEND prolog
-| [ "prolog" "[" open_constr_list(l) "]" int_or_var(n) ] -> [ prolog_tac l n ]
+| [ "prolog" "[" open_constr_list(l) "]" int_or_var(n) ] -> [ Proofview.V82.tactic (prolog_tac l n) ]
END
open Auto
@@ -100,7 +100,7 @@ let unify_e_resolve flags (c,clenv) gls =
let rec e_trivial_fail_db db_list local_db goal =
let tacl =
registered_e_assumption ::
- (tclTHEN Tactics.intro
+ (tclTHEN (Proofview.V82.of_tactic Tactics.intro)
(function g'->
let d = pf_last_hyp g' in
let hintl = make_resolve_hyp (pf_env g') (project g') d in
@@ -134,7 +134,7 @@ and e_my_find_search db_list local_db hdc concl =
tclTHEN (unify_e_resolve st (term,cl))
(e_trivial_fail_db db_list local_db)
| Unfold_nth c -> h_reduce (Unfold [AllOccurrences,c]) onConcl
- | Extern tacast -> conclPattern concl p tacast
+ | Extern tacast -> Proofview.V82.of_tactic (conclPattern concl p tacast)
in
(tac,lazy (pr_autotactic t)))
in
@@ -241,7 +241,7 @@ module SearchProblem = struct
{ depth = s.depth; tacres = res;
last_tactic = pp; dblist = s.dblist;
localdb = ldb :: List.tl s.localdb; prev = ps })
- (filter_tactics s.tacres [Tactics.intro,lazy (str "intro")])
+ (filter_tactics s.tacres [Proofview.V82.of_tactic Tactics.intro,lazy (str "intro")])
in
let rec_tacs =
let l =
@@ -412,7 +412,7 @@ END
TACTIC EXTEND eauto
| [ "eauto" int_or_var_opt(n) int_or_var_opt(p) auto_using(lems)
hintbases(db) ] ->
- [ gen_eauto (make_dimension n p) lems db ]
+ [ Proofview.V82.tactic (gen_eauto (make_dimension n p) lems db) ]
END
TACTIC EXTEND new_eauto
@@ -426,19 +426,19 @@ END
TACTIC EXTEND debug_eauto
| [ "debug" "eauto" int_or_var_opt(n) int_or_var_opt(p) auto_using(lems)
hintbases(db) ] ->
- [ gen_eauto ~debug:Debug (make_dimension n p) lems db ]
+ [ Proofview.V82.tactic (gen_eauto ~debug:Debug (make_dimension n p) lems db) ]
END
TACTIC EXTEND info_eauto
| [ "info_eauto" int_or_var_opt(n) int_or_var_opt(p) auto_using(lems)
hintbases(db) ] ->
- [ gen_eauto ~debug:Info (make_dimension n p) lems db ]
+ [ Proofview.V82.tactic (gen_eauto ~debug:Info (make_dimension n p) lems db) ]
END
TACTIC EXTEND dfs_eauto
| [ "dfs" "eauto" int_or_var_opt(p) auto_using(lems)
hintbases(db) ] ->
- [ gen_eauto (true, make_depth p) lems db ]
+ [ Proofview.V82.tactic (gen_eauto (true, make_depth p) lems db) ]
END
let cons a l = a :: l
@@ -472,7 +472,7 @@ let autounfold_tac db cls gl =
open Extraargs
TACTIC EXTEND autounfold
-| [ "autounfold" hintbases(db) in_arg_hyp(id) ] -> [ autounfold_tac db id ]
+| [ "autounfold" hintbases(db) in_arg_hyp(id) ] -> [ Proofview.V82.tactic (autounfold_tac db id) ]
END
let unfold_head env (ids, csts) c =
@@ -537,28 +537,30 @@ let autounfold_one db cl gl =
TACTIC EXTEND autounfold_one
| [ "autounfold_one" hintbases(db) "in" hyp(id) ] ->
- [ autounfold_one (match db with None -> ["core"] | Some x -> "core"::x) (Some (id, InHyp)) ]
+ [ Proofview.V82.tactic (autounfold_one (match db with None -> ["core"] | Some x -> "core"::x) (Some (id, InHyp))) ]
| [ "autounfold_one" hintbases(db) ] ->
- [ autounfold_one (match db with None -> ["core"] | Some x -> "core"::x) None ]
+ [ Proofview.V82.tactic (autounfold_one (match db with None -> ["core"] | Some x -> "core"::x) None) ]
END
TACTIC EXTEND autounfoldify
| [ "autounfoldify" constr(x) ] -> [
+ Proofview.V82.tactic (
let db = match kind_of_term x with
| Const c -> Label.to_string (con_label c)
| _ -> assert false
- in autounfold ["core";db] onConcl ]
+ in autounfold ["core";db] onConcl
+ )]
END
TACTIC EXTEND unify
-| ["unify" constr(x) constr(y) ] -> [ unify x y ]
+| ["unify" constr(x) constr(y) ] -> [ Proofview.V82.tactic (unify x y) ]
| ["unify" constr(x) constr(y) "with" preident(base) ] -> [
- unify ~state:(Hint_db.transparent_state (searchtable_map base)) x y ]
+ Proofview.V82.tactic (unify ~state:(Hint_db.transparent_state (searchtable_map base)) x y) ]
END
TACTIC EXTEND convert_concl_no_check
-| ["convert_concl_no_check" constr(x) ] -> [ convert_concl_no_check x DEFAULTcast ]
+| ["convert_concl_no_check" constr(x) ] -> [ Proofview.V82.tactic (convert_concl_no_check x DEFAULTcast) ]
END
diff --git a/tactics/elim.ml b/tactics/elim.ml
index faa32ab86..0e30daf6d 100644
--- a/tactics/elim.ml
+++ b/tactics/elim.ml
@@ -17,6 +17,7 @@ open Tacmach
open Tacticals
open Tactics
open Misctypes
+open Proofview.Notations
let introElimAssumsThen tac ba =
let nassums =
@@ -24,8 +25,8 @@ let introElimAssumsThen tac ba =
(fun acc b -> if b then acc+2 else acc+1)
0 ba.branchsign
in
- let introElimAssums = tclDO nassums intro in
- (tclTHEN introElimAssums (elim_on_ba tac ba))
+ let introElimAssums = Tacticals.New.tclDO nassums intro in
+ (Tacticals.New.tclTHEN introElimAssums (Tacticals.New.elim_on_ba tac ba))
let introCaseAssumsThen tac ba =
let case_thin_sign =
@@ -41,8 +42,8 @@ let introCaseAssumsThen tac ba =
(ba.branchnames, []),
if n1 > n2 then snd (List.chop n2 case_thin_sign) else [] in
let introCaseAssums =
- tclTHEN (intros_pattern MoveLast l1) (intros_clearing l3) in
- (tclTHEN introCaseAssums (case_on_ba (tac l2) ba))
+ Tacticals.New.tclTHEN (intros_pattern MoveLast l1) (intros_clearing l3) in
+ (Tacticals.New.tclTHEN introCaseAssums (Tacticals.New.case_on_ba (tac l2) ba))
(* The following tactic Decompose repeatedly applies the
elimination(s) rule(s) of the types satisfying the predicate
@@ -62,18 +63,18 @@ Another example :
Qed.
*)
-let elimHypThen tac id gl =
- elimination_then tac ([],[]) (mkVar id) gl
+let elimHypThen tac id =
+ Tacticals.New.elimination_then tac ([],[]) (mkVar id)
let rec general_decompose_on_hyp recognizer =
- ifOnHyp recognizer (general_decompose_aux recognizer) (fun _ -> tclIDTAC)
+ Tacticals.New.ifOnHyp recognizer (general_decompose_aux recognizer) (fun _ -> Proofview.tclUNIT())
and general_decompose_aux recognizer id =
elimHypThen
(introElimAssumsThen
(fun bas ->
- tclTHEN (clear [id])
- (tclMAP (general_decompose_on_hyp recognizer)
+ Tacticals.New.tclTHEN (Proofview.V82.tactic (clear [id]))
+ (Tacticals.New.tclMAP (general_decompose_on_hyp recognizer)
(ids_of_named_context bas.assums))))
id
@@ -84,42 +85,49 @@ and general_decompose_aux recognizer id =
let tmphyp_name = Id.of_string "_TmpHyp"
let up_to_delta = ref false (* true *)
-let general_decompose recognizer c gl =
- let typc = pf_type_of gl c in
- tclTHENSV (cut typc)
- [| tclTHEN (intro_using tmphyp_name)
- (onLastHypId
- (ifOnHyp recognizer (general_decompose_aux recognizer)
- (fun id -> clear [id])));
- exact_no_check c |] gl
-
-let head_in gls indl t =
- try
- let ity,_ =
- if !up_to_delta
- then find_mrectype (pf_env gls) (project gls) t
- else extract_mrectype t
- in List.mem ity indl
- with Not_found -> false
-
-let decompose_these c l gls =
+let general_decompose recognizer c =
+ Tacmach.New.pf_apply Typing.type_of >>- fun type_of ->
+ let typc = type_of c in
+ Tacticals.New.tclTHENS (Proofview.V82.tactic (cut typc))
+ [ Tacticals.New.tclTHEN (intro_using tmphyp_name)
+ (Tacticals.New.onLastHypId
+ (Tacticals.New.ifOnHyp recognizer (general_decompose_aux recognizer)
+ (fun id -> Proofview.V82.tactic (clear [id]))));
+ Proofview.V82.tactic (exact_no_check c) ]
+
+let head_in =
+ Goal.env >- fun env ->
+ Goal.defs >- fun sigma ->
+ Goal.return begin
+ fun indl t ->
+ try
+ let ity,_ =
+ if !up_to_delta
+ then find_mrectype env sigma t
+ else extract_mrectype t
+ in List.mem ity indl
+ with Not_found -> false
+ end
+
+let decompose_these c l =
let indl = (*List.map inductive_of*) l in
- general_decompose (fun (_,t) -> head_in gls indl t) c gls
+ head_in >>- fun head_in ->
+ general_decompose (fun (_,t) -> head_in indl t) c
-let decompose_nonrec c gls =
+let decompose_nonrec c =
general_decompose
(fun (_,t) -> is_non_recursive_type t)
- c gls
+ c
-let decompose_and c gls =
+let decompose_and c =
general_decompose
(fun (_,t) -> is_record t)
- c gls
+ c
-let decompose_or c gls =
+let decompose_or c =
general_decompose
(fun (_,t) -> is_disjunction t)
- c gls
+ c
let h_decompose l c = decompose_these c l
@@ -133,10 +141,11 @@ let simple_elimination c gls =
simple_elimination_then (fun _ -> tclIDTAC) c gls
let induction_trailer abs_i abs_j bargs =
- tclTHEN
- (tclDO (abs_j - abs_i) intro)
- (onLastHypId
- (fun id gls ->
+ Tacticals.New.tclTHEN
+ (Tacticals.New.tclDO (abs_j - abs_i) intro)
+ (Tacticals.New.onLastHypId
+ (fun id ->
+ Proofview.V82.tactic begin fun gls ->
let idty = pf_type_of gls (mkVar id) in
let fvty = global_vars (pf_env gls) idty in
let possible_bring_hyps =
@@ -153,22 +162,24 @@ let induction_trailer abs_i abs_j bargs =
let ids = List.rev (ids_of_named_context hyps) in
(tclTHENSEQ
[bring_hyps hyps; tclTRY (clear ids);
- simple_elimination (mkVar id)])
- gls))
-
-let double_ind h1 h2 gls =
- let abs_i = depth_of_quantified_hypothesis true h1 gls in
- let abs_j = depth_of_quantified_hypothesis true h2 gls in
- let (abs_i,abs_j) =
- if abs_i < abs_j then (abs_i,abs_j) else
- if abs_i > abs_j then (abs_j,abs_i) else
- error "Both hypotheses are the same." in
- (tclTHEN (tclDO abs_i intro)
- (onLastHypId
+ simple_elimination (mkVar id)]) gls
+ end
+ ))
+
+let double_ind h1 h2 =
+ Tacmach.New.of_old (depth_of_quantified_hypothesis true h1) >>- fun abs_i ->
+ Tacmach.New.of_old (depth_of_quantified_hypothesis true h2) >>- fun abs_j ->
+ let abs =
+ if abs_i < abs_j then Proofview.tclUNIT (abs_i,abs_j) else
+ if abs_i > abs_j then Proofview.tclUNIT (abs_j,abs_i) else
+ Proofview.tclZERO (Errors.UserError ("", Pp.str"Both hypotheses are the same.")) in
+ abs >= fun (abs_i,abs_j) ->
+ (Tacticals.New.tclTHEN (Tacticals.New.tclDO abs_i intro)
+ (Tacticals.New.onLastHypId
(fun id ->
- elimination_then
+ Tacticals.New.elimination_then
(introElimAssumsThen (induction_trailer abs_i abs_j))
- ([],[]) (mkVar id)))) gls
+ ([],[]) (mkVar id))))
let h_double_induction = double_ind
diff --git a/tactics/elim.mli b/tactics/elim.mli
index d135997cd..35b7b2602 100644
--- a/tactics/elim.mli
+++ b/tactics/elim.mli
@@ -17,19 +17,19 @@ open Misctypes
(** Eliminations tactics. *)
val introElimAssumsThen :
- (branch_assumptions -> tactic) -> branch_args -> tactic
+ (branch_assumptions -> unit Proofview.tactic) -> branch_args -> unit Proofview.tactic
val introCaseAssumsThen :
- (intro_pattern_expr Loc.located list -> branch_assumptions -> tactic) ->
- branch_args -> tactic
+ (intro_pattern_expr Loc.located list -> branch_assumptions -> unit Proofview.tactic) ->
+ branch_args -> unit Proofview.tactic
-val general_decompose : (Id.t * constr -> bool) -> constr -> tactic
-val decompose_nonrec : constr -> tactic
-val decompose_and : constr -> tactic
-val decompose_or : constr -> tactic
-val h_decompose : inductive list -> constr -> tactic
-val h_decompose_or : constr -> tactic
-val h_decompose_and : constr -> tactic
+val general_decompose : (Id.t * constr -> bool) -> constr -> unit Proofview.tactic
+val decompose_nonrec : constr -> unit Proofview.tactic
+val decompose_and : constr -> unit Proofview.tactic
+val decompose_or : constr -> unit Proofview.tactic
+val h_decompose : inductive list -> constr -> unit Proofview.tactic
+val h_decompose_or : constr -> unit Proofview.tactic
+val h_decompose_and : constr -> unit Proofview.tactic
-val double_ind : quantified_hypothesis -> quantified_hypothesis -> tactic
-val h_double_induction : quantified_hypothesis -> quantified_hypothesis->tactic
+val double_ind : quantified_hypothesis -> quantified_hypothesis -> unit Proofview.tactic
+val h_double_induction : quantified_hypothesis -> quantified_hypothesis-> unit Proofview.tactic
diff --git a/tactics/eqdecide.ml4 b/tactics/eqdecide.ml4
index a5f8831a0..ac6e944fb 100644
--- a/tactics/eqdecide.ml4
+++ b/tactics/eqdecide.ml4
@@ -59,18 +59,18 @@ let choose_noteq eqonleft =
if eqonleft then h_simplest_right else h_simplest_left
let mkBranches c1 c2 =
- tclTHENSEQ
- [generalize [c2];
+ Tacticals.New.tclTHENLIST
+ [Proofview.V82.tactic (generalize [c2]);
h_simplest_elim c1;
intros;
- onLastHyp h_simplest_case;
- clear_last;
+ Tacticals.New.onLastHyp h_simplest_case;
+ Proofview.V82.tactic clear_last;
intros]
let solveNoteqBranch side =
- tclTHEN (choose_noteq side)
- (tclTHEN introf
- (onLastHypId (fun id -> Extratactics.discrHyp id)))
+ Tacticals.New.tclTHEN (choose_noteq side)
+ (Tacticals.New.tclTHEN introf
+ (Tacticals.New.onLastHypId (fun id -> Extratactics.discrHyp id)))
(* Constructs the type {c1=c2}+{~c1=c2} *)
@@ -93,42 +93,58 @@ let mkGenDecideEqGoal rectype g =
rectype (mkVar xname) (mkVar yname) g)))
let eqCase tac =
- (tclTHEN intro
- (tclTHEN (onLastHyp Equality.rewriteLR)
- (tclTHEN clear_last
+ (Tacticals.New.tclTHEN intro
+ (Tacticals.New.tclTHEN (Tacticals.New.onLastHyp Equality.rewriteLR)
+ (Tacticals.New.tclTHEN (Proofview.V82.tactic clear_last)
tac)))
let diseqCase eqonleft =
let diseq = Id.of_string "diseq" in
let absurd = Id.of_string "absurd" in
- (tclTHEN (intro_using diseq)
- (tclTHEN (choose_noteq eqonleft)
- (tclTHEN red_in_concl
- (tclTHEN (intro_using absurd)
- (tclTHEN (h_simplest_apply (mkVar diseq))
- (tclTHEN (Extratactics.injHyp absurd)
+ (Tacticals.New.tclTHEN (intro_using diseq)
+ (Tacticals.New.tclTHEN (choose_noteq eqonleft)
+ (Tacticals.New.tclTHEN (Proofview.V82.tactic red_in_concl)
+ (Tacticals.New.tclTHEN (intro_using absurd)
+ (Tacticals.New.tclTHEN (Proofview.V82.tactic (h_simplest_apply (mkVar diseq)))
+ (Tacticals.New.tclTHEN (Extratactics.injHyp absurd)
(full_trivial [])))))))
-let solveArg eqonleft op a1 a2 tac g =
- let rectype = pf_type_of g a1 in
- let decide = mkDecideEqGoal eqonleft op rectype a1 a2 g in
+open Proofview.Notations
+
+(* spiwack: a small wrapper around [Hipattern]. *)
+
+let match_eqdec c =
+ try Proofview.tclUNIT (match_eqdec c)
+ with PatternMatchingFailure -> Proofview.tclZERO PatternMatchingFailure
+
+(* /spiwack *)
+
+let solveArg eqonleft op a1 a2 tac =
+ Tacmach.New.of_old (fun g -> pf_type_of g a1) >>- fun rectype ->
+ Tacmach.New.of_old (fun g -> mkDecideEqGoal eqonleft op rectype a1 a2 g) >>- fun decide ->
let subtacs =
if eqonleft then [eqCase tac;diseqCase eqonleft;default_auto]
else [diseqCase eqonleft;eqCase tac;default_auto] in
- (tclTHENS (h_elim_type decide) subtacs) g
-
-let solveEqBranch rectype g =
- try
- let (eqonleft,op,lhs,rhs,_) = match_eqdec (pf_concl g) in
- let (mib,mip) = Global.lookup_inductive rectype in
- let nparams = mib.mind_nparams in
- let getargs l = List.skipn nparams (snd (decompose_app l)) in
- let rargs = getargs rhs
- and largs = getargs lhs in
- List.fold_right2
- (solveArg eqonleft op) largs rargs
- (tclTHEN (choose_eq eqonleft) h_reflexivity) g
- with PatternMatchingFailure -> error "Unexpected conclusion!"
+ (Tacticals.New.tclTHENS (Proofview.V82.tactic (h_elim_type decide)) subtacs)
+
+let solveEqBranch rectype =
+ Proofview.tclORELSE
+ begin
+ Goal.concl >>- fun concl ->
+ match_eqdec concl >= fun (eqonleft,op,lhs,rhs,_) ->
+ let (mib,mip) = Global.lookup_inductive rectype in
+ let nparams = mib.mind_nparams in
+ let getargs l = List.skipn nparams (snd (decompose_app l)) in
+ let rargs = getargs rhs
+ and largs = getargs lhs in
+ List.fold_right2
+ (solveArg eqonleft op) largs rargs
+ (Tacticals.New.tclTHEN (choose_eq eqonleft) h_reflexivity)
+ end
+ begin function
+ | PatternMatchingFailure -> Proofview.tclZERO (UserError ("",Pp.str"Unexpected conclusion!"))
+ | e -> Proofview.tclZERO e
+ end
(* The tactic Decide Equality *)
@@ -136,39 +152,43 @@ let hd_app c = match kind_of_term c with
| App (h,_) -> h
| _ -> c
-let decideGralEquality g =
- try
- let eqonleft,_,c1,c2,typ = match_eqdec (pf_concl g) in
- let headtyp = hd_app (pf_compute g typ) in
- let rectype =
- match kind_of_term headtyp with
- | Ind mi -> mi
- | _ -> error"This decision procedure only works for inductive objects."
- in
- (tclTHEN
- (mkBranches c1 c2)
- (tclORELSE (solveNoteqBranch eqonleft) (solveEqBranch rectype)))
- g
- with PatternMatchingFailure ->
- error "The goal must be of the form {x<>y}+{x=y} or {x=y}+{x<>y}."
-
-let decideEqualityGoal = tclTHEN intros decideGralEquality
-
-let decideEquality rectype g =
- let decide = mkGenDecideEqGoal rectype g in
- (tclTHENS (cut decide) [default_auto;decideEqualityGoal]) g
+let decideGralEquality =
+ Proofview.tclORELSE
+ begin
+ Goal.concl >>- fun concl ->
+ match_eqdec concl >= fun eqonleft,_,c1,c2,typ ->
+ Tacmach.New.of_old (fun g -> hd_app (pf_compute g typ)) >>- fun headtyp ->
+ begin match kind_of_term headtyp with
+ | Ind mi -> Proofview.tclUNIT mi
+ | _ -> Proofview.tclZERO (UserError ("",Pp.str"This decision procedure only works for inductive objects."))
+ end >= fun rectype ->
+ (Tacticals.New.tclTHEN
+ (mkBranches c1 c2)
+ (Tacticals.New.tclORELSE (solveNoteqBranch eqonleft) (solveEqBranch rectype)))
+ end
+ begin function
+ | PatternMatchingFailure ->
+ Proofview.tclZERO (UserError ("", Pp.str"The goal must be of the form {x<>y}+{x=y} or {x=y}+{x<>y}."))
+ | e -> Proofview.tclZERO e
+ end
+
+let decideEqualityGoal = Tacticals.New.tclTHEN intros decideGralEquality
+
+let decideEquality rectype =
+ Tacmach.New.of_old (mkGenDecideEqGoal rectype) >>- fun decide ->
+ (Tacticals.New.tclTHENS (Proofview.V82.tactic (cut decide)) [default_auto;decideEqualityGoal])
(* The tactic Compare *)
-let compare c1 c2 g =
- let rectype = pf_type_of g c1 in
- let decide = mkDecideEqGoal true (build_coq_sumbool ()) rectype c1 c2 g in
- (tclTHENS (cut decide)
- [(tclTHEN intro
- (tclTHEN (onLastHyp simplest_case)
- clear_last));
- decideEquality (pf_type_of g c1)]) g
+let compare c1 c2 =
+ Tacmach.New.of_old (fun g -> pf_type_of g c1) >>- fun rectype ->
+ Tacmach.New.of_old (fun g -> mkDecideEqGoal true (build_coq_sumbool ()) rectype c1 c2 g) >>- fun decide ->
+ (Tacticals.New.tclTHENS (Proofview.V82.tactic (cut decide))
+ [(Tacticals.New.tclTHEN intro
+ (Tacticals.New.tclTHEN (Tacticals.New.onLastHyp simplest_case)
+ (Proofview.V82.tactic clear_last)));
+ decideEquality rectype])
(* User syntax *)
diff --git a/tactics/equality.ml b/tactics/equality.ml
index 0864048f9..897285f25 100644
--- a/tactics/equality.ml
+++ b/tactics/equality.ml
@@ -42,6 +42,7 @@ open Eqschemes
open Locus
open Locusops
open Misctypes
+open Proofview.Notations
(* Options *)
@@ -131,7 +132,7 @@ let make_flags frzevars sigma flags clause =
let side_tac tac sidetac =
match sidetac with
| None -> tac
- | Some sidetac -> tclTHENSFIRSTn tac [|tclIDTAC|] sidetac
+ | Some sidetac -> Tacticals.New.tclTHENSFIRSTn tac [|Proofview.tclUNIT ()|] sidetac
let instantiate_lemma_all frzevars env gl c ty l l2r concl =
let eqclause = Clenv.make_clenv_binding gl (c,ty) l in
@@ -198,42 +199,52 @@ let rewrite_elim_in with_evars frzevars id c e gl =
(elimination_in_clause_scheme with_evars ~flags id) c e gl
(* Ad hoc asymmetric general_elim_clause *)
-let general_elim_clause with_evars frzevars cls rew elim =
+let general_elim_clause with_evars frzevars cls rew elim gl =
try
(match cls with
| None ->
(* was tclWEAK_PROGRESS which only fails for tactics generating one
subgoal and did not fail for useless conditional rewritings generating
an extra condition *)
- tclNOTSAMEGOAL (rewrite_elim with_evars frzevars rew elim)
- | Some id -> rewrite_elim_in with_evars frzevars id rew elim)
+ tclNOTSAMEGOAL (rewrite_elim with_evars frzevars rew elim) gl
+ | Some id -> rewrite_elim_in with_evars frzevars id rew elim gl)
with Pretype_errors.PretypeError (env,evd,
Pretype_errors.NoOccurrenceFound (c', _)) ->
raise (Pretype_errors.PretypeError
(env,evd,Pretype_errors.NoOccurrenceFound (c', cls)))
-let general_elim_clause with_evars frzevars tac cls c t l l2r elim gl =
+let general_elim_clause with_evars frzevars tac cls c t l l2r elim =
let all, firstonly, tac =
match tac with
| None -> false, false, None
| Some (tac, Naive) -> false, false, Some tac
- | Some (tac, FirstSolved) -> true, true, Some (tclCOMPLETE tac)
- | Some (tac, AllMatches) -> true, false, Some (tclCOMPLETE tac)
+ | Some (tac, FirstSolved) -> true, true, Some (Tacticals.New.tclCOMPLETE tac)
+ | Some (tac, AllMatches) -> true, false, Some (Tacticals.New.tclCOMPLETE tac)
in
let cs =
- (if not all then instantiate_lemma else instantiate_lemma_all frzevars)
- (pf_env gl) gl c t l l2r
- (match cls with None -> pf_concl gl | Some id -> pf_get_hyp_typ gl id)
+ Goal.env >- fun env ->
+ Goal.concl >- fun concl ->
+ Tacmach.New.of_old (
+ (if not all then instantiate_lemma else instantiate_lemma_all frzevars)
+ env ) >- fun instantiate_lemma ->
+ let typ =
+ match cls with None -> Goal.return concl | Some id -> Tacmach.New.pf_get_hyp_typ id
+ in
+ typ >- fun typ ->
+ Goal.return (instantiate_lemma c t l l2r typ)
in
let try_clause c =
side_tac
- (tclTHEN
- (Refiner.tclEVARS c.evd)
- (general_elim_clause with_evars frzevars cls c elim)) tac
+ (Tacticals.New.tclTHEN
+ (Proofview.V82.tactic (Refiner.tclEVARS c.evd))
+ (Proofview.V82.tactic (general_elim_clause with_evars frzevars cls c elim)))
+ tac
in
+ cs >>- fun cs ->
if firstonly then
- tclFIRST (List.map try_clause cs) gl
- else tclMAP try_clause cs gl
+ Tacticals.New.tclFIRST (List.map try_clause cs)
+ else
+ Tacticals.New.tclMAP try_clause cs
(* The next function decides in particular whether to try a regular
rewrite or a generalized rewrite.
@@ -282,12 +293,12 @@ let find_elim hdcncl lft2rgt dep cls ot gl =
begin
try
let _ = Global.lookup_constant c1' in
- mkConst c1', gl
+ mkConst c1', Declareops.no_seff
with Not_found ->
let rwr_thm = Label.to_string l' in
error ("Cannot find rewrite principle "^rwr_thm^".")
end
- | _ -> pr1, gl
+ | _ -> pr1, Declareops.no_seff
end
| _ ->
(* cannot occur since we checked that we are in presence of
@@ -309,22 +320,23 @@ let find_elim hdcncl lft2rgt dep cls ot gl =
match kind_of_term hdcncl with
| Ind ind ->
let c, eff = find_scheme scheme_name ind in
- let gl = {gl with sigma = Evd.emit_side_effects eff gl.sigma } in
- mkConst c, gl
+ mkConst c , eff
| _ -> assert false
-let type_of_clause gl = function
- | None -> pf_concl gl
- | Some id -> pf_get_hyp_typ gl id
+let type_of_clause = function
+ | None -> Goal.concl
+ | Some id -> Tacmach.New.pf_get_hyp_typ id
-let leibniz_rewrite_ebindings_clause cls lft2rgt tac c t l with_evars frzevars dep_proof_ok gl hdcncl =
+let leibniz_rewrite_ebindings_clause cls lft2rgt tac c t l with_evars frzevars dep_proof_ok hdcncl =
let isatomic = isProd (whd_zeta hdcncl) in
let dep_fun = if isatomic then dependent else dependent_no_evar in
- let dep = dep_proof_ok && dep_fun c (type_of_clause gl cls) in
- let elim, gl = find_elim hdcncl lft2rgt dep cls (Some t) gl in
+ type_of_clause cls >>- fun type_of_cls ->
+ let dep = dep_proof_ok && dep_fun c type_of_cls in
+ Tacmach.New.of_old (find_elim hdcncl lft2rgt dep cls (Some t)) >>- fun (elim,effs) ->
+ Proofview.V82.tactic (Tactics.emit_side_effects effs) <*>
general_elim_clause with_evars frzevars tac cls c t l
(match lft2rgt with None -> false | Some b -> b)
- {elimindex = None; elimbody = (elim,NoBindings)} gl
+ {elimindex = None; elimbody = (elim,NoBindings)}
let adjust_rewriting_direction args lft2rgt =
match args with
@@ -343,35 +355,42 @@ let rewrite_side_tac tac sidetac = side_tac tac (Option.map fst sidetac)
(* Main function for dispatching which kind of rewriting it is about *)
let general_rewrite_ebindings_clause cls lft2rgt occs frzevars dep_proof_ok ?tac
- ((c,l) : constr with_bindings) with_evars gl =
+ ((c,l) : constr with_bindings) with_evars =
if occs != AllOccurrences then (
- rewrite_side_tac (Hook.get forward_general_rewrite_clause cls lft2rgt occs (c,l) ~new_goals:[]) tac gl)
+ rewrite_side_tac (Hook.get forward_general_rewrite_clause cls lft2rgt occs (c,l) ~new_goals:[]) tac)
else
- let env = pf_env gl in
- let sigma = project gl in
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
let ctype = get_type_of env sigma c in
let rels, t = decompose_prod_assum (whd_betaiotazeta sigma ctype) in
match match_with_equality_type t with
| Some (hdcncl,args) -> (* Fast path: direct leibniz-like rewrite *)
let lft2rgt = adjust_rewriting_direction args lft2rgt in
leibniz_rewrite_ebindings_clause cls lft2rgt tac c (it_mkProd_or_LetIn t rels)
- l with_evars frzevars dep_proof_ok gl hdcncl
+ l with_evars frzevars dep_proof_ok hdcncl
| None ->
- try
- rewrite_side_tac (Hook.get forward_general_rewrite_clause cls
- lft2rgt occs (c,l) ~new_goals:[]) tac gl
- with e when Errors.noncritical e ->
- (* Try to see if there's an equality hidden *)
- let e = Errors.push e in
- let env' = push_rel_context rels env in
- let rels',t' = splay_prod_assum env' sigma t in (* Search for underlying eq *)
- match match_with_equality_type t' with
- | Some (hdcncl,args) ->
+ Proofview.tclORELSE
+ begin
+ rewrite_side_tac (Hook.get forward_general_rewrite_clause cls
+ lft2rgt occs (c,l) ~new_goals:[]) tac
+ end
+ begin function
+ | e ->
+ (* Try to see if there's an equality hidden *)
+ (* spiwack: [Errors.push] here is unlikely to do
+ what it's intended to, or anything meaningful for
+ that matter. *)
+ let e = Errors.push e in
+ let env' = push_rel_context rels env in
+ let rels',t' = splay_prod_assum env' sigma t in (* Search for underlying eq *)
+ match match_with_equality_type t' with
+ | Some (hdcncl,args) ->
let lft2rgt = adjust_rewriting_direction args lft2rgt in
- leibniz_rewrite_ebindings_clause cls lft2rgt tac c
- (it_mkProd_or_LetIn t' (rels' @ rels)) l with_evars frzevars dep_proof_ok gl hdcncl
- | None -> raise e
- (* error "The provided term does not end with an equality or a declared rewrite relation." *)
+ leibniz_rewrite_ebindings_clause cls lft2rgt tac c
+ (it_mkProd_or_LetIn t' (rels' @ rels)) l with_evars frzevars dep_proof_ok hdcncl
+ | None -> Proofview.tclZERO e
+ (* error "The provided term does not end with an equality or a declared rewrite relation." *)
+ end
let general_rewrite_ebindings =
general_rewrite_ebindings_clause None
@@ -406,35 +425,38 @@ let general_multi_rewrite l2r with_evars ?tac c cl =
(* If a precise list of locations is given, success is mandatory for
each of these locations. *)
let rec do_hyps = function
- | [] -> tclIDTAC
+ | [] -> Proofview.tclUNIT ()
| ((occs,id),_) :: l ->
- tclTHENFIRST
+ Tacticals.New.tclTHENFIRST
(general_rewrite_ebindings_in l2r (occs_of occs) false true ?tac id c with_evars)
(do_hyps l)
in
if cl.concl_occs == NoOccurrences then do_hyps l else
- tclTHENFIRST
- (general_rewrite_ebindings l2r (occs_of cl.concl_occs) false true ?tac c with_evars)
- (do_hyps l)
+ Tacticals.New.tclTHENFIRST
+ (general_rewrite_ebindings l2r (occs_of cl.concl_occs) false true ?tac c with_evars)
+ (do_hyps l)
| None ->
(* Otherwise, if we are told to rewrite in all hypothesis via the
syntax "* |-", we fail iff all the different rewrites fail *)
let rec do_hyps_atleastonce = function
- | [] -> (fun gl -> error "Nothing to rewrite.")
+ | [] -> Proofview.tclZERO (Errors.UserError ("",Pp.str"Nothing to rewrite."))
| id :: l ->
- tclIFTHENTRYELSEMUST
+ Tacticals.New.tclIFTHENTRYELSEMUST
(general_rewrite_ebindings_in l2r AllOccurrences false true ?tac id c with_evars)
(do_hyps_atleastonce l)
in
- let do_hyps gl =
+ let do_hyps =
(* If the term to rewrite uses an hypothesis H, don't rewrite in H *)
let ids =
let ids_in_c = Environ.global_vars_set (Global.env()) (fst c) in
- Id.Set.fold (fun id l -> List.remove Id.equal id l) ids_in_c (pf_ids_of_hyps gl)
- in do_hyps_atleastonce ids gl
+ Tacmach.New.pf_ids_of_hyps >- fun ids_of_hyps ->
+ Goal.return (Id.Set.fold (fun id l -> List.remove Id.equal id l) ids_in_c ids_of_hyps)
+ in
+ ids >>- fun ids ->
+ do_hyps_atleastonce ids
in
if cl.concl_occs == NoOccurrences then do_hyps else
- tclIFTHENTRYELSEMUST
+ Tacticals.New.tclIFTHENTRYELSEMUST
(general_rewrite_ebindings l2r (occs_of cl.concl_occs) false true ?tac c with_evars)
do_hyps
@@ -442,22 +464,24 @@ type delayed_open_constr_with_bindings =
env -> evar_map -> evar_map * constr with_bindings
let general_multi_multi_rewrite with_evars l cl tac =
- let do1 l2r f gl =
- let sigma,c = f (pf_env gl) (project gl) in
- Refiner.tclWITHHOLES with_evars
- (general_multi_rewrite l2r with_evars ?tac c) sigma cl gl in
+ let do1 l2r f =
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
+ let sigma,c = f env sigma in
+ Tacticals.New.tclWITHHOLES with_evars
+ (general_multi_rewrite l2r with_evars ?tac c) sigma cl in
let rec doN l2r c = function
- | Precisely n when n <= 0 -> tclIDTAC
+ | Precisely n when n <= 0 -> Proofview.tclUNIT ()
| Precisely 1 -> do1 l2r c
- | Precisely n -> tclTHENFIRST (do1 l2r c) (doN l2r c (Precisely (n-1)))
- | RepeatStar -> tclREPEAT_MAIN (do1 l2r c)
- | RepeatPlus -> tclTHENFIRST (do1 l2r c) (doN l2r c RepeatStar)
- | UpTo n when n<=0 -> tclIDTAC
- | UpTo n -> tclTHENFIRST (tclTRY (do1 l2r c)) (doN l2r c (UpTo (n-1)))
+ | Precisely n -> Tacticals.New.tclTHENFIRST (do1 l2r c) (doN l2r c (Precisely (n-1)))
+ | RepeatStar -> Tacticals.New.tclREPEAT_MAIN (do1 l2r c)
+ | RepeatPlus -> Tacticals.New.tclTHENFIRST (do1 l2r c) (doN l2r c RepeatStar)
+ | UpTo n when n<=0 -> Proofview.tclUNIT ()
+ | UpTo n -> Tacticals.New.tclTHENFIRST (Tacticals.New.tclTRY (do1 l2r c)) (doN l2r c (UpTo (n-1)))
in
let rec loop = function
- | [] -> tclIDTAC
- | (l2r,m,c)::l -> tclTHENFIRST (doN l2r c m) (loop l)
+ | [] -> Proofview.tclUNIT ()
+ | (l2r,m,c)::l -> Tacticals.New.tclTHENFIRST (doN l2r c m) (loop l)
in loop l
let rewriteLR = general_rewrite true AllOccurrences true true
@@ -471,43 +495,45 @@ let rewriteRL = general_rewrite false AllOccurrences true true
tac : Used to prove the equality c1 = c2
gl : goal *)
-let multi_replace clause c2 c1 unsafe try_prove_eq_opt gl =
+let multi_replace clause c2 c1 unsafe try_prove_eq_opt =
let try_prove_eq =
match try_prove_eq_opt with
- | None -> tclIDTAC
- | Some tac -> tclCOMPLETE tac
+ | None -> Proofview.tclUNIT ()
+ | Some tac -> Tacticals.New.tclCOMPLETE tac
in
- let t1 = pf_apply get_type_of gl c1
- and t2 = pf_apply get_type_of gl c2 in
- if unsafe || (pf_conv_x gl t1 t2) then
+ Tacmach.New.pf_apply get_type_of >>- fun get_type_of ->
+ let t1 = get_type_of c1
+ and t2 = get_type_of c2 in
+ Tacmach.New.pf_apply is_conv >>- fun is_conv ->
+ if unsafe || (is_conv t1 t2) then
let e = build_coq_eq () in
let sym = build_coq_eq_sym () in
let eq = applist (e, [t1;c1;c2]) in
- tclTHENS (assert_as false None eq)
- [onLastHypId (fun id ->
- tclTHEN
- (tclTRY (general_multi_rewrite false false (mkVar id,NoBindings) clause))
- (clear [id]));
- tclFIRST
- [assumption;
- tclTHEN (apply sym) assumption;
+ Tacticals.New.tclTHENS (assert_as false None eq)
+ [Tacticals.New.onLastHypId (fun id ->
+ Tacticals.New.tclTHEN
+ (Tacticals.New.tclTRY (general_multi_rewrite false false (mkVar id,NoBindings) clause))
+ (Proofview.V82.tactic (clear [id])));
+ Tacticals.New.tclFIRST
+ [Proofview.V82.tactic assumption;
+ Tacticals.New.tclTHEN (Proofview.V82.tactic (apply sym)) (Proofview.V82.tactic assumption);
try_prove_eq
]
- ] gl
+ ]
else
- error "Terms do not have convertible types."
+ Tacticals.New.tclFAIL 0 (str"Terms do not have convertible types.")
-let replace c2 c1 gl = multi_replace onConcl c2 c1 false None gl
+let replace c2 c1 = multi_replace onConcl c2 c1 false None
-let replace_in id c2 c1 gl = multi_replace (onHyp id) c2 c1 false None gl
+let replace_in id c2 c1 = multi_replace (onHyp id) c2 c1 false None
-let replace_by c2 c1 tac gl = multi_replace onConcl c2 c1 false (Some tac) gl
+let replace_by c2 c1 tac = multi_replace onConcl c2 c1 false (Some tac)
-let replace_in_by id c2 c1 tac gl = multi_replace (onHyp id) c2 c1 false (Some tac) gl
+let replace_in_by id c2 c1 tac = multi_replace (onHyp id) c2 c1 false (Some tac)
-let replace_in_clause_maybe_by c2 c1 cl tac_opt gl =
- multi_replace cl c2 c1 false tac_opt gl
+let replace_in_clause_maybe_by c2 c1 cl tac_opt =
+ multi_replace cl c2 c1 false tac_opt
(* End of Eduardo's code. The rest of this file could be improved
using the functions match_with_equation, etc that I defined
@@ -760,13 +786,13 @@ let rec build_discriminator sigma env dirn c sort = function
Goal ~ c _ 0 0 = c _ 0 1. intro. discriminate H.
*)
-let gen_absurdity id gl =
- if is_empty_type (pf_get_hyp_typ gl id)
+let gen_absurdity id =
+ Tacmach.New.pf_get_hyp_typ id >>- fun hyp_typ ->
+ if is_empty_type hyp_typ
then
- simplest_elim (mkVar id) gl
+ simplest_elim (mkVar id)
else
- errorlabstrm "Equality.gen_absurdity"
- (str "Not the negation of an equality.")
+ Proofview.tclZERO (Errors.UserError ("Equality.gen_absurdity" , str "Not the negation of an equality."))
(* Precondition: eq is leibniz equality
@@ -804,7 +830,7 @@ let apply_on_clause (f,t) clause =
| _ -> errorlabstrm "" (str "Ill-formed clause applicator.")) in
clenv_fchain argmv f_clause clause
-let discr_positions env sigma (lbeq,eqn,(t,t1,t2)) eq_clause cpath dirn sort gl=
+let discr_positions env sigma (lbeq,eqn,(t,t1,t2)) eq_clause cpath dirn sort =
let e = next_ident_away eq_baseid (ids_of_context env) in
let e_env = push_named (e,None,t) env in
let discriminator =
@@ -814,40 +840,48 @@ let discr_positions env sigma (lbeq,eqn,(t,t1,t2)) eq_clause cpath dirn sort gl=
let pf_ty = mkArrow eqn absurd_term in
let absurd_clause = apply_on_clause (pf,pf_ty) eq_clause in
let pf = clenv_value_cast_meta absurd_clause in
- let gl = {gl with sigma = Evd.emit_side_effects eff gl.sigma } in
- tclTHENS (cut_intro absurd_term)
- [onLastHypId gen_absurdity; refine pf] gl
+ Proofview.V82.tactic (Tactics.emit_side_effects eff) <*>
+ Tacticals.New.tclTHENS (cut_intro absurd_term)
+ [Tacticals.New.onLastHypId gen_absurdity; (Proofview.V82.tactic (refine pf))]
-let discrEq (lbeq,_,(t,t1,t2) as u) eq_clause gls =
+let discrEq (lbeq,_,(t,t1,t2) as u) eq_clause =
let sigma = eq_clause.evd in
- let env = pf_env gls in
+ Goal.env >>- fun env ->
+ Goal.concl >>- fun concl ->
match find_positions env sigma t1 t2 with
| Inr _ ->
- errorlabstrm "discr" (str"Not a discriminable equality.")
+ Proofview.tclZERO (Errors.UserError ("discr" , str"Not a discriminable equality."))
| Inl (cpath, (_,dirn), _) ->
- let sort = pf_apply get_type_of gls (pf_concl gls) in
- discr_positions env sigma u eq_clause cpath dirn sort gls
-
-let onEquality with_evars tac (c,lbindc) gls =
- let t = pf_type_of gls c in
- let t' = try snd (pf_reduce_to_quantified_ind gls t) with UserError _ -> t in
- let eq_clause = make_clenv_binding gls (c,t') lbindc in
+ Tacmach.New.pf_apply get_type_of >>- fun get_type_of ->
+ let sort = get_type_of concl in
+ discr_positions env sigma u eq_clause cpath dirn sort
+
+let onEquality with_evars tac (c,lbindc) =
+ Tacmach.New.pf_apply Typing.type_of >>- fun type_of ->
+ Tacmach.New.pf_apply Tacred.reduce_to_quantified_ind >>- fun reduce_to_quantified_ind ->
+ let t = type_of c in
+ let t' = try snd (reduce_to_quantified_ind t) with UserError _ -> t in
+ Tacmach.New.of_old make_clenv_binding >>- fun make_clenv_binding ->
+ let eq_clause = make_clenv_binding (c,t') lbindc in
let eq_clause' = clenv_pose_dependent_evars with_evars eq_clause in
let eqn = clenv_type eq_clause' in
- let eq,eq_args = find_this_eq_data_decompose gls eqn in
- tclTHEN
- (Refiner.tclEVARS eq_clause'.evd)
- (tac (eq,eqn,eq_args) eq_clause') gls
-
-let onNegatedEquality with_evars tac gls =
- let ccl = pf_concl gls in
- match kind_of_term (hnf_constr (pf_env gls) (project gls) ccl) with
+ Tacmach.New.of_old find_this_eq_data_decompose >>- fun find_this_eq_data_decompose ->
+ let eq,eq_args = find_this_eq_data_decompose eqn in
+ Tacticals.New.tclTHEN
+ (Proofview.V82.tactic (Refiner.tclEVARS eq_clause'.evd))
+ (tac (eq,eqn,eq_args) eq_clause')
+
+let onNegatedEquality with_evars tac =
+ Goal.concl >>- fun ccl ->
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
+ match kind_of_term (hnf_constr env sigma ccl) with
| Prod (_,t,u) when is_empty_type u ->
- tclTHEN introf
- (onLastHypId (fun id ->
- onEquality with_evars tac (mkVar id,NoBindings))) gls
+ Tacticals.New.tclTHEN introf
+ (Tacticals.New.onLastHypId (fun id ->
+ onEquality with_evars tac (mkVar id,NoBindings)))
| _ ->
- errorlabstrm "" (str "Not a negated primitive equality.")
+ Proofview.tclZERO (Errors.UserError ("" , str "Not a negated primitive equality."))
let discrSimpleClause with_evars = function
| None -> onNegatedEquality with_evars discrEq
@@ -855,19 +889,19 @@ let discrSimpleClause with_evars = function
let discr with_evars = onEquality with_evars discrEq
-let discrClause with_evars = onClause (discrSimpleClause with_evars)
+let discrClause with_evars = Tacticals.New.onClause (discrSimpleClause with_evars)
let discrEverywhere with_evars =
(*
tclORELSE
*)
(if discr_do_intro () then
- (tclTHEN
- (tclREPEAT introf)
- (Tacticals.tryAllHyps
- (fun id -> tclCOMPLETE (discr with_evars (mkVar id,NoBindings)))))
+ (Tacticals.New.tclTHEN
+ (Tacticals.New.tclREPEAT introf)
+ (Tacticals.New.tryAllHyps
+ (fun id -> Tacticals.New.tclCOMPLETE (discr with_evars (mkVar id,NoBindings)))))
else (* <= 8.2 compat *)
- Tacticals.tryAllHypsAndConcl (discrSimpleClause with_evars))
+ Tacticals.New.tryAllHypsAndConcl (discrSimpleClause with_evars))
(* (fun gls ->
errorlabstrm "DiscrEverywhere" (str"No discriminable equalities."))
*)
@@ -875,8 +909,8 @@ let discr_tac with_evars = function
| None -> discrEverywhere with_evars
| Some c -> onInductionArg (discr with_evars) c
-let discrConcl gls = discrClause false onConcl gls
-let discrHyp id gls = discrClause false (onHyp id) gls
+let discrConcl = discrClause false onConcl
+let discrHyp id = discrClause false (onHyp id)
(* returns the sigma type (sigS, sigT) with the respective
constructor depending on the sort *)
@@ -1131,12 +1165,13 @@ let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac =
in
let injectors = List.map_filter filter posns in
if List.is_empty injectors then
- errorlabstrm "Equality.inj" (str "Failed to decompose the equality.");
- tclTHEN
- (tclMAP
- (fun (pf,ty) -> tclTHENS (cut ty) [tclIDTAC; refine pf])
- (if l2r then List.rev injectors else injectors))
- (tac (List.length injectors))
+ Proofview.tclZERO (Errors.UserError ("Equality.inj" , str "Failed to decompose the equality."))
+ else
+ Tacticals.New.tclTHEN
+ (Proofview.V82.tactic (tclMAP
+ (fun (pf,ty) -> tclTHENS (cut ty) [tclIDTAC; refine pf])
+ (if l2r then List.rev injectors else injectors)))
+ (tac (List.length injectors))
exception Not_dep_pair
@@ -1170,29 +1205,31 @@ let inject_if_homogenous_dependent_pair env sigma (eq,_,(t,t1,t2)) gl =
tclTHENS (cut (mkApp (ceq,new_eq_args)))
[tclIDTAC; tclTHEN (apply (
mkApp(inj2,[|ar1.(0);mkConst c;ar1.(1);ar1.(2);ar1.(3);ar2.(3)|])
- )) (Auto.trivial [] [])
+ )) (Proofview.V82.of_tactic (Auto.trivial [] []))
] gl
(* not a dep eq or no decidable type found *)
end
else raise Not_dep_pair
-let injEqThen tac l2r (eq,_,(t,t1,t2) as u) eq_clause gl =
+let injEqThen tac l2r (eq,_,(t,t1,t2) as u) eq_clause =
let sigma = eq_clause.evd in
let env = eq_clause.env in
match find_positions env sigma t1 t2 with
| Inl _ ->
- errorlabstrm "Inj"
- (str"Not a projectable equality but a discriminable one.")
+ Proofview.tclZERO (Errors.UserError ("Inj",str"Not a projectable equality but a discriminable one."))
| Inr [] ->
- errorlabstrm "Equality.inj"
- (str"Nothing to do, it is an equality between convertible terms.")
+ Proofview.tclZERO (Errors.UserError ("Equality.inj",str"Nothing to do, it is an equality between convertible terms."))
| Inr [([],_,_)] when Flags.version_strictly_greater Flags.V8_3 ->
- errorlabstrm "Equality.inj" (str"Nothing to inject.")
+ Proofview.tclZERO (Errors.UserError ("Equality.inj" , str"Nothing to inject."))
| Inr posns ->
- try inject_if_homogenous_dependent_pair env sigma u gl
- with Not_dep_pair as e | e when Errors.noncritical e ->
- inject_at_positions env sigma l2r u eq_clause posns
- (tac (clenv_value eq_clause)) gl
+ Proofview.tclORELSE
+ begin Proofview.V82.tactic (inject_if_homogenous_dependent_pair env sigma u) end
+ begin function
+ | Not_dep_pair as e |e when Errors.noncritical e ->
+ inject_at_positions env sigma l2r u eq_clause posns
+ (tac (clenv_value eq_clause))
+ | reraise -> Proofview.tclZERO reraise
+ end
let postInjEqTac ipats c n =
match ipats with
@@ -1207,14 +1244,14 @@ let postInjEqTac ipats c n =
else ipats in
tclTHEN
(clear_tac)
- (intros_pattern MoveLast ipats)
+ (Proofview.V82.of_tactic (intros_pattern MoveLast ipats))
| None -> tclIDTAC
let injEq ipats =
let l2r =
if use_injection_pattern_l2r_order () && not (Option.is_empty ipats) then true else false
in
- injEqThen (postInjEqTac ipats) l2r
+ injEqThen (fun c i -> Proofview.V82.tactic (postInjEqTac ipats c i)) l2r
let inj ipats with_evars = onEquality with_evars (injEq ipats)
@@ -1222,27 +1259,27 @@ let injClause ipats with_evars = function
| None -> onNegatedEquality with_evars (injEq ipats)
| Some c -> onInductionArg (inj ipats with_evars) c
-let injConcl gls = injClause None false None gls
-let injHyp id gls = injClause None false (Some (ElimOnIdent (Loc.ghost,id))) gls
-
-let decompEqThen ntac (lbeq,_,(t,t1,t2) as u) clause gls =
- let sort = pf_apply get_type_of gls (pf_concl gls) in
- let sigma = clause.evd in
- let env = pf_env gls in
- match find_positions env sigma t1 t2 with
- | Inl (cpath, (_,dirn), _) ->
- discr_positions env sigma u clause cpath dirn sort gls
- | Inr [] -> (* Change: do not fail, simplify clear this trivial hyp *)
- ntac (clenv_value clause) 0 gls
+let injConcl = injClause None false None
+let injHyp id = injClause None false (Some (ElimOnIdent (Loc.ghost,id)))
+
+let decompEqThen ntac (lbeq,_,(t,t1,t2) as u) clause =
+ Tacmach.New.of_old (fun gls -> pf_apply get_type_of gls (pf_concl gls)) >>- fun sort ->
+ let sigma = clause.evd in
+ Goal.env >>- fun env ->
+ match find_positions env sigma t1 t2 with
+ | Inl (cpath, (_,dirn), _) ->
+ discr_positions env sigma u clause cpath dirn sort
+ | Inr [] -> (* Change: do not fail, simplify clear this trivial hyp *)
+ ntac (clenv_value clause) 0
| Inr posns ->
inject_at_positions env sigma true u clause posns
- (ntac (clenv_value clause)) gls
+ (ntac (clenv_value clause))
let dEqThen with_evars ntac = function
| None -> onNegatedEquality with_evars (decompEqThen ntac)
| Some c -> onInductionArg (onEquality with_evars (decompEqThen ntac)) c
-let dEq with_evars = dEqThen with_evars (fun c x -> tclIDTAC)
+let dEq with_evars = dEqThen with_evars (fun c x -> Proofview.tclUNIT ())
let _ = declare_intro_decomp_eq (fun tac -> decompEqThen (fun _ -> tac))
@@ -1266,7 +1303,7 @@ let swapEquandsInConcl gls =
let bareRevSubstInConcl lbeq body (t,e1,e2) gls =
(* find substitution scheme *)
- let eq_elim, gls = find_elim lbeq.eq (Some false) false None None gls in
+ let eq_elim, effs = find_elim lbeq.eq (Some false) false None None gls in
(* build substitution predicate *)
let p = lambda_create (pf_env gls) (t,body) in
(* apply substitution scheme *)
@@ -1471,16 +1508,20 @@ let is_eq_x gl x (id,_,c) =
(* Rewrite "hyp:x=rhs" or "hyp:rhs=x" (if dir=false) everywhere and
erase hyp and x; proceed by generalizing all dep hyps *)
-let subst_one dep_proof_ok x (hyp,rhs,dir) gl =
+let subst_one dep_proof_ok x (hyp,rhs,dir) =
+ Goal.env >>- fun env ->
+ Goal.hyps >>- fun hyps ->
+ Goal.concl >>- fun concl ->
+ let hyps = Environ.named_context_of_val hyps in
(* The set of hypotheses using x *)
let depdecls =
let test (id,_,c as dcl) =
- if not (Id.equal id hyp) && occur_var_in_decl (pf_env gl) x dcl then Some dcl
+ if not (Id.equal id hyp) && occur_var_in_decl env x dcl then Some dcl
else None in
- List.rev (List.map_filter test (pf_hyps gl)) in
+ List.rev (List.map_filter test hyps) in
let dephyps = List.map (fun (id,_,_) -> id) depdecls in
(* Decides if x appears in conclusion *)
- let depconcl = occur_var (pf_env gl) x (pf_concl gl) in
+ let depconcl = occur_var env x concl in
(* The set of non-defined hypothesis: they must be abstracted,
rewritten and reintroduced *)
let abshyps =
@@ -1496,28 +1537,32 @@ let subst_one dep_proof_ok x (hyp,rhs,dir) gl =
(Some (replace_term (mkVar x) rhs htyp)) nowhere
in
let need_rewrite = not (List.is_empty dephyps) || depconcl in
- tclTHENLIST
+ Tacticals.New.tclTHENLIST
((if need_rewrite then
- [generalize abshyps;
+ [Proofview.V82.tactic (generalize abshyps);
general_rewrite dir AllOccurrences true dep_proof_ok (mkVar hyp);
- thin dephyps;
- tclMAP introtac depdecls]
+ Proofview.V82.tactic (thin dephyps);
+ (Tacticals.New.tclMAP introtac depdecls)]
else
- [tclIDTAC]) @
- [tclTRY (clear [x;hyp])]) gl
+ [Proofview.tclUNIT ()]) @
+ [Proofview.V82.tactic (tclTRY (clear [x;hyp]))])
(* Look for an hypothesis hyp of the form "x=rhs" or "rhs=x", rewrite
it everywhere, and erase hyp and x; proceed by generalizing all dep hyps *)
-let subst_one_var dep_proof_ok x gl =
- let hyps = pf_hyps gl in
- let (_,xval,_) = pf_get_hyp gl x in
+(* arnaud: il va y avoir un bug là-dedans. Le try ne se déclenche pas
+ au bon endroit. Il faut convertir test en une Proofview.tactic
+ pour la gestion des exceptions. *)
+let subst_one_var dep_proof_ok x =
+ Goal.hyps >>- fun hyps ->
+ let hyps = Environ.named_context_of_val hyps in
+ Tacmach.New.pf_get_hyp x >>- fun (_,xval,_) ->
(* If x has a body, simply replace x with body and clear x *)
- if not (Option.is_empty xval) then tclTHEN (unfold_body x) (clear [x]) gl else
+ if not (Option.is_empty xval) then Proofview.V82.tactic (tclTHEN (unfold_body x) (clear [x])) else
(* x is a variable: *)
let varx = mkVar x in
(* Find a non-recursive definition for x *)
- let (hyp,rhs,dir) =
+ let found gl =
try
let test hyp _ = is_eq_x gl varx hyp in
Context.fold_named_context test ~init:() hyps;
@@ -1525,10 +1570,11 @@ let subst_one_var dep_proof_ok x gl =
(str "Cannot find any non-recursive equality over " ++ pr_id x ++
str".")
with FoundHyp res -> res in
- subst_one dep_proof_ok x (hyp,rhs,dir) gl
+ Tacmach.New.of_old found >>- fun (hyp,rhs,dir) ->
+ subst_one dep_proof_ok x (hyp,rhs,dir)
let subst_gen dep_proof_ok ids =
- tclTHEN tclNORMEVAR (tclMAP (subst_one_var dep_proof_ok) ids)
+ Tacticals.New.tclTHEN (Proofview.V82.tactic tclNORMEVAR) (Tacticals.New.tclMAP (subst_one_var dep_proof_ok) ids)
(* For every x, look for an hypothesis hyp of the form "x=rhs" or "rhs=x",
rewrite it everywhere, and erase hyp and x; proceed by generalizing
@@ -1547,10 +1593,12 @@ let default_subst_tactic_flags () =
else
{ only_leibniz = true; rewrite_dependent_proof = false }
-let subst_all ?(flags=default_subst_tactic_flags ()) gl =
+(* arnaud: encore des erreurs potentiels avec ces try/with *)
+let subst_all ?(flags=default_subst_tactic_flags ()) =
+ Tacmach.New.of_old find_eq_data_decompose >>- fun find_eq_data_decompose ->
let test (_,c) =
try
- let lbeq,(_,x,y) = find_eq_data_decompose gl c in
+ let lbeq,(_,x,y) = find_eq_data_decompose c in
if flags.only_leibniz then restrict_to_eq_and_identity lbeq.eq;
(* J.F.: added to prevent failure on goal containing x=x as an hyp *)
if eq_constr x y then failwith "caught";
@@ -1559,9 +1607,10 @@ let subst_all ?(flags=default_subst_tactic_flags ()) gl =
with PatternMatchingFailure -> failwith "caught"
in
let test p = try Some (test p) with Failure _ -> None in
- let ids = List.map_filter test (pf_hyps_types gl) in
+ Tacmach.New.pf_hyps_types >>- fun hyps ->
+ let ids = List.map_filter test hyps in
let ids = List.uniquize ids in
- subst_gen flags.rewrite_dependent_proof ids gl
+ subst_gen flags.rewrite_dependent_proof ids
(* Rewrite the first assumption for which the condition faildir does not fail
and gives the direction of the rewrite *)
@@ -1586,18 +1635,25 @@ let cond_eq_term c t gl =
else failwith "not convertible"
with PatternMatchingFailure -> failwith "not an equality"
-let rewrite_multi_assumption_cond cond_eq_term cl gl =
+let cond_eq_term_left c t = Tacmach.New.of_old (cond_eq_term_left c t)
+let cond_eq_term_right c t = Tacmach.New.of_old (cond_eq_term_right c t)
+let cond_eq_term c t = Tacmach.New.of_old (cond_eq_term c t)
+
+(* arnaud: toujours des histoires de try/with *)
+let rewrite_multi_assumption_cond cond_eq_term cl =
let rec arec = function
| [] -> error "No such assumption."
| (id,_,t) ::rest ->
begin
try
- let dir = cond_eq_term t gl in
- general_multi_rewrite dir false (mkVar id,NoBindings) cl gl
+ cond_eq_term t >>- fun dir ->
+ general_multi_rewrite dir false (mkVar id,NoBindings) cl
with | Failure _ | UserError _ -> arec rest
end
in
- arec (pf_hyps gl)
+ Goal.hyps >>- fun hyps ->
+ let hyps = Environ.named_context_of_val hyps in
+ arec hyps
let replace_multi_term dir_opt c =
let cond_eq_fun =
diff --git a/tactics/equality.mli b/tactics/equality.mli
index 70480d26e..3bb204caf 100644
--- a/tactics/equality.mli
+++ b/tactics/equality.mli
@@ -40,66 +40,66 @@ type conditions =
val general_rewrite_bindings :
orientation -> occurrences -> freeze_evars_flag -> dep_proof_flag ->
- ?tac:(tactic * conditions) -> constr with_bindings -> evars_flag -> tactic
+ ?tac:(unit Proofview.tactic * conditions) -> constr with_bindings -> evars_flag -> unit Proofview.tactic
val general_rewrite :
orientation -> occurrences -> freeze_evars_flag -> dep_proof_flag ->
- ?tac:(tactic * conditions) -> constr -> tactic
+ ?tac:(unit Proofview.tactic * conditions) -> constr -> unit Proofview.tactic
(* Equivalent to [general_rewrite l2r] *)
-val rewriteLR : ?tac:(tactic * conditions) -> constr -> tactic
-val rewriteRL : ?tac:(tactic * conditions) -> constr -> tactic
+val rewriteLR : ?tac:(unit Proofview.tactic * conditions) -> constr -> unit Proofview.tactic
+val rewriteRL : ?tac:(unit Proofview.tactic * conditions) -> constr -> unit Proofview.tactic
(* Warning: old [general_rewrite_in] is now [general_rewrite_bindings_in] *)
val general_rewrite_clause :
(Id.t option -> orientation ->
- occurrences -> constr with_bindings -> new_goals:constr list -> tactic) Hook.t
+ occurrences -> constr with_bindings -> new_goals:constr list -> unit Proofview.tactic) Hook.t
val is_applied_rewrite_relation : (env -> evar_map -> rel_context -> constr -> constr option) Hook.t
val general_rewrite_ebindings_clause : Id.t option ->
orientation -> occurrences -> freeze_evars_flag -> dep_proof_flag ->
- ?tac:(tactic * conditions) -> constr with_bindings -> evars_flag -> tactic
+ ?tac:(unit Proofview.tactic * conditions) -> constr with_bindings -> evars_flag -> unit Proofview.tactic
val general_rewrite_bindings_in :
orientation -> occurrences -> freeze_evars_flag -> dep_proof_flag ->
- ?tac:(tactic * conditions) ->
- Id.t -> constr with_bindings -> evars_flag -> tactic
+ ?tac:(unit Proofview.tactic * conditions) ->
+ Id.t -> constr with_bindings -> evars_flag -> unit Proofview.tactic
val general_rewrite_in :
orientation -> occurrences -> freeze_evars_flag -> dep_proof_flag ->
- ?tac:(tactic * conditions) -> Id.t -> constr -> evars_flag -> tactic
+ ?tac:(unit Proofview.tactic * conditions) -> Id.t -> constr -> evars_flag -> unit Proofview.tactic
val general_multi_rewrite :
- orientation -> evars_flag -> ?tac:(tactic * conditions) -> constr with_bindings -> clause -> tactic
+ orientation -> evars_flag -> ?tac:(unit Proofview.tactic * conditions) -> constr with_bindings -> clause -> unit Proofview.tactic
type delayed_open_constr_with_bindings =
env -> evar_map -> evar_map * constr with_bindings
val general_multi_multi_rewrite :
evars_flag -> (bool * multi * delayed_open_constr_with_bindings) list ->
- clause -> (tactic * conditions) option -> tactic
-
-val replace_in_clause_maybe_by : constr -> constr -> clause -> tactic option -> tactic
-val replace : constr -> constr -> tactic
-val replace_in : Id.t -> constr -> constr -> tactic
-val replace_by : constr -> constr -> tactic -> tactic
-val replace_in_by : Id.t -> constr -> constr -> tactic -> tactic
-
-val discr : evars_flag -> constr with_bindings -> tactic
-val discrConcl : tactic
-val discrClause : evars_flag -> clause -> tactic
-val discrHyp : Id.t -> tactic
-val discrEverywhere : evars_flag -> tactic
+ clause -> (unit Proofview.tactic * conditions) option -> unit Proofview.tactic
+
+val replace_in_clause_maybe_by : constr -> constr -> clause -> unit Proofview.tactic option -> unit Proofview.tactic
+val replace : constr -> constr -> unit Proofview.tactic
+val replace_in : Id.t -> constr -> constr -> unit Proofview.tactic
+val replace_by : constr -> constr -> unit Proofview.tactic -> unit Proofview.tactic
+val replace_in_by : Id.t -> constr -> constr -> unit Proofview.tactic -> unit Proofview.tactic
+
+val discr : evars_flag -> constr with_bindings -> unit Proofview.tactic
+val discrConcl : unit Proofview.tactic
+val discrClause : evars_flag -> clause -> unit Proofview.tactic
+val discrHyp : Id.t -> unit Proofview.tactic
+val discrEverywhere : evars_flag -> unit Proofview.tactic
val discr_tac : evars_flag ->
- constr with_bindings induction_arg option -> tactic
+ constr with_bindings induction_arg option -> unit Proofview.tactic
val inj : intro_pattern_expr Loc.located list option -> evars_flag ->
- constr with_bindings -> tactic
+ constr with_bindings -> unit Proofview.tactic
val injClause : intro_pattern_expr Loc.located list option -> evars_flag ->
- constr with_bindings induction_arg option -> tactic
-val injHyp : Id.t -> tactic
-val injConcl : tactic
+ constr with_bindings induction_arg option -> unit Proofview.tactic
+val injHyp : Id.t -> unit Proofview.tactic
+val injConcl : unit Proofview.tactic
-val dEq : evars_flag -> constr with_bindings induction_arg option -> tactic
-val dEqThen : evars_flag -> (constr -> int -> tactic) -> constr with_bindings induction_arg option -> tactic
+val dEq : evars_flag -> constr with_bindings induction_arg option -> unit Proofview.tactic
+val dEqThen : evars_flag -> (constr -> int -> unit Proofview.tactic) -> constr with_bindings induction_arg option -> unit Proofview.tactic
val make_iterated_tuple :
env -> evar_map -> constr -> (constr * types) -> constr * constr * constr
@@ -126,15 +126,15 @@ type subst_tactic_flags = {
only_leibniz : bool;
rewrite_dependent_proof : bool
}
-val subst_gen : bool -> Id.t list -> tactic
-val subst : Id.t list -> tactic
-val subst_all : ?flags:subst_tactic_flags -> tactic
+val subst_gen : bool -> Id.t list -> unit Proofview.tactic
+val subst : Id.t list -> unit Proofview.tactic
+val subst_all : ?flags:subst_tactic_flags -> unit Proofview.tactic
(* Replace term *)
(* [replace_multi_term dir_opt c cl]
perfoms replacement of [c] by the first value found in context
(according to [dir] if given to get the rewrite direction) in the clause [cl]
*)
-val replace_multi_term : bool option -> constr -> clause -> tactic
+val replace_multi_term : bool option -> constr -> clause -> unit Proofview.tactic
val set_eq_dec_scheme_kind : mutual scheme_kind -> unit
diff --git a/tactics/evar_tactics.ml b/tactics/evar_tactics.ml
index ece0104b2..1234fe72b 100644
--- a/tactics/evar_tactics.ml
+++ b/tactics/evar_tactics.ml
@@ -50,8 +50,11 @@ let instantiate n (ist,rawc) ido gl =
tclNORMEVAR
gl
-let let_evar name typ gls =
+open Proofview.Notations
+let let_evar name typ =
let src = (Loc.ghost,Evar_kinds.GoalEvar) in
- let sigma',evar = Evarutil.new_evar gls.sigma (pf_env gls) ~src typ in
- Refiner.tclTHEN (Refiner.tclEVARS sigma')
- (Tactics.letin_tac None name evar None Locusops.nowhere) gls
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
+ let sigma',evar = Evarutil.new_evar sigma env ~src typ in
+ Tacticals.New.tclTHEN (Proofview.V82.tactic (Refiner.tclEVARS sigma'))
+ (Tactics.letin_tac None name evar None Locusops.nowhere)
diff --git a/tactics/evar_tactics.mli b/tactics/evar_tactics.mli
index b6ce3e30c..0e653e5b7 100644
--- a/tactics/evar_tactics.mli
+++ b/tactics/evar_tactics.mli
@@ -14,4 +14,4 @@ open Locus
val instantiate : int -> Tacinterp.interp_sign * Glob_term.glob_constr ->
(Id.t * hyp_location_flag, unit) location -> tactic
-val let_evar : Name.t -> Term.types -> tactic
+val let_evar : Name.t -> Term.types -> unit Proofview.tactic
diff --git a/tactics/extratactics.ml4 b/tactics/extratactics.ml4
index 32affcbe7..ae0bdfe44 100644
--- a/tactics/extratactics.ml4
+++ b/tactics/extratactics.ml4
@@ -31,13 +31,13 @@ TACTIC EXTEND admit
END
let replace_in_clause_maybe_by (sigma1,c1) c2 in_hyp tac =
- Refiner.tclWITHHOLES false
+ Tacticals.New.tclWITHHOLES false
(replace_in_clause_maybe_by c1 c2 (glob_in_arg_hyp_to_clause in_hyp))
sigma1
(Option.map Tacinterp.eval_tactic tac)
let replace_multi_term dir_opt (sigma,c) in_hyp =
- Refiner.tclWITHHOLES false
+ Tacticals.New.tclWITHHOLES false
(replace_multi_term dir_opt c)
sigma
(glob_in_arg_hyp_to_clause in_hyp)
@@ -49,12 +49,12 @@ END
TACTIC EXTEND replace_term_left
[ "replace" "->" open_constr(c) in_arg_hyp(in_hyp) ]
- -> [ replace_multi_term (Some true) c in_hyp]
+ -> [ replace_multi_term (Some true) c in_hyp ]
END
TACTIC EXTEND replace_term_right
[ "replace" "<-" open_constr(c) in_arg_hyp(in_hyp) ]
- -> [replace_multi_term (Some false) c in_hyp]
+ -> [ replace_multi_term (Some false) c in_hyp ]
END
TACTIC EXTEND replace_term
@@ -71,7 +71,7 @@ let induction_arg_of_quantified_hyp = function
ElimOnIdent and not as "constr" *)
let elimOnConstrWithHoles tac with_evars c =
- Refiner.tclWITHHOLES with_evars (tac with_evars)
+ Tacticals.New.tclWITHHOLES with_evars (tac with_evars)
c.sigma (Some (ElimOnConstr c.it))
TACTIC EXTEND simplify_eq_main
@@ -114,8 +114,10 @@ TACTIC EXTEND ediscriminate
[ discr_tac true (Some (induction_arg_of_quantified_hyp h)) ]
END
-let discrHyp id gl =
- discr_main {it = Term.mkVar id,NoBindings; sigma = Refiner.project gl;} gl
+open Proofview.Notations
+let discrHyp id =
+ Goal.defs >>- fun sigma ->
+ discr_main {it = Term.mkVar id,NoBindings; sigma = sigma;}
let injection_main c =
elimOnConstrWithHoles (injClause None) false c
@@ -158,19 +160,20 @@ TACTIC EXTEND einjection_as
[ injClause (Some ipat) true (Some (induction_arg_of_quantified_hyp h)) ]
END
-let injHyp id gl =
- injection_main { it = Term.mkVar id,NoBindings; sigma = Refiner.project gl; } gl
+let injHyp id =
+ Goal.defs >>- fun sigma ->
+ injection_main { it = Term.mkVar id,NoBindings; sigma = sigma; }
TACTIC EXTEND dependent_rewrite
-| [ "dependent" "rewrite" orient(b) constr(c) ] -> [ rewriteInConcl b c ]
+| [ "dependent" "rewrite" orient(b) constr(c) ] -> [ Proofview.V82.tactic (rewriteInConcl b c) ]
| [ "dependent" "rewrite" orient(b) constr(c) "in" hyp(id) ]
- -> [ rewriteInHyp b c id ]
+ -> [ Proofview.V82.tactic (rewriteInHyp b c id) ]
END
TACTIC EXTEND cut_rewrite
-| [ "cutrewrite" orient(b) constr(eqn) ] -> [ cutRewriteInConcl b eqn ]
+| [ "cutrewrite" orient(b) constr(eqn) ] -> [ Proofview.V82.tactic (cutRewriteInConcl b eqn) ]
| [ "cutrewrite" orient(b) constr(eqn) "in" hyp(id) ]
- -> [ cutRewriteInHyp b eqn id ]
+ -> [ Proofview.V82.tactic (cutRewriteInHyp b eqn id) ]
END
(**********************************************************************)
@@ -184,7 +187,7 @@ END
let onSomeWithHoles tac = function
| None -> tac None
- | Some c -> Refiner.tclWITHHOLES false tac c.sigma (Some c.it)
+ | Some c -> Tacticals.New.tclWITHHOLES false tac c.sigma (Some c.it)
TACTIC EXTEND contradiction
[ "contradiction" constr_with_bindings_opt(c) ] ->
@@ -231,7 +234,7 @@ END
let rewrite_star clause orient occs (sigma,c) (tac : glob_tactic_expr option) =
let tac' = Option.map (fun t -> Tacinterp.eval_tactic t, FirstSolved) tac in
- Refiner. tclWITHHOLES false
+ Tacticals.New. tclWITHHOLES false
(general_rewrite_ebindings_clause clause orient occs ?tac:tac' true true (c,NoBindings)) sigma true
TACTIC EXTEND rewrite_star
@@ -383,7 +386,7 @@ END
TACTIC EXTEND subst
| [ "subst" ne_var_list(l) ] -> [ subst l ]
-| [ "subst" ] -> [ fun gl -> subst_all gl ]
+| [ "subst" ] -> [ subst_all ?flags:None ]
END
let simple_subst_tactic_flags =
@@ -407,8 +410,8 @@ open Tacticals
TACTIC EXTEND instantiate
[ "instantiate" "(" integer(i) ":=" glob(c) ")" hloc(hl) ] ->
- [instantiate i c hl ]
-| [ "instantiate" ] -> [ tclNORMEVAR ]
+ [ Proofview.V82.tactic (instantiate i c hl) ]
+| [ "instantiate" ] -> [ Proofview.V82.tactic (tclNORMEVAR) ]
END
@@ -434,12 +437,12 @@ let transitivity_left_table = Summary.ref [] ~name:"transitivity-steps-l"
let step left x tac =
let l =
List.map (fun lem ->
- tclTHENLAST
- (apply_with_bindings (lem, ImplicitBindings [x]))
+ Tacticals.New.tclTHENLAST
+ (Proofview.V82.tactic (apply_with_bindings (lem, ImplicitBindings [x])))
tac)
!(if left then transitivity_left_table else transitivity_right_table)
in
- tclFIRST l
+ Tacticals.New.tclFIRST l
(* Main function to push lemmas in persistent environment *)
@@ -468,12 +471,12 @@ let add_transitivity_lemma left lem =
TACTIC EXTEND stepl
| ["stepl" constr(c) "by" tactic(tac) ] -> [ step true c (Tacinterp.eval_tactic tac) ]
-| ["stepl" constr(c) ] -> [ step true c tclIDTAC ]
+| ["stepl" constr(c) ] -> [ step true c (Proofview.tclUNIT ()) ]
END
TACTIC EXTEND stepr
| ["stepr" constr(c) "by" tactic(tac) ] -> [ step false c (Tacinterp.eval_tactic tac) ]
-| ["stepr" constr(c) ] -> [ step false c tclIDTAC ]
+| ["stepr" constr(c) ] -> [ step false c (Proofview.tclUNIT ()) ]
END
VERNAC COMMAND EXTEND AddStepl CLASSIFIED AS SIDEFF
@@ -527,7 +530,7 @@ END
during dependent induction. For internal use. *)
TACTIC EXTEND specialize_eqs
-[ "specialize_eqs" hyp(id) ] -> [ specialize_eqs id ]
+[ "specialize_eqs" hyp(id) ] -> [ Proofview.V82.tactic (specialize_eqs id) ]
END
(**********************************************************************)
@@ -608,8 +611,8 @@ let hResolve_auto id c t gl =
resolve_auto 1
TACTIC EXTEND hresolve_core
-| [ "hresolve_core" "(" ident(id) ":=" constr(c) ")" "at" int_or_var(occ) "in" constr(t) ] -> [ hResolve id c (out_arg occ) t ]
-| [ "hresolve_core" "(" ident(id) ":=" constr(c) ")" "in" constr(t) ] -> [ hResolve_auto id c t ]
+| [ "hresolve_core" "(" ident(id) ":=" constr(c) ")" "at" int_or_var(occ) "in" constr(t) ] -> [ Proofview.V82.tactic (hResolve id c (out_arg occ) t) ]
+| [ "hresolve_core" "(" ident(id) ":=" constr(c) ")" "in" constr(t) ] -> [ Proofview.V82.tactic (hResolve_auto id c t) ]
END
(**
@@ -627,7 +630,7 @@ let hget_evar n gl =
change_in_concl None (mkLetIn (Anonymous,mkEvar ev,ev_type,pf_concl gl)) gl
TACTIC EXTEND hget_evar
-| [ "hget_evar" int_or_var(n) ] -> [ hget_evar (out_arg n) ]
+| [ "hget_evar" int_or_var(n) ] -> [ Proofview.V82.tactic (hget_evar (out_arg n)) ]
END
(**********************************************************************)
@@ -640,12 +643,13 @@ END
(* Contributed by Julien Forest and Pierre Courtieu (july 2010) *)
(**********************************************************************)
-exception Found of tactic
+exception Found of unit Proofview.tactic
-let rewrite_except h g =
- tclMAP (fun id -> if Id.equal id h then tclIDTAC else
- tclTRY (Equality.general_rewrite_in true Locus.AllOccurrences true true id (mkVar h) false))
- (Tacmach.pf_ids_of_hyps g) g
+let rewrite_except h =
+ Tacmach.New.pf_ids_of_hyps >>- fun hyps ->
+ Tacticals.New.tclMAP (fun id -> if Id.equal id h then Proofview.tclUNIT () else
+ Tacticals.New.tclTRY (Equality.general_rewrite_in true Locus.AllOccurrences true true id (mkVar h) false))
+ hyps
let refl_equal =
@@ -658,31 +662,40 @@ let refl_equal =
(* This is simply an implementation of the case_eq tactic. this code
should be replaced by a call to the tactic but I don't know how to
call it before it is defined. *)
-let mkCaseEq a : tactic =
- (fun g ->
- let type_of_a = Tacmach.pf_type_of g a in
- tclTHENLIST
- [Hiddentac.h_generalize [mkApp(delayed_force refl_equal, [| type_of_a; a|])];
- (fun g2 ->
- change_in_concl None
- (Tacred.pattern_occs [Locus.OnlyOccurrences [1], a] (Tacmach.pf_env g2) Evd.empty (Tacmach.pf_concl g2))
- g2);
- simplest_case a] g);;
-
-
-let case_eq_intros_rewrite x g =
- let n = nb_prod (Tacmach.pf_concl g) in
+let mkCaseEq a : unit Proofview.tactic =
+ Tacmach.New.of_old (fun g -> Tacmach.pf_type_of g a) >>- fun type_of_a ->
+ Tacticals.New.tclTHENLIST
+ [Proofview.V82.tactic (Hiddentac.h_generalize [mkApp(delayed_force refl_equal, [| type_of_a; a|])]);
+ begin
+ Goal.concl >>- fun concl ->
+ Goal.env >>- fun env ->
+ Proofview.V82.tactic begin
+ change_in_concl None
+ (Tacred.pattern_occs [Locus.OnlyOccurrences [1], a] env Evd.empty concl)
+ end
+ end;
+ simplest_case a]
+
+
+let case_eq_intros_rewrite x =
+ begin
+ Goal.concl >- fun concl ->
+ Goal.return (nb_prod concl)
+ end >>- fun n ->
(* Pp.msgnl (Printer.pr_lconstr x); *)
- tclTHENLIST [
+ Tacticals.New.tclTHENLIST [
mkCaseEq x;
- (fun g ->
- let n' = nb_prod (Tacmach.pf_concl g) in
- let h = fresh_id (Tacmach.pf_ids_of_hyps g) (Id.of_string "heq") g in
- tclTHENLIST [ (tclDO (n'-n-1) intro);
- Tacmach.introduction h;
- rewrite_except h] g
- )
- ] g
+ begin
+ Goal.concl >>- fun concl ->
+ Tacmach.New.pf_ids_of_hyps >>- fun hyps ->
+ let n' = nb_prod concl in
+ Tacmach.New.of_old (fun g -> fresh_id hyps (Id.of_string "heq") g) >>- fun h ->
+ Tacticals.New.tclTHENLIST [
+ Tacticals.New.tclDO (n'-n-1) intro;
+ Proofview.V82.tactic (Tacmach.introduction h);
+ rewrite_except h]
+ end
+ ]
let rec find_a_destructable_match t =
match kind_of_term t with
@@ -698,17 +711,17 @@ let rec find_a_destructable_match t =
let destauto t =
try find_a_destructable_match t;
- error "No destructable match found"
+ Proofview.tclZERO (UserError ("", str"No destructable match found"))
with Found tac -> tac
-let destauto_in id g =
- let ctype = Tacmach.pf_type_of g (mkVar id) in
+let destauto_in id =
+ Tacmach.New.of_old (fun g -> Tacmach.pf_type_of g (mkVar id)) >>- fun ctype ->
(* Pp.msgnl (Printer.pr_lconstr (mkVar id)); *)
(* Pp.msgnl (Printer.pr_lconstr (ctype)); *)
- destauto ctype g
+ destauto ctype
TACTIC EXTEND destauto
-| [ "destauto" ] -> [ (fun g -> destauto (Tacmach.pf_concl g) g) ]
+| [ "destauto" ] -> [ Goal.concl >>- fun concl -> destauto concl ]
| [ "destauto" "in" hyp(id) ] -> [ destauto_in id ]
END
@@ -717,14 +730,14 @@ END
TACTIC EXTEND constr_eq
| [ "constr_eq" constr(x) constr(y) ] -> [
- if eq_constr x y then tclIDTAC else tclFAIL 0 (str "Not equal") ]
+ if eq_constr x y then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (str "Not equal") ]
END
TACTIC EXTEND is_evar
| [ "is_evar" constr(x) ] ->
[ match kind_of_term x with
- | Evar _ -> tclIDTAC
- | _ -> tclFAIL 0 (str "Not an evar")
+ | Evar _ -> Proofview.tclUNIT ()
+ | _ -> Tacticals.New.tclFAIL 0 (str "Not an evar")
]
END
@@ -750,21 +763,21 @@ and has_evar_prec (_, ts1, ts2) =
TACTIC EXTEND has_evar
| [ "has_evar" constr(x) ] ->
- [ if has_evar x then tclIDTAC else tclFAIL 0 (str "No evars") ]
+ [ if has_evar x then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (str "No evars") ]
END
TACTIC EXTEND is_hyp
| [ "is_var" constr(x) ] ->
[ match kind_of_term x with
- | Var _ -> tclIDTAC
- | _ -> tclFAIL 0 (str "Not a variable or hypothesis") ]
+ | Var _ -> Proofview.tclUNIT ()
+ | _ -> Tacticals.New.tclFAIL 0 (str "Not a variable or hypothesis") ]
END
TACTIC EXTEND is_fix
| [ "is_fix" constr(x) ] ->
[ match kind_of_term x with
- | Fix _ -> Tacticals.tclIDTAC
- | _ -> Tacticals.tclFAIL 0 (Pp.str "not a fix definition") ]
+ | Fix _ -> Proofview.tclUNIT ()
+ | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a fix definition") ]
END;;
(* Command to grab the evars left unresolved at the end of a proof. *)
diff --git a/tactics/extratactics.mli b/tactics/extratactics.mli
index 306067ff0..7a0fabe1f 100644
--- a/tactics/extratactics.mli
+++ b/tactics/extratactics.mli
@@ -8,9 +8,9 @@
open Proof_type
-val discrHyp : Names.Id.t -> tactic
-val injHyp : Names.Id.t -> tactic
+val discrHyp : Names.Id.t -> unit Proofview.tactic
+val injHyp : Names.Id.t -> unit Proofview.tactic
-val refine_tac : Evd.open_constr -> tactic
+val refine_tac : Evd.open_constr -> unit Proofview.tactic
-val onSomeWithHoles : ('a option -> tactic) -> 'a Evd.sigma option -> tactic
+val onSomeWithHoles : ('a option -> unit Proofview.tactic) -> 'a Evd.sigma option -> unit Proofview.tactic
diff --git a/tactics/g_class.ml4 b/tactics/g_class.ml4
index 970f635b7..c7867a83c 100644
--- a/tactics/g_class.ml4
+++ b/tactics/g_class.ml4
@@ -63,8 +63,8 @@ VERNAC COMMAND EXTEND Typeclasses_Settings CLASSIFIED AS SIDEFF
END
TACTIC EXTEND typeclasses_eauto
-| [ "typeclasses" "eauto" "with" ne_preident_list(l) ] -> [ typeclasses_eauto l ]
-| [ "typeclasses" "eauto" ] -> [ typeclasses_eauto ~only_classes:true [Auto.typeclasses_db] ]
+| [ "typeclasses" "eauto" "with" ne_preident_list(l) ] -> [ Proofview.V82.tactic (typeclasses_eauto l) ]
+| [ "typeclasses" "eauto" ] -> [ Proofview.V82.tactic (typeclasses_eauto ~only_classes:true [Auto.typeclasses_db]) ]
END
TACTIC EXTEND head_of_constr
@@ -76,9 +76,9 @@ TACTIC EXTEND not_evar
END
TACTIC EXTEND is_ground
- [ "is_ground" constr(ty) ] -> [ is_ground ty ]
+ [ "is_ground" constr(ty) ] -> [ Proofview.V82.tactic (is_ground ty) ]
END
TACTIC EXTEND autoapply
- [ "autoapply" constr(c) "using" preident(i) ] -> [ autoapply c i ]
+ [ "autoapply" constr(c) "using" preident(i) ] -> [ Proofview.V82.tactic (autoapply c i) ]
END
diff --git a/tactics/g_rewrite.ml4 b/tactics/g_rewrite.ml4
index f4c5aeda8..f889f9cb2 100644
--- a/tactics/g_rewrite.ml4
+++ b/tactics/g_rewrite.ml4
@@ -112,10 +112,10 @@ let db_strat db = StratUnary ("topdown", StratHints (false, db))
let cl_rewrite_clause_db db = cl_rewrite_clause_strat (strategy_of_ast (db_strat db))
TACTIC EXTEND rewrite_strat
-| [ "rewrite_strat" rewstrategy(s) "in" hyp(id) ] -> [ cl_rewrite_clause_strat s (Some id) ]
-| [ "rewrite_strat" rewstrategy(s) ] -> [ cl_rewrite_clause_strat s None ]
-| [ "rewrite_db" preident(db) "in" hyp(id) ] -> [ cl_rewrite_clause_db db (Some id) ]
-| [ "rewrite_db" preident(db) ] -> [ cl_rewrite_clause_db db None ]
+| [ "rewrite_strat" rewstrategy(s) "in" hyp(id) ] -> [ Proofview.V82.tactic (cl_rewrite_clause_strat s (Some id)) ]
+| [ "rewrite_strat" rewstrategy(s) ] -> [ Proofview.V82.tactic (cl_rewrite_clause_strat s None) ]
+| [ "rewrite_db" preident(db) "in" hyp(id) ] -> [ Proofview.V82.tactic (cl_rewrite_clause_db db (Some id)) ]
+| [ "rewrite_db" preident(db) ] -> [ Proofview.V82.tactic (cl_rewrite_clause_db db None) ]
END
let clsubstitute o c =
@@ -126,9 +126,8 @@ let clsubstitute o c =
| Some id when is_tac id -> tclIDTAC
| _ -> cl_rewrite_clause c o AllOccurrences cl)
-
TACTIC EXTEND substitute
-| [ "substitute" orient(o) glob_constr_with_bindings(c) ] -> [ clsubstitute o c ]
+| [ "substitute" orient(o) glob_constr_with_bindings(c) ] -> [ Proofview.V82.tactic (clsubstitute o c) ]
END
@@ -136,15 +135,15 @@ END
TACTIC EXTEND setoid_rewrite
[ "setoid_rewrite" orient(o) glob_constr_with_bindings(c) ]
- -> [ cl_rewrite_clause c o AllOccurrences None ]
+ -> [ Proofview.V82.tactic (cl_rewrite_clause c o AllOccurrences None) ]
| [ "setoid_rewrite" orient(o) glob_constr_with_bindings(c) "in" hyp(id) ] ->
- [ cl_rewrite_clause c o AllOccurrences (Some id)]
+ [ Proofview.V82.tactic (cl_rewrite_clause c o AllOccurrences (Some id))]
| [ "setoid_rewrite" orient(o) glob_constr_with_bindings(c) "at" occurrences(occ) ] ->
- [ cl_rewrite_clause c o (occurrences_of occ) None]
+ [ Proofview.V82.tactic (cl_rewrite_clause c o (occurrences_of occ) None)]
| [ "setoid_rewrite" orient(o) glob_constr_with_bindings(c) "at" occurrences(occ) "in" hyp(id)] ->
- [ cl_rewrite_clause c o (occurrences_of occ) (Some id)]
+ [ Proofview.V82.tactic (cl_rewrite_clause c o (occurrences_of occ) (Some id))]
| [ "setoid_rewrite" orient(o) glob_constr_with_bindings(c) "in" hyp(id) "at" occurrences(occ)] ->
- [ cl_rewrite_clause c o (occurrences_of occ) (Some id)]
+ [ Proofview.V82.tactic (cl_rewrite_clause c o (occurrences_of occ) (Some id))]
END
let cl_rewrite_clause_newtac_tac c o occ cl gl =
@@ -153,15 +152,15 @@ let cl_rewrite_clause_newtac_tac c o occ cl gl =
TACTIC EXTEND GenRew
| [ "rew" orient(o) glob_constr_with_bindings(c) "in" hyp(id) "at" occurrences(occ) ] ->
- [ cl_rewrite_clause_newtac_tac c o (occurrences_of occ) (Some id) ]
+ [ Proofview.V82.tactic (cl_rewrite_clause_newtac_tac c o (occurrences_of occ) (Some id)) ]
| [ "rew" orient(o) glob_constr_with_bindings(c) "at" occurrences(occ) "in" hyp(id) ] ->
- [ cl_rewrite_clause_newtac_tac c o (occurrences_of occ) (Some id) ]
+ [ Proofview.V82.tactic (cl_rewrite_clause_newtac_tac c o (occurrences_of occ) (Some id)) ]
| [ "rew" orient(o) glob_constr_with_bindings(c) "in" hyp(id) ] ->
- [ cl_rewrite_clause_newtac_tac c o AllOccurrences (Some id) ]
+ [ Proofview.V82.tactic (cl_rewrite_clause_newtac_tac c o AllOccurrences (Some id)) ]
| [ "rew" orient(o) glob_constr_with_bindings(c) "at" occurrences(occ) ] ->
- [ cl_rewrite_clause_newtac_tac c o (occurrences_of occ) None ]
+ [ Proofview.V82.tactic (cl_rewrite_clause_newtac_tac c o (occurrences_of occ) None) ]
| [ "rew" orient(o) glob_constr_with_bindings(c) ] ->
- [ cl_rewrite_clause_newtac_tac c o AllOccurrences None ]
+ [ Proofview.V82.tactic (cl_rewrite_clause_newtac_tac c o AllOccurrences None) ]
END
VERNAC COMMAND EXTEND AddRelation CLASSIFIED AS SIDEFF
@@ -269,17 +268,17 @@ TACTIC EXTEND setoid_transitivity
END
TACTIC EXTEND implify
-[ "implify" hyp(n) ] -> [ implify n ]
+[ "implify" hyp(n) ] -> [ Proofview.V82.tactic (implify n) ]
END
TACTIC EXTEND fold_match
-[ "fold_match" constr(c) ] -> [ fold_match_tac c ]
+[ "fold_match" constr(c) ] -> [ Proofview.V82.tactic (fold_match_tac c) ]
END
TACTIC EXTEND fold_matches
-| [ "fold_matches" constr(c) ] -> [ fold_matches_tac c ]
+| [ "fold_matches" constr(c) ] -> [ Proofview.V82.tactic (fold_matches_tac c) ]
END
TACTIC EXTEND myapply
-| [ "myapply" global(id) constr_list(l) ] -> [ myapply id l ]
+| [ "myapply" global(id) constr_list(l) ] -> [ Proofview.V82.tactic (myapply id l) ]
END
diff --git a/tactics/hiddentac.mli b/tactics/hiddentac.mli
index b4491770e..f81698370 100644
--- a/tactics/hiddentac.mli
+++ b/tactics/hiddentac.mli
@@ -25,9 +25,9 @@ open Misctypes
(** Basic tactics *)
-val h_intro_move : Id.t option -> Id.t move_location -> tactic
-val h_intro : Id.t -> tactic
-val h_intros_until : quantified_hypothesis -> tactic
+val h_intro_move : Id.t option -> Id.t move_location -> unit Proofview.tactic
+val h_intro : Id.t -> unit Proofview.tactic
+val h_intros_until : quantified_hypothesis -> unit Proofview.tactic
val h_assumption : tactic
val h_exact : constr -> tactic
@@ -38,12 +38,12 @@ val h_apply : advanced_flag -> evars_flag ->
constr with_bindings located list -> tactic
val h_apply_in : advanced_flag -> evars_flag ->
constr with_bindings located list ->
- Id.t * intro_pattern_expr located option -> tactic
+ Id.t * intro_pattern_expr located option -> unit Proofview.tactic
val h_elim : evars_flag -> constr with_bindings ->
- constr with_bindings option -> tactic
+ constr with_bindings option -> unit Proofview.tactic
val h_elim_type : constr -> tactic
-val h_case : evars_flag -> constr with_bindings -> tactic
+val h_case : evars_flag -> constr with_bindings -> unit Proofview.tactic
val h_case_type : constr -> tactic
val h_mutual_fix : Id.t -> int ->
@@ -57,31 +57,31 @@ val h_generalize : constr list -> tactic
val h_generalize_gen : (constr Locus.with_occurrences * Name.t) list -> tactic
val h_generalize_dep : ?with_let:bool -> constr -> tactic
val h_let_tac : letin_flag -> Name.t -> constr -> Locus.clause ->
- intro_pattern_expr located option -> tactic
+ intro_pattern_expr located option -> unit Proofview.tactic
val h_let_pat_tac : letin_flag -> Name.t -> evar_map * constr ->
Locus.clause -> intro_pattern_expr located option ->
- tactic
+ unit Proofview.tactic
(** Derived basic tactics *)
-val h_simple_induction : quantified_hypothesis -> tactic
-val h_simple_destruct : quantified_hypothesis -> tactic
-val h_simple_induction_destruct : rec_flag -> quantified_hypothesis -> tactic
+val h_simple_induction : quantified_hypothesis -> unit Proofview.tactic
+val h_simple_destruct : quantified_hypothesis -> unit Proofview.tactic
+val h_simple_induction_destruct : rec_flag -> quantified_hypothesis -> unit Proofview.tactic
val h_new_induction : evars_flag ->
(evar_map * constr with_bindings) induction_arg ->
intro_pattern_expr located option * intro_pattern_expr located option ->
constr with_bindings option ->
- Locus.clause option -> tactic
+ Locus.clause option -> unit Proofview.tactic
val h_new_destruct : evars_flag ->
(evar_map * constr with_bindings) induction_arg ->
intro_pattern_expr located option * intro_pattern_expr located option ->
constr with_bindings option ->
- Locus.clause option -> tactic
+ Locus.clause option -> unit Proofview.tactic
val h_induction_destruct : rec_flag -> evars_flag ->
((evar_map * constr with_bindings) induction_arg *
(intro_pattern_expr located option * intro_pattern_expr located option)) list
* constr with_bindings option
- * Locus.clause option -> tactic
+ * Locus.clause option -> unit Proofview.tactic
val h_specialize : int option -> constr with_bindings -> tactic
val h_lapply : constr -> tactic
@@ -97,14 +97,14 @@ val h_rename : (Id.t*Id.t) list -> tactic
val h_revert : Id.t list -> tactic
(** Constructors *)
-val h_constructor : evars_flag -> int -> constr bindings -> tactic
-val h_left : evars_flag -> constr bindings -> tactic
-val h_right : evars_flag -> constr bindings -> tactic
-val h_split : evars_flag -> constr bindings list -> tactic
+val h_constructor : evars_flag -> int -> constr bindings -> unit Proofview.tactic
+val h_left : evars_flag -> constr bindings -> unit Proofview.tactic
+val h_right : evars_flag -> constr bindings -> unit Proofview.tactic
+val h_split : evars_flag -> constr bindings list -> unit Proofview.tactic
-val h_one_constructor : int -> tactic
-val h_simplest_left : tactic
-val h_simplest_right : tactic
+val h_one_constructor : int -> unit Proofview.tactic
+val h_simplest_left : unit Proofview.tactic
+val h_simplest_right : unit Proofview.tactic
(** Conversion *)
@@ -113,13 +113,13 @@ val h_change :
Pattern.constr_pattern option -> constr -> Locus.clause -> tactic
(** Equivalence relations *)
-val h_reflexivity : tactic
-val h_symmetry : Locus.clause -> tactic
-val h_transitivity : constr option -> tactic
+val h_reflexivity : unit Proofview.tactic
+val h_symmetry : Locus.clause -> unit Proofview.tactic
+val h_transitivity : constr option -> unit Proofview.tactic
val h_simplest_apply : constr -> tactic
val h_simplest_eapply : constr -> tactic
-val h_simplest_elim : constr -> tactic
-val h_simplest_case : constr -> tactic
+val h_simplest_elim : constr -> unit Proofview.tactic
+val h_simplest_case : constr -> unit Proofview.tactic
-val h_intro_patterns : intro_pattern_expr located list -> tactic
+val h_intro_patterns : intro_pattern_expr located list -> unit Proofview.tactic
diff --git a/tactics/inv.ml b/tactics/inv.ml
index ec96a887d..c9d54f84f 100644
--- a/tactics/inv.ml
+++ b/tactics/inv.ml
@@ -28,6 +28,7 @@ open Elim
open Equality
open Misctypes
open Tacexpr
+open Proofview.Notations
let collect_meta_variables c =
let rec collrec acc c = match kind_of_term c with
@@ -272,19 +273,18 @@ Summary: nine useless hypotheses!
Nota: with Inversion_clear, only four useless hypotheses
*)
-let generalizeRewriteIntros tac depids id gls =
- let dids = dependent_hyps id depids gls in
- (tclTHENSEQ
- [bring_hyps dids; tac;
+let generalizeRewriteIntros tac depids id =
+ Tacmach.New.of_old (dependent_hyps id depids) >>- fun dids ->
+ (Tacticals.New.tclTHENLIST
+ [Proofview.V82.tactic (bring_hyps dids); tac;
(* may actually fail to replace if dependent in a previous eq *)
intros_replacing (ids_of_named_context dids)])
- gls
let rec tclMAP_i n tacfun = function
- | [] -> tclDO n (tacfun None)
+ | [] -> Tacticals.New.tclDO n (tacfun None)
| a::l ->
- if Int.equal n 0 then error "Too many names."
- else tclTHEN (tacfun (Some a)) (tclMAP_i (n-1) tacfun l)
+ if Int.equal n 0 then Proofview.tclZERO (Errors.UserError ("", Pp.str"Too many names."))
+ else Tacticals.New.tclTHEN (tacfun (Some a)) (tclMAP_i (n-1) tacfun l)
let remember_first_eq id x = if !x == MoveLast then x := MoveAfter id
@@ -296,29 +296,29 @@ let remember_first_eq id x = if !x == MoveLast then x := MoveAfter id
If it can discriminate then the goal is proved, if not tries to use it as
a rewrite rule. It erases the clause which is given as input *)
-let projectAndApply thin id eqname names depids gls =
+let projectAndApply thin id eqname names depids =
let subst_hyp l2r id =
tclTHEN (tclTRY(rewriteInConcl l2r (mkVar id)))
(if thin then clear [id] else (remember_first_eq id eqname; tclIDTAC))
in
- let substHypIfVariable tac id gls =
- let (t,t1,t2) = Hipattern.dest_nf_eq gls (pf_get_hyp_typ gls id) in
+ let substHypIfVariable tac id =
+ Tacmach.New.of_old (fun gls -> Hipattern.dest_nf_eq gls (pf_get_hyp_typ gls id)) >>- fun (t,t1,t2) ->
match (kind_of_term t1, kind_of_term t2) with
- | Var id1, _ -> generalizeRewriteIntros (subst_hyp true id) depids id1 gls
- | _, Var id2 -> generalizeRewriteIntros (subst_hyp false id) depids id2 gls
- | _ -> tac id gls
+ | Var id1, _ -> generalizeRewriteIntros (Proofview.V82.tactic (subst_hyp true id)) depids id1
+ | _, Var id2 -> generalizeRewriteIntros (Proofview.V82.tactic (subst_hyp false id)) depids id2
+ | _ -> tac id
in
let deq_trailer id _ neqns =
- tclTHENSEQ
- [(if not (List.is_empty names) then clear [id] else tclIDTAC);
+ Tacticals.New.tclTHENLIST
+ [Proofview.V82.tactic (if not (List.is_empty names) then clear [id] else tclIDTAC);
(tclMAP_i neqns (fun idopt ->
- tclTRY (tclTHEN
+ Tacticals.New.tclTRY (Tacticals.New.tclTHEN
(intro_move idopt MoveLast)
(* try again to substitute and if still not a variable after *)
(* decomposition, arbitrarily try to rewrite RL !? *)
- (tclTRY (onLastHypId (substHypIfVariable (subst_hyp false))))))
+ (Tacticals.New.tclTRY (Tacticals.New.onLastHypId (substHypIfVariable (fun id -> Proofview.V82.tactic (subst_hyp false id)))))))
names);
- (if List.is_empty names then clear [id] else tclIDTAC)]
+ Proofview.V82.tactic (if List.is_empty names then clear [id] else tclIDTAC)]
in
substHypIfVariable
(* If no immediate variable in the equation, try to decompose it *)
@@ -327,42 +327,40 @@ let projectAndApply thin id eqname names depids gls =
dEqThen false (deq_trailer id)
(Some (ElimOnConstr (mkVar id,NoBindings))))
id
- gls
(* Inversion qui n'introduit pas les hypotheses, afin de pouvoir les nommer
soi-meme (proposition de Valerie). *)
-let rewrite_equations_gene othin neqns ba gl =
- let (depids,nodepids) = split_dep_and_nodep ba.assums gl in
+let rewrite_equations_gene othin neqns ba =
+ Tacmach.New.of_old (fun gl -> split_dep_and_nodep ba.assums gl) >>- fun (depids,nodepids) ->
let rewrite_eqns =
match othin with
| Some thin ->
- onLastHypId
+ Tacticals.New.onLastHypId
(fun last ->
- tclTHENSEQ
- [tclDO neqns
- (tclTHEN intro
- (onLastHypId
+ Tacticals.New.tclTHENLIST
+ [Tacticals.New.tclDO neqns
+ (Tacticals.New.tclTHEN intro
+ (Tacticals.New.onLastHypId
(fun id ->
- tclTRY
+ Tacticals.New.tclTRY
(projectAndApply thin id (ref MoveLast)
[] depids))));
- onHyps (compose List.rev (afterHyp last)) bring_hyps;
- onHyps (afterHyp last)
- (compose clear ids_of_named_context)])
- | None -> tclIDTAC
+ Proofview.V82.tactic (onHyps (compose List.rev (afterHyp last)) bring_hyps);
+ Proofview.V82.tactic (onHyps (afterHyp last)
+ (compose clear ids_of_named_context))])
+ | None -> Proofview.tclUNIT ()
in
- (tclTHENSEQ
- [tclDO neqns intro;
- bring_hyps nodepids;
- clear (ids_of_named_context nodepids);
- onHyps (compose List.rev (nLastDecls neqns)) bring_hyps;
- onHyps (nLastDecls neqns) (compose clear ids_of_named_context);
+ (Tacticals.New.tclTHENLIST
+ [Tacticals.New.tclDO neqns intro;
+ Proofview.V82.tactic (bring_hyps nodepids);
+ Proofview.V82.tactic (clear (ids_of_named_context nodepids));
+ Proofview.V82.tactic (onHyps (compose List.rev (nLastDecls neqns)) bring_hyps);
+ Proofview.V82.tactic (onHyps (nLastDecls neqns) (compose clear ids_of_named_context));
rewrite_eqns;
- tclMAP (fun (id,_,_ as d) ->
+ Proofview.V82.tactic (tclMAP (fun (id,_,_ as d) ->
(tclORELSE (clear [id])
(tclTHEN (bring_hyps [d]) (clear [id]))))
- depids])
- gl
+ depids)])
(* Introduction of the equations on arguments
othin: discriminates Simple Inversion, Inversion and Inversion_clear
@@ -398,35 +396,34 @@ let extract_eqn_names = function
| None -> None,[]
| Some x -> x
-let rewrite_equations othin neqns names ba gl =
+let rewrite_equations othin neqns names ba =
let names = List.map (get_names true) names in
- let (depids,nodepids) = split_dep_and_nodep ba.assums gl in
+ Tacmach.New.of_old (fun gl -> split_dep_and_nodep ba.assums gl) >>- fun (depids,nodepids) ->
let rewrite_eqns =
let first_eq = ref MoveLast in
match othin with
| Some thin ->
- tclTHENSEQ
- [onHyps (compose List.rev (nLastDecls neqns)) bring_hyps;
- onHyps (nLastDecls neqns) (compose clear ids_of_named_context);
+ Tacticals.New.tclTHENLIST
+ [Proofview.V82.tactic (onHyps (compose List.rev (nLastDecls neqns)) bring_hyps);
+ Proofview.V82.tactic (onHyps (nLastDecls neqns) (compose clear ids_of_named_context));
tclMAP_i neqns (fun o ->
let idopt,names = extract_eqn_names o in
- (tclTHEN
+ (Tacticals.New.tclTHEN
(intro_move idopt MoveLast)
- (onLastHypId (fun id ->
- tclTRY (projectAndApply thin id first_eq names depids)))))
+ (Tacticals.New.onLastHypId (fun id ->
+ Tacticals.New.tclTRY (projectAndApply thin id first_eq names depids)))))
names;
- tclMAP (fun (id,_,_) gl ->
- intro_move None (if thin then MoveLast else !first_eq) gl)
+ Tacticals.New.tclMAP (fun (id,_,_) -> Proofview.tclUNIT () >= fun () -> (* delay for [first_eq]. *)
+ intro_move None (if thin then MoveLast else !first_eq))
nodepids;
- tclMAP (fun (id,_,_) -> tclTRY (clear [id])) depids]
- | None -> tclIDTAC
+ Proofview.V82.tactic (tclMAP (fun (id,_,_) -> tclTRY (clear [id])) depids)]
+ | None -> Proofview.tclUNIT ()
in
- (tclTHENSEQ
- [tclDO neqns intro;
- bring_hyps nodepids;
- clear (ids_of_named_context nodepids);
+ (Tacticals.New.tclTHENLIST
+ [Tacticals.New.tclDO neqns intro;
+ Proofview.V82.tactic (bring_hyps nodepids);
+ Proofview.V82.tactic (clear (ids_of_named_context nodepids));
rewrite_eqns])
- gl
let interp_inversion_kind = function
| SimpleInversion -> None
@@ -440,61 +437,66 @@ let rewrite_equations_tac (gene, othin) id neqns names ba =
else rewrite_equations othin neqns names ba in
match othin with
| Some true (* if Inversion_clear, clear the hypothesis *) ->
- tclTHEN tac (tclTRY (clear [id]))
+ Tacticals.New.tclTHEN tac (Proofview.V82.tactic (tclTRY (clear [id])))
| _ ->
tac
-let raw_inversion inv_kind id status names gl =
- let env = pf_env gl and sigma = project gl in
+let raw_inversion inv_kind id status names =
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
+ Goal.concl >>- fun concl ->
let c = mkVar id in
+ Tacmach.New.pf_apply Tacred.reduce_to_atomic_ind >>- fun reduce_to_atomic_ind ->
+ Tacmach.New.pf_apply Typing.type_of >>- fun type_of ->
let (ind,t) =
- try pf_reduce_to_atomic_ind gl (pf_type_of gl c)
+ try
+ reduce_to_atomic_ind (type_of c)
with UserError _ ->
errorlabstrm "raw_inversion"
(str ("The type of "^(Id.to_string id)^" is not inductive.")) in
- let indclause = mk_clenv_from gl (c,t) in
+ Tacmach.New.of_old (fun gl -> mk_clenv_from gl (c,t)) >>- fun indclause ->
let ccl = clenv_type indclause in
check_no_metas indclause ccl;
let IndType (indf,realargs) = find_rectype env sigma ccl in
let (elim_predicate,neqns) =
- make_inv_predicate env sigma indf realargs id status (pf_concl gl) in
+ make_inv_predicate env sigma indf realargs id status concl in
let (cut_concl,case_tac) =
- if status != NoDep && (dependent c (pf_concl gl)) then
+ if status != NoDep & (dependent c concl) then
Reduction.beta_appvect elim_predicate (Array.of_list (realargs@[c])),
- case_then_using
+ Tacticals.New.case_then_using
else
Reduction.beta_appvect elim_predicate (Array.of_list realargs),
- case_nodep_then_using
+ Tacticals.New.case_nodep_then_using
in
- (tclTHENS
+ (Tacticals.New.tclTHENS
(assert_tac Anonymous cut_concl)
[case_tac names
(introCaseAssumsThen (rewrite_equations_tac inv_kind id neqns))
(Some elim_predicate) ([],[]) ind indclause;
- onLastHypId
+ Tacticals.New.onLastHypId
(fun id ->
- (tclTHEN
- (apply_term (mkVar id)
- (List.init neqns (fun _ -> Evarutil.mk_new_meta())))
+ (Tacticals.New.tclTHEN
+ (Proofview.V82.tactic (apply_term (mkVar id)
+ (List.init neqns (fun _ -> Evarutil.mk_new_meta()))))
reflexivity))])
- gl
(* Error messages of the inversion tactics *)
let wrap_inv_error id = function
| Indrec.RecursionSchemeError
(Indrec.NotAllowedCaseAnalysis (_,(Type _ | Prop Pos as k),i)) ->
- errorlabstrm ""
+ Proofview.tclZERO (Errors.UserError ("",
(strbrk "Inversion would require case analysis on sort " ++
pr_sort k ++
strbrk " which is not allowed for inductive definition " ++
- pr_inductive (Global.env()) i ++ str ".")
- | e -> raise e
+ pr_inductive (Global.env()) i ++ str ".")))
+ | e -> Proofview.tclZERO e
(* The most general inversion tactic *)
-let inversion inv_kind status names id gls =
- try (raw_inversion inv_kind id status names) gls
- with e when Errors.noncritical e -> wrap_inv_error id e
+let inversion inv_kind status names id =
+ Proofview.tclORELSE
+ (raw_inversion inv_kind id status names)
+ (wrap_inv_error id)
(* Specializing it... *)
@@ -519,25 +521,26 @@ let dinv_clear_tac id = dinv FullInversionClear None None (NamedHyp id)
* perform inversion on the named hypothesis. After, it will intro them
* back to their places in the hyp-list. *)
-let invIn k names ids id gls =
- let hyps = List.map (pf_get_hyp gls) ids in
- let nb_prod_init = nb_prod (pf_concl gls) in
- let intros_replace_ids gls =
+let invIn k names ids id =
+ Goal.sensitive_list_map Tacmach.New.pf_get_hyp ids >>- fun hyps ->
+ Goal.concl >>- fun concl ->
+ let nb_prod_init = nb_prod concl in
+ let intros_replace_ids =
+ Goal.concl >>- fun concl ->
let nb_of_new_hyp =
- nb_prod (pf_concl gls) - (List.length hyps + nb_prod_init)
+ nb_prod concl - (List.length hyps + nb_prod_init)
in
if nb_of_new_hyp < 1 then
- intros_replacing ids gls
+ intros_replacing ids
else
- tclTHEN (tclDO nb_of_new_hyp intro) (intros_replacing ids) gls
+ Tacticals.New.tclTHEN (Tacticals.New.tclDO nb_of_new_hyp intro) (intros_replacing ids)
in
- try
- (tclTHENSEQ
- [bring_hyps hyps;
+ Proofview.tclORELSE
+ (Tacticals.New.tclTHENLIST
+ [Proofview.V82.tactic (bring_hyps hyps);
inversion (false,k) NoDep names id;
intros_replace_ids])
- gls
- with e when Errors.noncritical e -> wrap_inv_error id e
+ (wrap_inv_error id)
let invIn_gen k names idl = try_intros_until (invIn k names idl)
diff --git a/tactics/inv.mli b/tactics/inv.mli
index 52db199ee..9491b7d7b 100644
--- a/tactics/inv.mli
+++ b/tactics/inv.mli
@@ -19,24 +19,24 @@ type inversion_status = Dep of constr option | NoDep
val inv_gen :
bool -> inversion_kind -> inversion_status ->
- intro_pattern_expr located option -> quantified_hypothesis -> tactic
+ intro_pattern_expr located option -> quantified_hypothesis -> unit Proofview.tactic
val invIn_gen :
inversion_kind -> intro_pattern_expr located option -> Id.t list ->
- quantified_hypothesis -> tactic
+ quantified_hypothesis -> unit Proofview.tactic
val inv_clause :
inversion_kind -> intro_pattern_expr located option -> Id.t list ->
- quantified_hypothesis -> tactic
+ quantified_hypothesis -> unit Proofview.tactic
val inv : inversion_kind -> intro_pattern_expr located option ->
- quantified_hypothesis -> tactic
+ quantified_hypothesis -> unit Proofview.tactic
val dinv : inversion_kind -> constr option ->
- intro_pattern_expr located option -> quantified_hypothesis -> tactic
+ intro_pattern_expr located option -> quantified_hypothesis -> unit Proofview.tactic
-val half_inv_tac : Id.t -> tactic
-val inv_tac : Id.t -> tactic
-val inv_clear_tac : Id.t -> tactic
-val half_dinv_tac : Id.t -> tactic
-val dinv_tac : Id.t -> tactic
-val dinv_clear_tac : Id.t -> tactic
+val half_inv_tac : Id.t -> unit Proofview.tactic
+val inv_tac : Id.t -> unit Proofview.tactic
+val inv_clear_tac : Id.t -> unit Proofview.tactic
+val half_dinv_tac : Id.t -> unit Proofview.tactic
+val dinv_tac : Id.t -> unit Proofview.tactic
+val dinv_clear_tac : Id.t -> unit Proofview.tactic
diff --git a/tactics/leminv.ml b/tactics/leminv.ml
index 19e7153b5..00cf4e997 100644
--- a/tactics/leminv.ml
+++ b/tactics/leminv.ml
@@ -28,6 +28,7 @@ open Declare
open Tacticals
open Tactics
open Decl_kinds
+open Proofview.Notations
let no_inductive_inconstr env constr =
(str "Cannot recognize an inductive predicate in " ++
@@ -196,8 +197,10 @@ let inversion_scheme env sigma t sort dep_option inv_op =
(str"Computed inversion goal was not closed in initial signature.");
*)
let pf = Proof.start [invEnv,invGoal] in
- let pf = Proof.run_tactic env
- (Proofview.V82.tactic (tclTHEN intro (onLastHypId inv_op))) pf in
+ let pf =
+ Proof.run_tactic env (
+ Tacticals.New.tclTHEN intro (Tacticals.New.onLastHypId inv_op)) pf
+ in
let pfterm = List.hd (Proof.partial_proof pf) in
let global_named_context = Global.named_context () in
let ownSign = ref begin
@@ -278,19 +281,20 @@ let lemInv id c gls =
(str "Cannot refine current goal with the lemma " ++
pr_lconstr_env (Global.env()) c)
-let lemInv_gen id c = try_intros_until (fun id -> lemInv id c) id
+let lemInv_gen id c = try_intros_until (fun id -> Proofview.V82.tactic (lemInv id c)) id
-let lemInvIn id c ids gls =
- let hyps = List.map (pf_get_hyp gls) ids in
- let intros_replace_ids gls =
- let nb_of_new_hyp = nb_prod (pf_concl gls) - List.length ids in
+let lemInvIn id c ids =
+ Goal.sensitive_list_map Tacmach.New.pf_get_hyp ids >>- fun hyps ->
+ let intros_replace_ids =
+ Goal.concl >>- fun concl ->
+ let nb_of_new_hyp = nb_prod concl - List.length ids in
if nb_of_new_hyp < 1 then
- intros_replacing ids gls
+ intros_replacing ids
else
- (tclTHEN (tclDO nb_of_new_hyp intro) (intros_replacing ids)) gls
+ (Tacticals.New.tclTHEN (Tacticals.New.tclDO nb_of_new_hyp intro) (intros_replacing ids))
in
- ((tclTHEN (tclTHEN (bring_hyps hyps) (lemInv id c))
- (intros_replace_ids)) gls)
+ ((Tacticals.New.tclTHEN (Proofview.V82.tactic (tclTHEN (bring_hyps hyps) (lemInv id c)))
+ (intros_replace_ids)))
let lemInvIn_gen id c l = try_intros_until (fun id -> lemInvIn id c l) id
diff --git a/tactics/leminv.mli b/tactics/leminv.mli
index 5019ceda5..16695fcd7 100644
--- a/tactics/leminv.mli
+++ b/tactics/leminv.mli
@@ -15,15 +15,15 @@ open Proof_type
open Constrexpr
open Misctypes
-val lemInv_gen : quantified_hypothesis -> constr -> tactic
-val lemInvIn_gen : quantified_hypothesis -> constr -> Id.t list -> tactic
+val lemInv_gen : quantified_hypothesis -> constr -> unit Proofview.tactic
+val lemInvIn_gen : quantified_hypothesis -> constr -> Id.t list -> unit Proofview.tactic
val lemInv_clause :
- quantified_hypothesis -> constr -> Id.t list -> tactic
+ quantified_hypothesis -> constr -> Id.t list -> unit Proofview.tactic
val inversion_lemma_from_goal :
int -> Id.t -> Id.t located -> sorts -> bool ->
- (Id.t -> tactic) -> unit
+ (Id.t -> unit Proofview.tactic) -> unit
val add_inversion_lemma_exn :
- Id.t -> constr_expr -> glob_sort -> bool -> (Id.t -> tactic) ->
+ Id.t -> constr_expr -> glob_sort -> bool -> (Id.t -> unit Proofview.tactic) ->
unit
diff --git a/tactics/refine.ml b/tactics/refine.ml
index 03b697e8f..4a601e41e 100644
--- a/tactics/refine.ml
+++ b/tactics/refine.ml
@@ -59,6 +59,7 @@ open Environ
open Tactics
open Tacticals
open Printer
+open Proofview.Notations
type term_with_holes = TH of constr * meta_type_map * sg_proofs
and sg_proofs = (term_with_holes option) list
@@ -262,85 +263,83 @@ let rec compute_metamap env sigma c = match kind_of_term c with
let ensure_products n =
let p = ref 0 in
- let rec aux n gl =
- if Int.equal n 0 then tclFAIL 0 (mt()) gl
+ let rec aux n =
+ if Int.equal n 0 then Tacticals.New.tclFAIL 0 (mt())
else
- tclTHEN
- (tclORELSE intro (fun gl -> incr p; introf gl))
- (aux (n-1)) gl in
- tclORELSE
+ Tacticals.New.tclTHEN
+ (Tacticals.New.tclORELSE intro (Proofview.tclUNIT () >= fun () -> incr p; introf))
+ (aux (n-1)) in
+ Tacticals.New.tclORELSE
(aux n)
(* Now we know how many red are needed *)
- (fun gl -> tclDO !p red_in_concl gl)
+ (Proofview.V82.tactic (fun gl -> tclDO !p red_in_concl gl))
-let rec tcc_aux subst (TH (c,mm,sgp) as _th) gl =
+let rec tcc_aux subst (TH (c,mm,sgp) as _th) : unit Proofview.tactic =
let c = substl subst c in
match (kind_of_term c,sgp) with
(* mv => sous-but : on ne fait rien *)
| Meta _ , _ ->
- tclIDTAC gl
+ Proofview.tclUNIT ()
| Cast (c,_,_), _ when isMeta c ->
- tclIDTAC gl
+ Proofview.tclUNIT ()
(* terme pur => refine *)
| _,[] ->
- refine c gl
+ Proofview.V82.tactic (refine c)
(* abstraction => intro *)
| Lambda (Name id,_,m), _ ->
assert (isMeta (strip_outer_cast m));
begin match sgp with
- | [None] -> intro_mustbe_force id gl
+ | [None] -> intro_mustbe_force id
| [Some th] ->
- tclTHEN (introduction id)
- (onLastHypId (fun id -> tcc_aux (mkVar id::subst) th)) gl
+ Tacticals.New.tclTHEN (Proofview.V82.tactic (introduction id))
+ (Tacticals.New.onLastHypId (fun id -> tcc_aux (mkVar id::subst) th))
| _ -> assert false
end
| Lambda (Anonymous,_,m), _ -> (* if anon vars are allowed in evars *)
assert (isMeta (strip_outer_cast m));
begin match sgp with
- | [None] -> tclTHEN intro (onLastHypId (fun id -> clear [id])) gl
+ | [None] -> Tacticals.New.tclTHEN intro (Proofview.V82.tactic (onLastHypId (fun id -> clear [id])))
| [Some th] ->
- tclTHEN
+ Tacticals.New.tclTHEN
intro
- (onLastHypId (fun id ->
- tclTHEN
- (clear [id])
- (tcc_aux (mkVar (*dummy*) id::subst) th))) gl
+ (Tacticals.New.onLastHypId (fun id ->
+ Tacticals.New.tclTHEN
+ (Proofview.V82.tactic (clear [id]))
+ (tcc_aux (mkVar (*dummy*) id::subst) th)))
| _ -> assert false
end
(* let in without holes in the body => possibly dependent intro *)
| LetIn (Name id,c1,t1,c2), _ when not (isMeta (strip_outer_cast c1)) ->
- let c = pf_concl gl in
+ Goal.concl >>- fun c ->
let newc = mkNamedLetIn id c1 t1 c in
- tclTHEN
- (change_in_concl None newc)
+ Tacticals.New.tclTHEN
+ (Proofview.V82.tactic (change_in_concl None newc))
(match sgp with
- | [None] -> introduction id
+ | [None] -> Proofview.V82.tactic (introduction id)
| [Some th] ->
- tclTHEN (introduction id)
- (onLastHypId (fun id -> tcc_aux (mkVar id::subst) th))
+ Tacticals.New.tclTHEN (Proofview.V82.tactic (introduction id))
+ (Tacticals.New.onLastHypId (fun id -> tcc_aux (mkVar id::subst) th))
| _ -> assert false)
- gl
(* let in with holes in the body => unable to handle dependency
because of evars limitation, use non dependent assert instead *)
| LetIn (Name id,c1,t1,c2), _ ->
- tclTHENS
+ Tacticals.New.tclTHENS
(assert_tac (Name id) t1)
[(match List.hd sgp with
- | None -> tclIDTAC
- | Some th -> onLastHypId (fun id -> tcc_aux (mkVar id::subst) th));
+ | None -> Proofview.tclUNIT ()
+ | Some th -> Tacticals.New.onLastHypId (fun id -> tcc_aux (mkVar id::subst) th));
(match List.tl sgp with
- | [] -> refine (subst1 (mkVar id) c2) (* a complete proof *)
- | [None] -> tclIDTAC (* a meta *)
+ | [] -> Proofview.V82.tactic (refine (subst1 (mkVar id) c2)) (* a complete proof *)
+ | [None] -> Proofview.tclUNIT () (* a meta *)
| [Some th] -> (* a partial proof *)
- onLastHypId (fun id -> tcc_aux (mkVar id::subst) th)
+ Tacticals.New.onLastHypId (fun id -> tcc_aux (mkVar id::subst) th)
| _ -> assert false)]
- gl
(* fix => tactique Fix *)
| Fix ((ni,j),(fi,ai,_)) , _ ->
@@ -350,14 +349,13 @@ let rec tcc_aux subst (TH (c,mm,sgp) as _th) gl =
in
let fixes = Array.map3 (fun f n c -> (out_name f,succ n,c)) fi ni ai in
let firsts,lasts = List.chop j (Array.to_list fixes) in
- tclTHENS
- (tclTHEN
+ Tacticals.New.tclTHENS
+ (Tacticals.New.tclTHEN
(ensure_products (succ ni.(j)))
- (mutual_fix (out_name fi.(j)) (succ ni.(j)) (firsts@List.tl lasts) j))
- (List.map (function
- | None -> tclIDTAC
- | Some th -> tcc_aux subst th) sgp)
- gl
+ (Proofview.V82.tactic (mutual_fix (out_name fi.(j)) (succ ni.(j)) (firsts@List.tl lasts) j)))
+ ((List.map (function
+ | None -> Proofview.tclUNIT ()
+ | Some th -> tcc_aux subst th) sgp))
(* cofix => tactique CoFix *)
| CoFix (j,(fi,ai,_)) , _ ->
@@ -367,30 +365,29 @@ let rec tcc_aux subst (TH (c,mm,sgp) as _th) gl =
in
let cofixes = Array.map2 (fun f c -> (out_name f,c)) fi ai in
let firsts,lasts = List.chop j (Array.to_list cofixes) in
- tclTHENS
- (mutual_cofix (out_name fi.(j)) (firsts@List.tl lasts) j)
+ Tacticals.New.tclTHENS
+ (Proofview.V82.tactic (mutual_cofix (out_name fi.(j)) (firsts@List.tl lasts) j))
(List.map (function
- | None -> tclIDTAC
+ | None -> Proofview.tclUNIT ()
| Some th -> tcc_aux subst th) sgp)
- gl
(* sinon on fait refine du terme puis appels rec. sur les sous-buts.
* c'est le cas pour App et MutCase. *)
| _ ->
- tclTHENS
- (refine c)
+ Tacticals.New.tclTHENS
+ (Proofview.V82.tactic (refine c))
(List.map
- (function None -> tclIDTAC | Some th -> tcc_aux subst th) sgp)
- gl
+ (function None -> Proofview.tclUNIT () | Some th -> tcc_aux subst th) sgp)
(* Et finalement la tactique refine elle-même : *)
-let refine (evd,c) gl =
- let sigma = project gl in
- let evd = Typeclasses.resolve_typeclasses ~filter:Typeclasses.no_goals (pf_env gl) evd in
+let refine (evd,c) =
+ Goal.defs >>- fun sigma ->
+ Goal.env >>- fun env ->
+ let evd = Typeclasses.resolve_typeclasses ~filter:Typeclasses.no_goals env evd in
let c = Evarutil.nf_evar evd c in
let (evd,c) = Evarutil.evars_to_metas sigma (evd,c) in
(* Relies on Cast's put on Meta's by evars_to_metas, because it is otherwise
complicated to update meta types when passing through a binder *)
- let th = compute_metamap (pf_env gl) evd c in
- tclTHEN (Refiner.tclEVARS evd) (tcc_aux [] th) gl
+ let th = compute_metamap env evd c in
+ Tacticals.New.tclTHEN (Proofview.V82.tactic (Refiner.tclEVARS evd)) (tcc_aux [] th)
diff --git a/tactics/refine.mli b/tactics/refine.mli
index fc6b401a0..1be6d1f01 100644
--- a/tactics/refine.mli
+++ b/tactics/refine.mli
@@ -8,4 +8,4 @@
open Tacmach
-val refine : Evd.open_constr -> tactic
+val refine : Evd.open_constr -> unit Proofview.tactic
diff --git a/tactics/rewrite.ml b/tactics/rewrite.ml
index d6ad598c5..65f29498c 100644
--- a/tactics/rewrite.ml
+++ b/tactics/rewrite.ml
@@ -361,7 +361,7 @@ let solve_remaining_by by env prf =
let evd' =
List.fold_right (fun mv evd ->
let ty = Clenv.clenv_nf_meta env (meta_type evd mv) in
- let c = Pfedit.build_by_tactic env.env ty (tclCOMPLETE tac) in
+ let c = Pfedit.build_by_tactic env.env ty (Tacticals.New.tclCOMPLETE tac) in
meta_assign mv (c, (Conv,TypeNotProcessed)) evd)
indep env.evd
in { env with evd = evd' }, prf
@@ -1276,10 +1276,10 @@ let assert_replacing id newt tac =
env ~init:[]
in
let (e, args) = destEvar ev in
- Goal.return
+ Goal.return
(mkEvar (e, Array.of_list inst)))))
in Goal.bind reft Goal.refine)
- in Proofview.tclTHEN (Proofview.tclSENSITIVE sens)
+ in Tacticals.New.tclTHEN (Proofview.tclSENSITIVE sens)
(Proofview.tclFOCUS 2 2 tac)
let newfail n s =
@@ -1337,7 +1337,7 @@ let tactic_init_setoid () =
init_setoid (); tclIDTAC
let cl_rewrite_clause_new_strat ?abs strat clause =
- Proofview.tclTHEN (newtactic_init_setoid ())
+ Tacticals.New.tclTHEN (newtactic_init_setoid ())
(try cl_rewrite_clause_newtac ?abs strat clause
with RewriteFailure s ->
newfail 0 (str"setoid rewrite failed: " ++ s))
@@ -1747,73 +1747,94 @@ let general_s_rewrite cl l2r occs (c,l) ~new_goals gl =
(pf_env gl,project gl,
Pretype_errors.NoOccurrenceFound ((if l2r then x else y), cl)))
+open Proofview.Notations
+
let general_s_rewrite_clause x =
- init_setoid ();
match x with
| None -> general_s_rewrite None
| Some id -> general_s_rewrite (Some id)
+let general_s_rewrite_clause x y z w ~new_goals =
+ newtactic_init_setoid () <*>
+ Proofview.V82.tactic (general_s_rewrite_clause x y z w ~new_goals)
let _ = Hook.set Equality.general_rewrite_clause general_s_rewrite_clause
(** [setoid_]{reflexivity,symmetry,transitivity} tactics *)
let not_declared env ty rel =
- tclFAIL 0 (str" The relation " ++ Printer.pr_constr_env env rel ++ str" is not a declared " ++
+ Tacticals.New.tclFAIL 0 (str" The relation " ++ Printer.pr_constr_env env rel ++ str" is not a declared " ++
str ty ++ str" relation. Maybe you need to require the Setoid library")
-let setoid_proof gl ty fn fallback =
- let env = pf_env gl in
- try
- let rel, args = decompose_app_rel env (project gl) (pf_concl gl) in
- let evm = project gl in
- let car = pi3 (List.hd (fst (Reduction.dest_prod env (Typing.type_of env evm rel)))) in
- fn env evm car rel gl
- with e when Errors.noncritical e ->
- try fallback gl
- with Hipattern.NoEquationFound ->
- let e = Errors.push e in
- match e with
- | Not_found ->
- let rel, args = decompose_app_rel env (project gl) (pf_concl gl) in
- not_declared env ty rel gl
- | _ -> raise e
-
-let setoid_reflexivity gl =
- setoid_proof gl "reflexive"
- (fun env evm car rel -> apply (get_reflexive_proof env evm car rel))
- (reflexivity_red true)
-
-let setoid_symmetry gl =
- setoid_proof gl "symmetric"
- (fun env evm car rel -> apply (get_symmetric_proof env evm car rel))
+let setoid_proof ty fn fallback =
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
+ Goal.concl >>- fun concl ->
+ Proofview.tclORELSE
+ begin
+ try
+ let rel, args = decompose_app_rel env sigma concl in
+ let evm = sigma in
+ let car = pi3 (List.hd (fst (Reduction.dest_prod env (Typing.type_of env evm rel)))) in
+ fn env evm car rel
+ with e -> Proofview.tclZERO e
+ end
+ begin function
+ | e ->
+ Proofview.tclORELSE
+ fallback
+ begin function
+ | Hipattern.NoEquationFound ->
+ (* spiwack: [Errors.push] here is unlikely to do what
+ it's intended to, or anything meaningful for that
+ matter. *)
+ let e = Errors.push e in
+ begin match e with
+ | Not_found ->
+ let rel, args = decompose_app_rel env sigma concl in
+ not_declared env ty rel
+ | _ -> Proofview.tclZERO e
+ end
+ | e' -> Proofview.tclZERO e'
+ end
+ end
+
+let setoid_reflexivity =
+ setoid_proof "reflexive"
+ (fun env evm car rel -> Proofview.V82.tactic (apply (get_reflexive_proof env evm car rel)))
+ (reflexivity_red true)
+
+let setoid_symmetry =
+ setoid_proof "symmetric"
+ (fun env evm car rel -> Proofview.V82.tactic (apply (get_symmetric_proof env evm car rel)))
(symmetry_red true)
-let setoid_transitivity c gl =
- setoid_proof gl "transitive"
+let setoid_transitivity c =
+ setoid_proof "transitive"
(fun env evm car rel ->
- let proof = get_transitive_proof env evm car rel in
- match c with
- | None -> eapply proof
- | Some c -> apply_with_bindings (proof,ImplicitBindings [ c ]))
+ Proofview.V82.tactic begin
+ let proof = get_transitive_proof env evm car rel in
+ match c with
+ | None -> eapply proof
+ | Some c -> apply_with_bindings (proof,ImplicitBindings [ c ])
+ end)
(transitivity_red true c)
-let setoid_symmetry_in id gl =
- let ctype = pf_type_of gl (mkVar id) in
- let binders,concl = decompose_prod_assum ctype in
- let (equiv, args) = decompose_app concl in
- let rec split_last_two = function
- | [c1;c2] -> [],(c1, c2)
- | x::y::z -> let l,res = split_last_two (y::z) in x::l, res
- | _ -> error "The term provided is not an equivalence."
- in
- let others,(c1,c2) = split_last_two args in
- let he,c1,c2 = mkApp (equiv, Array.of_list others),c1,c2 in
- let new_hyp' = mkApp (he, [| c2 ; c1 |]) in
- let new_hyp = it_mkProd_or_LetIn new_hyp' binders in
- tclTHENS (Tactics.cut new_hyp)
- [ intro_replacing id;
- tclTHENLIST [ intros; setoid_symmetry; apply (mkVar id); Tactics.assumption ] ]
- gl
+let setoid_symmetry_in id =
+ Tacmach.New.of_old (fun gl -> pf_type_of gl (mkVar id)) >>- fun ctype ->
+ let binders,concl = decompose_prod_assum ctype in
+ let (equiv, args) = decompose_app concl in
+ let rec split_last_two = function
+ | [c1;c2] -> [],(c1, c2)
+ | x::y::z -> let l,res = split_last_two (y::z) in x::l, res
+ | _ -> error "The term provided is not an equivalence."
+ in
+ let others,(c1,c2) = split_last_two args in
+ let he,c1,c2 = mkApp (equiv, Array.of_list others),c1,c2 in
+ let new_hyp' = mkApp (he, [| c2 ; c1 |]) in
+ let new_hyp = it_mkProd_or_LetIn new_hyp' binders in
+ Tacticals.New.tclTHENS (Proofview.V82.tactic (Tactics.cut new_hyp))
+ [ Proofview.V82.tactic (intro_replacing id);
+ Tacticals.New.tclTHENLIST [ intros; setoid_symmetry; Proofview.V82.tactic (apply (mkVar id)); Proofview.V82.tactic (Tactics.assumption) ] ]
let _ = Hook.set Tactics.setoid_reflexivity setoid_reflexivity
let _ = Hook.set Tactics.setoid_symmetry setoid_symmetry
@@ -1831,7 +1852,7 @@ let implify id gl =
let tyhd = Typing.type_of env sigma ty
and tyconcl = Typing.type_of (Environ.push_rel hd env) sigma concl in
let app, unfold = arrow_morphism tyhd (subst1 mkProp tyconcl) ty (subst1 mkProp concl) in
- it_mkProd_or_LetIn app tl
+ it_mkProd_or_LetIn app tl
| _ -> ctype
in convert_hyp_no_check (id, b, ctype') gl
diff --git a/tactics/rewrite.mli b/tactics/rewrite.mli
index 4d1e285f9..637bab5a6 100644
--- a/tactics/rewrite.mli
+++ b/tactics/rewrite.mli
@@ -71,13 +71,13 @@ val default_morphism :
(types * constr option) option list * (types * types option) option ->
constr -> constr * constr
-val setoid_symmetry : tactic
+val setoid_symmetry : unit Proofview.tactic
-val setoid_symmetry_in : Id.t -> tactic
+val setoid_symmetry_in : Id.t -> unit Proofview.tactic
-val setoid_reflexivity : tactic
+val setoid_reflexivity : unit Proofview.tactic
-val setoid_transitivity : constr option -> tactic
+val setoid_transitivity : constr option -> unit Proofview.tactic
val implify : Id.t -> tactic
diff --git a/tactics/tacinterp.ml b/tactics/tacinterp.ml
index 0e0ccac2e..04624e1a1 100644
--- a/tactics/tacinterp.ml
+++ b/tactics/tacinterp.ml
@@ -43,6 +43,7 @@ open Miscops
open Locus
open Tacintern
open Taccoerce
+open Proofview.Notations
let safe_msgnl s =
let _ =
@@ -55,8 +56,7 @@ type value = tlevel generic_argument
(* Values for interpretation *)
type tacvalue =
- | VRTactic of (goal list sigma) (* For Match results *)
- (* Not a true tacvalue *)
+ | VRTactic (* variant of unit returned by match. For historical reasons. *)
| VFun of ltac_trace * value Id.Map.t *
Id.t option list * glob_tactic_expr
| VRec of value Id.Map.t ref * glob_tactic_expr
@@ -71,19 +71,18 @@ module Value = Taccoerce.Value
let dloc = Loc.ghost
-let catch_error call_trace tac g =
- try tac g with e when Errors.noncritical e ->
+let catching_error call_trace fail e =
let e = Errors.push e in
let inner_trace, e = match Exninfo.get e ltac_trace_info with
| Some inner_trace -> inner_trace, e
| None -> [], e
in
- if List.is_empty call_trace && List.is_empty inner_trace then raise e
+ if List.is_empty call_trace & List.is_empty inner_trace then fail e
else begin
assert (Errors.noncritical e); (* preserved invariant *)
let new_trace = inner_trace @ call_trace in
let located_exc = Exninfo.add e ltac_trace_info new_trace in
- raise located_exc
+ fail located_exc
end
module TacStore = Geninterp.TacStore
@@ -93,6 +92,13 @@ let f_avoid_ids : Id.t list TacStore.field = TacStore.field ()
let f_debug : debug_info TacStore.field = TacStore.field ()
let f_trace : ltac_trace TacStore.field = TacStore.field ()
+let catch_error call_trace f x =
+ try f x with e when Errors.noncritical e -> catching_error call_trace raise e
+let catch_error_tac call_trace tac =
+ Proofview.tclORELSE
+ tac
+ (catching_error call_trace Proofview.tclZERO)
+
(* Signature for interpretation: val_interp and interpretation functions *)
type interp_sign = Geninterp.interp_sign = {
lfun : value Id.Map.t;
@@ -107,7 +113,7 @@ let check_is_value v =
if has_type v (topwit wit_tacvalue) then
let v = to_tacvalue v in
match v with
- | VRTactic _ -> (* These are goals produced by Match *)
+ | VRTactic -> (* These are goals produced by Match *)
error "Immediate match producing tactics not allowed in local definitions."
| _ -> ()
else ()
@@ -248,7 +254,7 @@ let coerce_to_tactic loc id v =
if has_type v (topwit wit_tacvalue) then
let tacv = to_tacvalue v in
match tacv with
- | VFun _ | VRTactic _ -> v
+ | VFun _ | VRTactic -> v
| _ -> fail ()
else fail ()
@@ -966,7 +972,17 @@ let verify_metas_coherence gl (ln1,lcm) (ln,lm) =
let adjust (l, lc) = (l, Id.Map.map (fun c -> [], c) lc)
-type 'a extended_matching_result =
+
+(* spiwack: a small wrapper around the [Matching] module *)
+
+type 'a _matching_result =
+ { s_sub : bound_ident_map * patvar_map ;
+ s_ctx : 'a ;
+ s_nxt : 'a matching_result }
+and 'a matching_result = 'a _matching_result Goal.sensitive Proofview.tactic
+
+
+type 'a _extended_matching_result =
{ e_ctx : 'a;
e_sub : bound_ident_map * extended_patvar_map; }
@@ -1051,168 +1067,156 @@ let extend_gl_hyps { it=gl ; sigma=sigma } sign =
Goal.V82.new_goal_with sigma gl sign
(* Interprets an l-tac expression into a value *)
-let rec val_interp ist gl (tac:glob_tactic_expr) : Evd.evar_map * typed_generic_argument =
+let rec val_interp ist (tac:glob_tactic_expr) : typed_generic_argument Goal.sensitive Proofview.tactic =
let value_interp ist = match tac with
(* Immediate evaluation *)
| TacFun (it,body) ->
let v = VFun (extract_trace ist,ist.lfun,it,body) in
- project gl, of_tacvalue v
- | TacLetIn (true,l,u) -> interp_letrec ist gl l u
- | TacLetIn (false,l,u) -> interp_letin ist gl l u
- | TacMatchGoal (lz,lr,lmr) -> interp_match_goal ist gl lz lr lmr
- | TacMatch (lz,c,lmr) -> interp_match ist gl lz c lmr
- | TacArg (loc,a) -> interp_tacarg ist gl a
+ Proofview.tclUNIT (Goal.return (of_tacvalue v))
+ | TacLetIn (true,l,u) -> interp_letrec ist l u
+ | TacLetIn (false,l,u) -> interp_letin ist l u
+ | TacMatchGoal (lz,lr,lmr) -> interp_match_goal ist lz lr lmr
+ | TacMatch (lz,c,lmr) -> interp_match ist lz c lmr
+ | TacArg (loc,a) -> interp_tacarg ist a
(* Delayed evaluation *)
| t ->
let v = VFun (extract_trace ist,ist.lfun,[],t) in
- project gl, of_tacvalue v
-
+ Proofview.tclUNIT (Goal.return (of_tacvalue v))
in check_for_interrupt ();
- match curr_debug ist with
- | DebugOn lev ->
- let eval v =
- let ist = { ist with extra = TacStore.set ist.extra f_debug v } in
- value_interp ist
- in
- debug_prompt lev gl tac eval
- | _ -> value_interp ist
+ match curr_debug ist with
+ (* arnaud: todo: debug
+ | DebugOn lev ->
+ debug_prompt lev gl tac (fun v -> value_interp {ist with debug=v})
+ *)
+ | _ -> value_interp ist
+
and eval_tactic ist = function
| TacAtom (loc,t) ->
- fun gl ->
- let call = LtacAtomCall t in
- let tac = (* catch error in the interpretation *)
- catch_error (push_trace(dloc,call)ist)
- (interp_atomic ist gl) t in
- (* catch error in the evaluation *)
- catch_error (push_trace(loc,call)ist) tac gl
+ let call = LtacAtomCall t in
+ catch_error_tac (push_trace(loc,call) ist) (interp_atomic ist t)
| TacFun _ | TacLetIn _ -> assert false
| TacMatchGoal _ | TacMatch _ -> assert false
- | TacId s -> fun gl ->
- let res = tclIDTAC_MESSAGE (interp_message_nl ist gl s) gl in
- db_breakpoint (curr_debug ist) s; res
- | TacFail (n,s) -> fun gl -> tclFAIL (interp_int_or_var ist n) (interp_message ist gl s) gl
- | TacProgress tac -> tclPROGRESS (interp_tactic ist tac)
- | TacShowHyps tac -> tclSHOWHYPS (interp_tactic ist tac)
+ | TacId s -> Proofview.V82.tactic begin fun gl ->
+ let res = tclIDTAC_MESSAGE (interp_message_nl ist gl s) gl in
+ db_breakpoint (curr_debug ist) s; res
+ end
+ | TacFail (n,s) ->
+ Proofview.V82.tactic begin fun gl ->
+ tclFAIL (interp_int_or_var ist n) (interp_message ist gl s) gl
+ end
+ | TacProgress tac -> Proofview.tclPROGRESS (interp_tactic ist tac)
+ | TacShowHyps tac ->
+ (* arnaud: todo:
+ tclSHOWHYPS (interp_tactic ist tac)*)
+ assert false
| TacAbstract (tac,ido) ->
- fun gl -> Tactics.tclABSTRACT
+ Proofview.V82.tactic begin fun gl -> Tactics.tclABSTRACT
(Option.map (pf_interp_ident ist gl) ido) (interp_tactic ist tac) gl
+ end
| TacThen (t1,tf,t,tl) ->
- tclTHENS3PARTS (interp_tactic ist t1)
+ Tacticals.New.tclTHENS3PARTS (interp_tactic ist t1)
(Array.map (interp_tactic ist) tf) (interp_tactic ist t) (Array.map (interp_tactic ist) tl)
- | TacThens (t1,tl) -> tclTHENS (interp_tactic ist t1) (List.map (interp_tactic ist) tl)
- | TacDo (n,tac) -> tclDO (interp_int_or_var ist n) (interp_tactic ist tac)
- | TacTimeout (n,tac) -> tclTIMEOUT (interp_int_or_var ist n) (interp_tactic ist tac)
- | TacTry tac -> tclTRY (interp_tactic ist tac)
- | TacRepeat tac -> tclREPEAT (interp_tactic ist tac)
+ | TacThens (t1,tl) -> Tacticals.New.tclTHENS (interp_tactic ist t1) (List.map (interp_tactic ist) tl)
+ | TacDo (n,tac) -> Tacticals.New.tclDO (interp_int_or_var ist n) (interp_tactic ist tac)
+ | TacTimeout (n,tac) -> Proofview.tclTIMEOUT (interp_int_or_var ist n) (interp_tactic ist tac)
+ | TacTry tac -> Tacticals.New.tclTRY (interp_tactic ist tac)
+ | TacRepeat tac -> Tacticals.New.tclREPEAT (interp_tactic ist tac)
| TacOrelse (tac1,tac2) ->
- tclORELSE (interp_tactic ist tac1) (interp_tactic ist tac2)
- | TacFirst l -> tclFIRST (List.map (interp_tactic ist) l)
- | TacSolve l -> tclSOLVE (List.map (interp_tactic ist) l)
- | TacComplete tac -> tclCOMPLETE (interp_tactic ist tac)
+ Tacticals.New.tclORELSE (interp_tactic ist tac1) (interp_tactic ist tac2)
+ | TacFirst l -> Tacticals.New.tclFIRST (List.map (interp_tactic ist) l)
+ | TacSolve l -> Tacticals.New.tclSOLVE (List.map (interp_tactic ist) l)
+ | TacComplete tac -> Tacticals.New.tclCOMPLETE (interp_tactic ist tac)
| TacArg a -> interp_tactic ist (TacArg a)
| TacInfo tac ->
msg_warning
(strbrk "The general \"info\" tactic is currently not working." ++ fnl () ++
- strbrk "Some specific verbose tactics may exist instead, such as info_trivial, info_auto, info_eauto.");
+ strbrk "Some specific verbose tactics may exist instead, such as info_trivial, info_auto, info_eauto.");
eval_tactic ist tac
-and force_vrec ist gl v =
+and force_vrec ist v =
let v = Value.normalize v in
if has_type v (topwit wit_tacvalue) then
let v = to_tacvalue v in
match v with
- | VRec (lfun,body) -> val_interp {ist with lfun = !lfun} gl body
- | v -> project gl , of_tacvalue v
- else project gl, v
+ | VRec (lfun,body) -> val_interp {ist with lfun = !lfun} body
+ | v -> Proofview.tclUNIT (Goal.return (of_tacvalue v))
+ else Proofview.tclUNIT (Goal.return v)
-and interp_ltac_reference loc' mustbetac ist gl = function
+and interp_ltac_reference loc' mustbetac ist = function
| ArgVar (loc,id) ->
let v =
try Id.Map.find id ist.lfun
with Not_found -> in_gen (topwit wit_var) id
in
- let (sigma,v) = force_vrec ist gl v in
+ force_vrec ist v >>== fun v ->
let v = propagate_trace ist loc id v in
- sigma , if mustbetac then coerce_to_tactic loc id v else v
+ if mustbetac then Proofview.tclUNIT (Goal.return (coerce_to_tactic loc id v)) else Proofview.tclUNIT (Goal.return v)
| ArgArg (loc,r) ->
let ids = extract_ids [] ist.lfun in
let loc_info = ((if Loc.is_ghost loc' then loc else loc'),LtacNameCall r) in
let extra = TacStore.set ist.extra f_avoid_ids ids in
let extra = TacStore.set extra f_trace (push_trace loc_info ist) in
let ist = { lfun = Id.Map.empty; extra = extra; } in
- val_interp ist gl (lookup_ltacref r)
+ val_interp ist (lookup_ltacref r)
-and interp_tacarg ist gl arg =
- let evdref = ref (project gl) in
- let v = match arg with
- | TacGeneric arg ->
- let gl = { gl with sigma = !evdref } in
- let (sigma, v) = Geninterp.generic_interp ist gl arg in
- evdref := sigma;
- v
- | Reference r ->
- let (sigma,v) = interp_ltac_reference dloc false ist gl r in
- evdref := sigma;
- v
- | ConstrMayEval c ->
- let (sigma,c_interp) = interp_constr_may_eval ist gl c in
- evdref := sigma;
- Value.of_constr c_interp
- | MetaIdArg (loc,_,id) -> assert false
- | TacCall (loc,r,[]) ->
- let (sigma,v) = interp_ltac_reference loc true ist gl r in
- evdref := sigma;
- v
- | TacCall (loc,f,l) ->
- let (sigma,fv) = interp_ltac_reference loc true ist gl f in
- let (sigma,largs) =
- List.fold_right begin fun a (sigma',acc) ->
- let (sigma', a_interp) = interp_tacarg ist gl a in
- sigma' , a_interp::acc
- end l (sigma,[])
- in
- List.iter check_is_value largs;
- let (sigma,v) = interp_app loc ist { gl with sigma=sigma } fv largs in
- evdref:= sigma;
- v
- | TacExternal (loc,com,req,la) ->
- let (sigma,la_interp) =
- List.fold_right begin fun a (sigma,acc) ->
- let (sigma,a_interp) = interp_tacarg ist {gl with sigma=sigma} a in
- sigma , a_interp::acc
- end la (project gl,[])
- in
- let (sigma,v) = interp_external loc ist { gl with sigma=sigma } com req la_interp in
- evdref := sigma;
- v
- | TacFreshId l ->
- let id = pf_interp_fresh_id ist gl l in
- in_gen (topwit wit_intro_pattern) (dloc, IntroIdentifier id)
- | Tacexp t ->
- let (sigma,v) = val_interp ist gl t in
- evdref := sigma;
- v
- | TacDynamic(_,t) ->
- let tg = (Dyn.tag t) in
- if String.equal tg "tactic" then
- let (sigma,v) = val_interp ist gl (tactic_out t ist) in
- evdref := sigma;
- v
- else if String.equal tg "value" then
- value_out t
- else if String.equal tg "constr" then
- Value.of_constr (constr_out t)
- else
- anomaly ~loc:dloc ~label:"Tacinterp.val_interp"
- (str "Unknown dynamic: <" ++ str (Dyn.tag t) ++ str ">")
+and interp_tacarg ist arg =
+ let tac_of_old f =
+ Tacmach.New.of_old f >>-- fun (sigma,v) ->
+ Proofview.V82.tactic (tclEVARS sigma) <*>
+ Proofview.tclUNIT (Goal.return v)
in
- !evdref , v
+ match arg with
+ | TacGeneric arg ->
+ Goal.defs >>-- fun sigma ->
+ tac_of_old (fun gl ->
+ Geninterp.generic_interp ist { gl with sigma = sigma } arg)
+ | Reference r -> interp_ltac_reference dloc false ist r
+ | ConstrMayEval c ->
+ tac_of_old (fun gl -> interp_constr_may_eval ist gl c) >>== fun c_interp ->
+ Proofview.tclUNIT (Goal.return (Value.of_constr c_interp))
+ | MetaIdArg (loc,_,id) -> assert false
+ | TacCall (loc,r,[]) ->
+ interp_ltac_reference loc true ist r
+ | TacCall (loc,f,l) ->
+ interp_ltac_reference loc true ist f >>== fun fv ->
+ Proofview.tclEVARMAP >= fun sigma ->
+ List.fold_right begin fun a acc ->
+ interp_tacarg ist a >>== fun a_interp ->
+ acc >>== fun acc -> Proofview.tclUNIT (Goal.return (a_interp::acc))
+ end l (Proofview.tclUNIT (Goal.return [])) >>== fun largs ->
+ interp_app loc ist fv largs
+ | TacExternal (loc,com,req,la) ->
+ List.fold_right begin fun a acc ->
+ interp_tacarg ist a >>== fun a_interp ->
+ acc >>== fun acc -> Proofview.tclUNIT (Goal.return (a_interp::acc))
+ end la (Proofview.tclUNIT (Goal.return [])) >>== fun la_interp ->
+ tac_of_old (fun gl -> interp_external loc ist { gl with sigma=project gl } com req la_interp)
+ | TacFreshId l ->
+ Tacmach.New.of_old (fun gl -> pf_interp_fresh_id ist gl l) >>-- fun id ->
+ Proofview.tclUNIT (Goal.return (in_gen (topwit wit_intro_pattern) (dloc, IntroIdentifier id)))
+ | Tacexp t -> val_interp ist t
+ | TacDynamic(_,t) ->
+ let tg = (Dyn.tag t) in
+ if String.equal tg "tactic" then
+ val_interp ist (tactic_out t ist)
+ else if String.equal tg "value" then
+ Proofview.tclUNIT (Goal.return (value_out t))
+ else if String.equal tg "constr" then
+ Proofview.tclUNIT (Goal.return (Value.of_constr (constr_out t)))
+ else
+ Errors.anomaly ~loc:dloc ~label:"Tacinterp.val_interp"
+ (str "Unknown dynamic: <" ++ str (Dyn.tag t) ++ str ">")
(* Interprets an application node *)
-and interp_app loc ist gl fv largs =
- let fail () = user_err_loc (loc, "Tacinterp.interp_app",
- (str"Illegal tactic application.")) in
+and interp_app loc ist fv largs =
+ let fail =
+ (* spiwack: quick hack, can be inlined. *)
+ try
+ user_err_loc (loc, "Tacinterp.interp_app",
+ (str"Illegal tactic application."))
+ with e -> Proofview.tclZERO e
+ in
let fv = Value.normalize fv in
if has_type fv (topwit wit_tacvalue) then
match to_tacvalue fv with
@@ -1226,72 +1230,78 @@ and interp_app loc ist gl fv largs =
let (extfun,lvar,lval)=head_with_value (var,largs) in
let fold accu (id, v) = Id.Map.add id v accu in
let newlfun = List.fold_left fold olfun extfun in
- if List.is_empty lvar then
- (* Check evaluation and report problems with current trace *)
- let (sigma,v) =
- try
+ if List.is_empty lvar then
+ begin Proofview.tclORELSE
+ begin
let ist = {
lfun = newlfun;
extra = TacStore.set ist.extra f_trace []; } in
- catch_error trace (val_interp ist gl) body
- with reraise ->
- let reraise = Errors.push reraise in
- debugging_exception_step ist false reraise
- (fun () -> str "evaluation");
- raise reraise
- in
+ catch_error_tac trace (val_interp ist body)
+ end
+ begin fun e ->
+ (* spiwack: [Errors.push] here is unlikely to do what
+ it's intended to, or anything meaningful for that
+ matter. *)
+ let e = Errors.push e in
+ (* arnaud: todo: debugging: debugging_exception_step ist false e (fun () -> str "evaluation"); *)
+ Proofview.tclZERO e
+ end
+ end >>== fun v ->
+ (* arnaud: todo: debugging:
(* No errors happened, we propagate the trace *)
let v = append_trace trace v in
-
- let gl = { gl with sigma=sigma } in
- debugging_step ist
- (fun () ->
- str"evaluation returns"++fnl()++pr_value (Some (pf_env gl)) v);
- if List.is_empty lval then sigma,v else interp_app loc ist gl v lval
- else
- project gl , of_tacvalue (VFun(trace,newlfun,lvar,body))
- | _ -> fail ()
- else fail ()
+ debugging_step ist
+ (fun () ->
+ str"evaluation returns"++fnl()++pr_value (Some (pf_env gl)) v);
+ *)
+ if List.is_empty lval then Proofview.tclUNIT (Goal.return v) else interp_app loc ist v lval
+ else
+ Proofview.tclUNIT (Goal.return (of_tacvalue (VFun(trace,newlfun,lvar,body))))
+ | _ -> fail
+ else fail
(* Gives the tactic corresponding to the tactic value *)
-and tactic_of_value ist vle g =
+and tactic_of_value ist vle =
let vle = Value.normalize vle in
if has_type vle (topwit wit_tacvalue) then
match to_tacvalue vle with
- | VRTactic res -> res
+ | VRTactic -> Proofview.tclUNIT ()
| VFun (trace,lfun,[],t) ->
let ist = {
lfun = lfun;
extra = TacStore.set ist.extra f_trace []; } in
let tac = eval_tactic ist t in
- catch_error trace tac g
- | (VFun _|VRec _) -> error "A fully applied tactic is expected."
- else errorlabstrm "" (str"Expression does not evaluate to a tactic.")
+ catch_error_tac trace tac
+ | (VFun _|VRec _) -> Proofview.tclZERO (UserError ("" , str "A fully applied tactic is expected."))
+ else Proofview.tclZERO (UserError ("" , str"Expression does not evaluate to a tactic."))
(* Evaluation with FailError catching *)
-and eval_with_fail ist is_lazy goal tac =
- try
- let (sigma,v) = val_interp ist goal tac in
- let v = Value.normalize v in
- sigma ,
- (if has_type v (topwit wit_tacvalue) then match to_tacvalue v with
- | VFun (trace,lfun,[],t) when not is_lazy ->
- let ist = {
- lfun = lfun;
- extra = TacStore.set ist.extra f_trace trace; } in
- let tac = eval_tactic ist t in
- of_tacvalue (VRTactic (catch_error trace tac { goal with sigma=sigma }))
- | _ -> v
- else v)
- with
- (** FIXME: Should we add [Errors.push]? *)
- | FailError (0,s) ->
- raise (Eval_fail (Lazy.force s))
- | FailError (lvl,s) as e ->
- raise (Exninfo.copy e (FailError (lvl - 1, s)))
+and eval_with_fail ist is_lazy tac =
+ Proofview.tclORELSE
+ begin
+ val_interp ist tac >>== fun v ->
+ (if has_type v (topwit wit_tacvalue) then match to_tacvalue v with
+ | VFun (trace,lfun,[],t) when not is_lazy ->
+ let ist = {
+ lfun = lfun;
+ extra = TacStore.set ist.extra f_trace trace; } in
+ let tac = eval_tactic ist t in
+ catch_error_tac trace (tac <*> Proofview.tclUNIT (Goal.return (of_tacvalue VRTactic)))
+ | _ -> Proofview.tclUNIT (Goal.return v)
+ else Proofview.tclUNIT (Goal.return v))
+ end
+ begin function
+ (** FIXME: Should we add [Errors.push]? *)
+ | FailError (0,s) ->
+ Proofview.tclZERO (Eval_fail (Lazy.force s))
+ | FailError (lvl,s) as e ->
+ Proofview.tclZERO (Exninfo.copy e (FailError (lvl - 1, s)))
+ | e -> Proofview.tclZERO e
+ end
(* Interprets the clauses of a recursive LetIn *)
-and interp_letrec ist gl llc u =
+and interp_letrec ist llc u =
+ Proofview.tclUNIT () >= fun () -> (* delay for the effects of [lref], just in case. *)
let lref = ref ist.lfun in
let fold accu ((_, id), b) =
let v = of_tacvalue (VRec (lref, TacArg (dloc, b))) in
@@ -1300,83 +1310,115 @@ and interp_letrec ist gl llc u =
let lfun = List.fold_left fold ist.lfun llc in
let () = lref := lfun in
let ist = { ist with lfun } in
- val_interp ist gl u
+ val_interp ist u
(* Interprets the clauses of a LetIn *)
-and interp_letin ist gl llc u =
- let fold ((_, id), body) (sigma, acc) =
- let (sigma, v) = interp_tacarg ist { gl with sigma } body in
+and interp_letin ist llc u =
+ let fold ((_, id), body) acc =
+ interp_tacarg ist body >>== fun v ->
+ acc >>== fun acc ->
let () = check_is_value v in
- sigma, Id.Map.add id v acc
+ Proofview.tclUNIT (Goal.return (Id.Map.add id v acc))
in
- let (sigma, lfun) = List.fold_right fold llc (project gl, ist.lfun) in
+ List.fold_right fold llc (Proofview.tclUNIT (Goal.return ist.lfun)) >>== fun lfun ->
let ist = { ist with lfun } in
- val_interp ist { gl with sigma } u
+ val_interp ist u
(* Interprets the Match Context expressions *)
-and interp_match_goal ist goal lz lr lmr =
+and interp_match_goal ist lz lr lmr =
+ (* arnaud: on va prier pour que je n'ai pas besoin de faire ça,
+ sinon, je fais un truc ad-hoc
let (gl,sigma) = Goal.V82.nf_evar (project goal) (sig_it goal) in
let goal = { it = gl ; sigma = sigma; } in
- let hyps = pf_hyps goal in
+ *)
+ Goal.hyps >>-- fun hyps ->
+ let hyps = Environ.named_context_of_val hyps in
let hyps = if lr then List.rev hyps else hyps in
- let concl = pf_concl goal in
- let env = pf_env goal in
- let apply_goal_sub app ist (id,c) csr mt mhyps hyps =
- let matches = match_subterm_gen app c csr in
+ Goal.concl >>-- fun concl ->
+ Goal.env >>-- fun env ->
+ let rec apply_goal_sub app ist (id,c) csr mt mhyps hyps =
let rec match_next_pattern next = match IStream.peek next with
- | None -> raise PatternMatchingFailure
+ | None -> Proofview.tclZERO PatternMatchingFailure
| Some ({ m_sub=lgoal; m_ctx=ctxt }, next) ->
let lctxt = give_context ctxt id in
- try apply_hyps_context ist env lz goal mt lctxt (adjust lgoal) mhyps hyps
- with e when is_match_catchable e -> match_next_pattern next in
- match_next_pattern matches
- in
- let rec apply_match_goal ist env goal nrs lex lpt =
+ Proofview.tclORELSE
+ (apply_hyps_context ist env lz mt lctxt (adjust lgoal) mhyps hyps)
+ begin function
+ | e when is_match_catchable e -> match_next_pattern next
+ | e -> Proofview.tclZERO e
+ end
+ in
+ match_next_pattern (match_subterm_gen app c csr) in
+ let rec apply_match_goal ist env nrs lex lpt =
begin
let () = match lex with
| r :: _ -> db_pattern_rule (curr_debug ist) nrs r
| _ -> ()
in
match lpt with
- | (All t)::tl ->
- begin
- db_mc_pattern_success (curr_debug ist);
- try eval_with_fail ist lz goal t
- with e when is_match_catchable e ->
- apply_match_goal ist env goal (nrs+1) (List.tl lex) tl
- end
- | (Pat (mhyps,mgoal,mt))::tl ->
- let mhyps = List.rev mhyps (* Sens naturel *) in
- (match mgoal with
- | Term mg ->
- (try
- let lmatch = extended_matches mg concl in
- db_matched_concl (curr_debug ist) env concl;
- apply_hyps_context ist env lz goal mt Id.Map.empty lmatch mhyps hyps
- with e when is_match_catchable e ->
- (match e with
- | PatternMatchingFailure -> db_matching_failure (curr_debug ist)
- | Eval_fail s -> db_eval_failure (curr_debug ist) s
- | _ -> db_logic_failure (curr_debug ist) e);
- apply_match_goal ist env goal (nrs+1) (List.tl lex) tl)
- | Subterm (b,id,mg) ->
- (try apply_goal_sub b ist (id,mg) concl mt mhyps hyps
- with
- | PatternMatchingFailure ->
- apply_match_goal ist env goal (nrs+1) (List.tl lex) tl))
- | _ ->
- errorlabstrm "Tacinterp.apply_match_goal"
- (v 0 (str "No matching clauses for match goal" ++
- (if (curr_debug ist) == DebugOff then
- fnl() ++ str "(use \"Set Ltac Debug\" for more info)"
- else mt()) ++ str"."))
+ | (All t)::tl ->
+ begin
+ db_mc_pattern_success (curr_debug ist);
+ Proofview.tclORELSE (eval_with_fail ist lz t)
+ begin function
+ | e when is_match_catchable e ->
+ apply_match_goal ist env (nrs+1) (List.tl lex) tl
+ | e -> Proofview.tclZERO e
+ end
+ end
+ | (Pat (mhyps,mgoal,mt))::tl ->
+ let mhyps = List.rev mhyps (* Sens naturel *) in
+ begin match mgoal with
+ | Term mg ->
+ let matches =
+ try Some (extended_matches mg concl)
+ with PatternMatchingFailure -> None
+ in
+ begin match matches with
+ | None ->
+ db_matching_failure (curr_debug ist);
+ apply_match_goal ist env (nrs+1) (List.tl lex) tl
+ | Some lmatch ->
+ Proofview.tclORELSE
+ begin
+ db_matched_concl (curr_debug ist) env concl;
+ apply_hyps_context ist env lz mt Id.Map.empty lmatch mhyps hyps
+ end
+ begin function
+ | e when is_match_catchable e ->
+ ((match e with
+ | PatternMatchingFailure -> db_matching_failure (curr_debug ist)
+ | Eval_fail s -> db_eval_failure (curr_debug ist) s
+ | _ -> db_logic_failure (curr_debug ist) e);
+ apply_match_goal ist env (nrs+1) (List.tl lex) tl)
+ | e -> Proofview.tclZERO e
+ end
+ end
+ | Subterm (b,id,mg) ->
+ Proofview.tclORELSE (apply_goal_sub b ist (id,mg) concl mt mhyps hyps)
+ begin function
+ | PatternMatchingFailure ->
+ apply_match_goal ist env (nrs+1) (List.tl lex) tl
+ | e -> Proofview.tclZERO e
+ end
+ end
+ | _ ->
+ Proofview.tclZERO (UserError (
+ "Tacinterp.apply_match_goal" ,
+ (v 0 (str "No matching clauses for match goal" ++
+ (if curr_debug ist==DebugOff then
+ fnl() ++ str "(use \"Set Ltac Debug\" for more info)"
+ else mt()) ++ str"."))
+ ))
end in
- apply_match_goal ist env goal 0 lmr
- (read_match_rule (extract_ltac_constr_values ist env)
- ist env (project goal) lmr)
+ Proofview.tclEVARMAP >= fun sigma ->
+ apply_match_goal ist env 0 lmr
+ (read_match_rule (extract_ltac_constr_values ist env)
+ ist env sigma lmr)
(* Tries to match the hypotheses in a Match Context *)
-and apply_hyps_context ist env lz goal mt lctxt lgmatch mhyps hyps =
+and apply_hyps_context ist env lz mt lctxt lgmatch mhyps hyps =
+ Goal.env >>-- fun env ->
let rec apply_hyps_context_rec lfun lmatch lhyps_rest = function
| hyp_pat::tl ->
let (hypname, _, pat as hyp_pat) =
@@ -1387,34 +1429,43 @@ and apply_hyps_context ist env lz goal mt lctxt lgmatch mhyps hyps =
let rec match_next_pattern next = match IStream.peek next with
| None ->
db_hyp_pattern_failure (curr_debug ist) env (hypname, pat);
- raise PatternMatchingFailure
+ Proofview.tclZERO PatternMatchingFailure
| Some (s, next) ->
- let lids,hyp_match = s.e_ctx in
- db_matched_hyp (curr_debug ist) (pf_env goal) hyp_match hypname;
- try
- let id_match = pi1 hyp_match in
- let select_match (id,_,_) = Id.equal id id_match in
- let nextlhyps = List.remove_first select_match lhyps_rest in
- let lfun = lfun +++ lids in
- apply_hyps_context_rec lfun s.e_sub nextlhyps tl
- with e when is_match_catchable e ->
- match_next_pattern next in
- let init_match_pattern = apply_one_mhyp_context goal lmatch hyp_pat lhyps_rest in
- match_next_pattern init_match_pattern
+ let lids, hyp_match = s.e_ctx in
+ db_matched_hyp (curr_debug ist) env hyp_match hypname;
+ Proofview.tclORELSE
+ begin
+ let id_match = pi1 hyp_match in
+ let select_match (id,_,_) = Id.equal id id_match in
+ let nextlhyps = List.remove_first select_match lhyps_rest in
+ let lfun = lfun +++ lids in
+ apply_hyps_context_rec lfun s.e_sub nextlhyps tl
+ end
+ begin function
+ | e when is_match_catchable e ->
+ match_next_pattern next
+ | e -> Proofview.tclZERO e
+ end
+ in
+ let init_match_pattern = Tacmach.New.of_old (fun gl ->
+ apply_one_mhyp_context gl lmatch hyp_pat lhyps_rest) in
+ init_match_pattern >>-- match_next_pattern
| [] ->
let lfun = lfun +++ ist.lfun in
let lfun = extend_values_with_bindings lmatch lfun in
db_mc_pattern_success (curr_debug ist);
- eval_with_fail { ist with lfun } lz goal mt
+ eval_with_fail {ist with lfun } lz mt
in
apply_hyps_context_rec lctxt lgmatch hyps mhyps
-and interp_external loc ist gl com req la =
+and interp_external loc ist gl com req la = assert false
+(* arnaud: todo
let f ch = extern_request ch req gl la in
let g ch = internalise_tacarg ch in
- interp_tacarg ist gl (System.connect f g com)
+ interp_tacarg ist (System.connect f g com)
+*)
- (* Interprets extended tactic generic arguments *)
+(* Interprets extended tactic generic arguments *)
and interp_genarg ist gl x =
let evdref = ref (project gl) in
let rec interp_genarg ist gl x =
@@ -1489,206 +1540,293 @@ and interp_genarg_var_list ist gl x =
in_gen (topwit (wit_list wit_var)) lc
(* Interprets the Match expressions *)
-and interp_match ist g lz constr lmr =
+and interp_match ist lz constr lmr =
let apply_match_subterm app ist (id,c) csr mt =
let rec match_next_pattern next = match IStream.peek next with
- | None -> raise PatternMatchingFailure
+ | None -> Proofview.tclZERO PatternMatchingFailure
| Some ({ m_sub=lmatch; m_ctx=ctxt; }, next) ->
let lctxt = give_context ctxt id in
let lfun = extend_values_with_bindings (adjust lmatch) (lctxt +++ ist.lfun) in
- try eval_with_fail {ist with lfun=lfun} lz g mt
- with e when is_match_catchable e -> match_next_pattern next in
+ Proofview.tclORELSE
+ (eval_with_fail {ist with lfun=lfun} lz mt)
+ begin function
+ | e when is_match_catchable e ->
+ match_next_pattern next
+ | e -> Proofview.tclZERO e
+ end
+ in
match_next_pattern (match_subterm_gen app c csr) in
- let rec apply_match ist sigma csr = let g = { g with sigma=sigma } in function
+
+ let rec apply_match ist csr = function
| (All t)::tl ->
- (try eval_with_fail ist lz g t
- with e when is_match_catchable e -> apply_match ist sigma csr tl)
+ Proofview.tclORELSE
+ (eval_with_fail ist lz t)
+ begin function
+ | e when is_match_catchable e -> apply_match ist csr tl
+ | e -> Proofview.tclZERO e
+ end
| (Pat ([],Term c,mt))::tl ->
- (try
- let lmatch =
- try extended_matches c csr
- with reraise ->
- let reraise = Errors.push reraise in
- debugging_exception_step ist false reraise (fun () ->
- str "matching with pattern" ++ fnl () ++
- pr_constr_pattern_env (pf_env g) c);
- raise reraise
- in
- try
- let lfun = extend_values_with_bindings lmatch ist.lfun in
- eval_with_fail { ist with lfun=lfun } lz g mt
- with reraise ->
- let reraise = Errors.push reraise in
- debugging_exception_step ist false reraise (fun () ->
- str "rule body for pattern" ++
- pr_constr_pattern_env (pf_env g) c);
- raise reraise
- with e when is_match_catchable e ->
- debugging_step ist (fun () -> str "switching to the next rule");
- apply_match ist sigma csr tl)
+ let matches =
+ try Some (extended_matches c csr)
+ with PatternMatchingFailure -> None
+ in
+ Proofview.tclORELSE begin match matches with
+ | None -> Proofview.tclZERO PatternMatchingFailure
+ | Some lmatch ->
+ Proofview.tclORELSE
+ begin
+ let lfun = extend_values_with_bindings lmatch ist.lfun in
+ eval_with_fail { ist with lfun=lfun } lz mt
+ end
+ begin function
+ | e ->
+ (* arnaud: todo: debugging
+ debugging_exception_step ist false e (fun () ->
+ str "rule body for pattern" ++
+ pr_constr_pattern_env (pf_env g) c);*)
+ Proofview.tclZERO e
+ end
+ end
+ begin function
+ | e when is_match_catchable e ->
+ debugging_step ist (fun () -> str "switching to the next rule");
+ apply_match ist csr tl
+ | e -> Proofview.tclZERO e
+ end
| (Pat ([],Subterm (b,id,c),mt))::tl ->
- (try apply_match_subterm b ist (id,c) csr mt
- with PatternMatchingFailure -> apply_match ist sigma csr tl)
+ Proofview.tclORELSE
+ (apply_match_subterm b ist (id,c) csr mt)
+ begin function
+ | PatternMatchingFailure -> apply_match ist csr tl
+ | e -> Proofview.tclZERO e
+ end
| _ ->
- errorlabstrm "Tacinterp.apply_match" (str
- "No matching clauses for match.") in
- let (sigma,csr) =
- try interp_ltac_constr ist g constr
- with reraise ->
- let reraise = Errors.push reraise in
- debugging_exception_step ist true reraise
- (fun () -> str "evaluation of the matched expression");
- raise reraise
- in
- let ilr = read_match_rule (extract_ltac_constr_values ist (pf_env g)) ist (pf_env g) sigma lmr in
- let res =
- try apply_match ist sigma csr ilr
- with reraise ->
- let reraise = Errors.push reraise in
- debugging_exception_step ist true reraise
- (fun () -> str "match expression");
- raise reraise
- in
+ Proofview.tclZERO (UserError ("Tacinterp.apply_match" , str
+ "No matching clauses for match.")) in
+ begin Proofview.tclORELSE
+ (interp_ltac_constr ist constr)
+ begin function
+ | e ->
+ (* spiwack: [Errors.push] here is unlikely to do what
+ it's intended to, or anything meaningful for that
+ matter. *)
+ let e = Errors.push e in
+ debugging_exception_step ist true e
+ (fun () -> str "evaluation of the matched expression");
+ Proofview.tclZERO e
+ end
+ end >>== fun csr ->
+ Goal.env >>-- fun env ->
+ Proofview.tclEVARMAP >= fun sigma ->
+ let ilr = read_match_rule (extract_ltac_constr_values ist env) ist env sigma lmr in
+ Proofview.tclORELSE
+ (apply_match ist csr ilr)
+ begin function
+ | e ->
+ debugging_exception_step ist true e (fun () -> str "match expression");
+ Proofview.tclZERO e
+ end >>== fun res ->
debugging_step ist (fun () ->
- str "match expression returns " ++ pr_value (Some (pf_env g)) (snd res));
- res
+ str "match expression returns " ++ pr_value (Some env) res);
+ Proofview.tclUNIT (Goal.return res)
(* Interprets tactic expressions : returns a "constr" *)
-and interp_ltac_constr ist gl e =
- let (sigma, result) =
- try val_interp ist gl e with Not_found ->
- debugging_step ist (fun () ->
- str "evaluation failed for" ++ fnl() ++
- Pptactic.pr_glob_tactic (pf_env gl) e);
- raise Not_found in
+and interp_ltac_constr ist e =
+ begin Proofview.tclORELSE
+ (val_interp ist e)
+ begin function
+ | Not_found ->
+ (* arnaud: todo: debugging
+ debugging_step ist (fun () ->
+ str "evaluation failed for" ++ fnl() ++
+ Pptactic.pr_glob_tactic (pf_env gl) e);
+ *)
+ Proofview.tclZERO Not_found
+ | e -> Proofview.tclZERO e
+ end
+ end >>== fun result ->
+ Goal.env >>-- fun env ->
let result = Value.normalize result in
try
- let cresult = coerce_to_closed_constr (pf_env gl) result in
+ let cresult = coerce_to_closed_constr env result in
debugging_step ist (fun () ->
- Pptactic.pr_glob_tactic (pf_env gl) e ++ fnl() ++
- str " has value " ++ fnl() ++
- pr_constr_env (pf_env gl) cresult);
- sigma, cresult
+ Pptactic.pr_glob_tactic env e ++ fnl() ++
+ str " has value " ++ fnl() ++
+ pr_constr_env env cresult);
+ Proofview.tclUNIT (Goal.return cresult)
with CannotCoerceTo _ ->
- errorlabstrm ""
+ Goal.env >>-- fun env ->
+ Proofview.tclZERO (UserError ( "",
+ errorlabstrm ""
(str "Must evaluate to a closed term" ++ fnl() ++
- str "offending expression: " ++ fnl() ++ pr_inspect (pf_env gl) e result)
+ str "offending expression: " ++ fnl() ++ pr_inspect env e result)))
(* Interprets tactic expressions : returns a "tactic" *)
-and interp_tactic ist tac gl =
- let (sigma,v) = val_interp ist gl tac in
- tactic_of_value ist v { gl with sigma=sigma }
+and interp_tactic ist tac =
+ val_interp ist tac >>= fun v ->
+ tactic_of_value ist v
(* Interprets a primitive tactic *)
-and interp_atomic ist gl tac =
- let env = pf_env gl and sigma = project gl in
+and interp_atomic ist tac =
match tac with
(* Basic tactics *)
| TacIntroPattern l ->
- h_intro_patterns (interp_intro_pattern_list_as_list ist gl l)
+ Tacmach.New.of_old (fun gl -> interp_intro_pattern_list_as_list ist gl l) >>- fun patterns ->
+ h_intro_patterns patterns
| TacIntrosUntil hyp ->
h_intros_until (interp_quantified_hypothesis ist hyp)
| TacIntroMove (ido,hto) ->
- h_intro_move (Option.map (interp_fresh_ident ist env) ido)
- (interp_move_location ist gl hto)
- | TacAssumption -> h_assumption
+ Goal.env >>- fun env ->
+ Tacmach.New.of_old (fun gl -> interp_move_location ist gl hto) >>- fun mloc ->
+ h_intro_move (Option.map (interp_fresh_ident ist env) ido) mloc
+ | TacAssumption -> Proofview.V82.tactic h_assumption
| TacExact c ->
- let (sigma,c_interp) = pf_interp_casted_constr ist gl c in
- tclTHEN
- (tclEVARS sigma)
- (h_exact c_interp)
+ Proofview.V82.tactic begin fun gl ->
+ let (sigma,c_interp) = pf_interp_casted_constr ist gl c in
+ tclTHEN
+ (tclEVARS sigma)
+ (h_exact c_interp)
+ gl
+ end
| TacExactNoCheck c ->
- let (sigma,c_interp) = pf_interp_constr ist gl c in
- tclTHEN
- (tclEVARS sigma)
- (h_exact_no_check c_interp)
+ Proofview.V82.tactic begin fun gl ->
+ let (sigma,c_interp) = pf_interp_constr ist gl c in
+ tclTHEN
+ (tclEVARS sigma)
+ (h_exact_no_check c_interp)
+ gl
+ end
| TacVmCastNoCheck c ->
- let (sigma,c_interp) = pf_interp_constr ist gl c in
- tclTHEN
- (tclEVARS sigma)
- (h_vm_cast_no_check c_interp)
+ Proofview.V82.tactic begin fun gl ->
+ let (sigma,c_interp) = pf_interp_constr ist gl c in
+ tclTHEN
+ (tclEVARS sigma)
+ (h_vm_cast_no_check c_interp)
+ gl
+ end
| TacApply (a,ev,cb,cl) ->
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
let sigma, l =
List.fold_map (interp_open_constr_with_bindings_loc ist env) sigma cb
in
let tac = match cl with
- | None -> h_apply a ev
+ | None -> fun l -> Proofview.V82.tactic (h_apply a ev l)
| Some cl ->
- (fun l -> h_apply_in a ev l (interp_in_hyp_as ist gl cl)) in
- tclWITHHOLES ev tac sigma l
+ (fun l ->
+ Tacmach.New.of_old (fun gl -> interp_in_hyp_as ist gl cl) >>- fun cl ->
+ h_apply_in a ev l cl) in
+ Tacticals.New.tclWITHHOLES ev tac sigma l
| TacElim (ev,cb,cbo) ->
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
let sigma, cb = interp_constr_with_bindings ist env sigma cb in
let sigma, cbo = Option.fold_map (interp_constr_with_bindings ist env) sigma cbo in
- tclWITHHOLES ev (h_elim ev cb) sigma cbo
+ Tacticals.New.tclWITHHOLES ev (h_elim ev cb) sigma cbo
| TacElimType c ->
- let (sigma,c_interp) = pf_interp_type ist gl c in
- tclTHEN
- (tclEVARS sigma)
- (h_elim_type c_interp)
+ Proofview.V82.tactic begin fun gl ->
+ let (sigma,c_interp) = pf_interp_type ist gl c in
+ tclTHEN
+ (tclEVARS sigma)
+ (h_elim_type c_interp)
+ gl
+ end
| TacCase (ev,cb) ->
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
let sigma, cb = interp_constr_with_bindings ist env sigma cb in
- tclWITHHOLES ev (h_case ev) sigma cb
+ Tacticals.New.tclWITHHOLES ev (h_case ev) sigma cb
| TacCaseType c ->
- let (sigma,c_interp) = pf_interp_type ist gl c in
- tclTHEN
- (tclEVARS sigma)
- (h_case_type c_interp)
- | TacFix (idopt,n) -> h_fix (Option.map (interp_fresh_ident ist env) idopt) n
+ Proofview.V82.tactic begin fun gl ->
+ let (sigma,c_interp) = pf_interp_type ist gl c in
+ tclTHEN
+ (tclEVARS sigma)
+ (h_case_type c_interp)
+ gl
+ end
+ | TacFix (idopt,n) ->
+ Goal.env >>- fun env ->
+ Proofview.V82.tactic (h_fix (Option.map (interp_fresh_ident ist env) idopt) n)
| TacMutualFix (id,n,l) ->
- let f sigma (id,n,c) =
- let (sigma,c_interp) = pf_interp_type ist { gl with sigma=sigma } c in
- sigma , (interp_fresh_ident ist env id,n,c_interp) in
- let (sigma,l_interp) =
- List.fold_right begin fun c (sigma,acc) ->
- let (sigma,c_interp) = f sigma c in
- sigma , c_interp::acc
- end l (project gl,[])
- in
- tclTHEN
- (tclEVARS sigma)
- (h_mutual_fix (interp_fresh_ident ist env id) n l_interp)
- | TacCofix idopt -> h_cofix (Option.map (interp_fresh_ident ist env) idopt)
+ Goal.env >>- fun env ->
+ Proofview.V82.tactic begin fun gl ->
+ let f sigma (id,n,c) =
+ let (sigma,c_interp) = pf_interp_type ist { gl with sigma=sigma } c in
+ sigma , (interp_fresh_ident ist env id,n,c_interp) in
+ let (sigma,l_interp) =
+ List.fold_right begin fun c (sigma,acc) ->
+ let (sigma,c_interp) = f sigma c in
+ sigma , c_interp::acc
+ end l (project gl,[])
+ in
+ tclTHEN
+ (tclEVARS sigma)
+ (h_mutual_fix (interp_fresh_ident ist env id) n l_interp)
+ gl
+ end
+ | TacCofix idopt ->
+ Goal.env >>- fun env ->
+ Proofview.V82.tactic (h_cofix (Option.map (interp_fresh_ident ist env) idopt))
| TacMutualCofix (id,l) ->
- let f sigma (id,c) =
- let (sigma,c_interp) = pf_interp_type ist { gl with sigma=sigma } c in
- sigma , (interp_fresh_ident ist env id,c_interp) in
- let (sigma,l_interp) =
- List.fold_right begin fun c (sigma,acc) ->
- let (sigma,c_interp) = f sigma c in
- sigma , c_interp::acc
- end l (project gl,[])
- in
- tclTHEN
- (tclEVARS sigma)
- (h_mutual_cofix (interp_fresh_ident ist env id) l_interp)
+ Goal.env >>- fun env ->
+ Proofview.V82.tactic begin fun gl ->
+ let f sigma (id,c) =
+ let (sigma,c_interp) = pf_interp_type ist { gl with sigma=sigma } c in
+ sigma , (interp_fresh_ident ist env id,c_interp) in
+ let (sigma,l_interp) =
+ List.fold_right begin fun c (sigma,acc) ->
+ let (sigma,c_interp) = f sigma c in
+ sigma , c_interp::acc
+ end l (project gl,[])
+ in
+ tclTHEN
+ (tclEVARS sigma)
+ (h_mutual_cofix (interp_fresh_ident ist env id) l_interp)
+ gl
+ end
| TacCut c ->
- let (sigma,c_interp) = pf_interp_type ist gl c in
- tclTHEN
- (tclEVARS sigma)
- (h_cut c_interp)
+ Proofview.V82.tactic begin fun gl ->
+ let (sigma,c_interp) = pf_interp_type ist gl c in
+ tclTHEN
+ (tclEVARS sigma)
+ (h_cut c_interp)
+ gl
+ end
| TacAssert (t,ipat,c) ->
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
let (sigma,c) = (if Option.is_empty t then interp_constr else interp_type) ist env sigma c in
- tclTHEN
- (tclEVARS sigma)
+ Tacmach.New.of_old (fun gl -> interp_intro_pattern ist gl) >>- fun patt ->
+ Tacticals.New.tclTHEN
+ (Proofview.V82.tactic (tclEVARS sigma))
(Tactics.forward (Option.map (interp_tactic ist) t)
- (Option.map (interp_intro_pattern ist gl) ipat) c)
+ (Option.map patt ipat) c)
| TacGeneralize cl ->
- let sigma, cl = interp_constr_with_occurrences_and_name_as_list ist env sigma cl in
- tclWITHHOLES false (h_generalize_gen) sigma cl
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
+ Proofview.V82.tactic begin fun gl ->
+ let sigma, cl = interp_constr_with_occurrences_and_name_as_list ist env sigma cl in
+ tclWITHHOLES false (h_generalize_gen) sigma cl gl
+ end
| TacGeneralizeDep c ->
- let (sigma,c_interp) = pf_interp_constr ist gl c in
- tclTHEN
- (tclEVARS sigma)
- (h_generalize_dep c_interp)
+ Proofview.V82.tactic begin fun gl ->
+ let (sigma,c_interp) = pf_interp_constr ist gl c in
+ tclTHEN
+ (tclEVARS sigma)
+ (h_generalize_dep c_interp)
+ gl
+ end
| TacLetTac (na,c,clp,b,eqpat) ->
- let clp = interp_clause ist gl clp in
- let eqpat = Option.map (interp_intro_pattern ist gl) eqpat in
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
+ Tacmach.New.of_old (fun gl -> interp_clause ist gl clp) >>- fun clp ->
+ Tacmach.New.of_old (fun gl -> Option.map (interp_intro_pattern ist gl) eqpat) >>- fun eqpat ->
if Locusops.is_nowhere clp then
(* We try to fully-typecheck the term *)
- let (sigma,c_interp) = pf_interp_constr ist gl c in
- tclTHEN
- (tclEVARS sigma)
+ Tacmach.New.of_old (fun gl -> pf_interp_constr ist gl c) >>- fun (sigma,c_interp) ->
+ Tacticals.New.tclTHEN
+ (Proofview.V82.tactic (tclEVARS sigma))
(h_let_tac b (interp_fresh_name ist env na) c_interp clp eqpat)
else
(* We try to keep the pattern structure as much as possible *)
@@ -1697,10 +1835,14 @@ and interp_atomic ist gl tac =
(* Automation tactics *)
| TacTrivial (debug,lems,l) ->
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
Auto.h_trivial ~debug
(interp_auto_lemmas ist env sigma lems)
(Option.map (List.map (interp_hint_base ist)) l)
| TacAuto (debug,n,lems,l) ->
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
Auto.h_auto ~debug (Option.map (interp_int_or_var ist) n)
(interp_auto_lemmas ist env sigma lems)
(Option.map (List.map (interp_hint_base ist)) l)
@@ -1709,241 +1851,321 @@ and interp_atomic ist gl tac =
| TacSimpleInductionDestruct (isrec,h) ->
h_simple_induction_destruct isrec (interp_quantified_hypothesis ist h)
| TacInductionDestruct (isrec,ev,(l,el,cls)) ->
- let sigma, l =
- List.fold_map (fun sigma (c,(ipato,ipats)) ->
- let c = interp_induction_arg ist gl c in
- (sigma,(c,
- (Option.map (interp_intro_pattern ist gl) ipato,
- Option.map (interp_intro_pattern ist gl) ipats)))) sigma l in
+ Goal.env >>- fun env ->
+ let l =
+ Goal.sensitive_list_map begin fun (c,(ipato,ipats)) ->
+ Tacmach.New.of_old (fun gl -> interp_induction_arg ist gl c) >- fun c ->
+ Tacmach.New.of_old (fun gl -> interp_intro_pattern ist gl) >- fun interp_intro_pattern ->
+ Goal.return begin
+ c,
+ (Option.map interp_intro_pattern ipato,
+ Option.map interp_intro_pattern ipats)
+ end
+ end l
+ in
+ l >>- fun l ->
+ Goal.defs >>- fun sigma ->
let sigma,el =
Option.fold_map (interp_constr_with_bindings ist env) sigma el in
- let cls = Option.map (interp_clause ist gl) cls in
- tclWITHHOLES ev (h_induction_destruct isrec ev) sigma (l,el,cls)
+ Tacmach.New.of_old (fun gl -> interp_clause ist gl) >>- fun interp_clause ->
+ let cls = Option.map interp_clause cls in
+ Tacticals.New.tclWITHHOLES ev (h_induction_destruct isrec ev) sigma (l,el,cls)
| TacDoubleInduction (h1,h2) ->
let h1 = interp_quantified_hypothesis ist h1 in
let h2 = interp_quantified_hypothesis ist h2 in
Elim.h_double_induction h1 h2
| TacDecomposeAnd c ->
- let (sigma,c_interp) = pf_interp_constr ist gl c in
- tclTHEN
- (tclEVARS sigma)
+ Tacmach.New.of_old (fun gl -> pf_interp_constr ist gl c) >>- fun (sigma,c_interp) ->
+ Tacticals.New.tclTHEN
+ (Proofview.V82.tactic (tclEVARS sigma))
(Elim.h_decompose_and c_interp)
| TacDecomposeOr c ->
- let (sigma,c_interp) = pf_interp_constr ist gl c in
- tclTHEN
- (tclEVARS sigma)
+ Tacmach.New.of_old (fun gl -> pf_interp_constr ist gl c) >>- fun (sigma,c_interp) ->
+ Tacticals.New.tclTHEN
+ (Proofview.V82.tactic (tclEVARS sigma))
(Elim.h_decompose_or c_interp)
| TacDecompose (l,c) ->
let l = List.map (interp_inductive ist) l in
- let (sigma,c_interp) = pf_interp_constr ist gl c in
- tclTHEN
- (tclEVARS sigma)
+ Tacmach.New.of_old (fun gl -> pf_interp_constr ist gl c) >>- fun (sigma,c_interp) ->
+ Tacticals.New.tclTHEN
+ (Proofview.V82.tactic (tclEVARS sigma))
(Elim.h_decompose l c_interp)
| TacSpecialize (n,cb) ->
- let sigma, cb = interp_constr_with_bindings ist env sigma cb in
- tclWITHHOLES false (h_specialize n) sigma cb
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
+ Proofview.V82.tactic begin fun gl ->
+ let sigma, cb = interp_constr_with_bindings ist env sigma cb in
+ tclWITHHOLES false (h_specialize n) sigma cb gl
+ end
| TacLApply c ->
- let (sigma,c_interp) = pf_interp_constr ist gl c in
- tclTHEN
- (tclEVARS sigma)
- (h_lapply c_interp)
+ Proofview.V82.tactic begin fun gl ->
+ let (sigma,c_interp) = pf_interp_constr ist gl c in
+ tclTHEN
+ (tclEVARS sigma)
+ (h_lapply c_interp)
+ gl
+ end
(* Context management *)
- | TacClear (b,l) -> h_clear b (interp_hyp_list ist gl l)
- | TacClearBody l -> h_clear_body (interp_hyp_list ist gl l)
+ | TacClear (b,l) ->
+ Proofview.V82.tactic begin fun gl ->
+ h_clear b (interp_hyp_list ist gl l) gl
+ end
+ | TacClearBody l ->
+ Proofview.V82.tactic begin fun gl ->
+ h_clear_body (interp_hyp_list ist gl l) gl
+ end
| TacMove (dep,id1,id2) ->
- h_move dep (interp_hyp ist gl id1) (interp_move_location ist gl id2)
+ Proofview.V82.tactic begin fun gl ->
+ h_move dep (interp_hyp ist gl id1) (interp_move_location ist gl id2) gl
+ end
| TacRename l ->
- h_rename (List.map (fun (id1,id2) ->
- interp_hyp ist gl id1,
- interp_fresh_ident ist env (snd id2)) l)
- | TacRevert l -> h_revert (interp_hyp_list ist gl l)
+ Goal.env >>- fun env ->
+ Proofview.V82.tactic begin fun gl ->
+ h_rename (List.map (fun (id1,id2) ->
+ interp_hyp ist gl id1,
+ interp_fresh_ident ist env (snd id2)) l)
+ gl
+ end
+ | TacRevert l ->
+ Proofview.V82.tactic begin fun gl ->
+ h_revert (interp_hyp_list ist gl l) gl
+ end
(* Constructors *)
| TacLeft (ev,bl) ->
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
let sigma, bl = interp_bindings ist env sigma bl in
- tclWITHHOLES ev (h_left ev) sigma bl
+ Tacticals.New.tclWITHHOLES ev (h_left ev) sigma bl
| TacRight (ev,bl) ->
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
let sigma, bl = interp_bindings ist env sigma bl in
- tclWITHHOLES ev (h_right ev) sigma bl
+ Tacticals.New.tclWITHHOLES ev (h_right ev) sigma bl
| TacSplit (ev,_,bll) ->
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
let sigma, bll = List.fold_map (interp_bindings ist env) sigma bll in
- tclWITHHOLES ev (h_split ev) sigma bll
+ Tacticals.New.tclWITHHOLES ev (h_split ev) sigma bll
| TacAnyConstructor (ev,t) ->
Tactics.any_constructor ev (Option.map (interp_tactic ist) t)
| TacConstructor (ev,n,bl) ->
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
let sigma, bl = interp_bindings ist env sigma bl in
- tclWITHHOLES ev (h_constructor ev (interp_int_or_var ist n)) sigma bl
+ Tacticals.New.tclWITHHOLES ev (h_constructor ev (interp_int_or_var ist n)) sigma bl
(* Conversion *)
| TacReduce (r,cl) ->
- let (sigma,r_interp) = pf_interp_red_expr ist gl r in
- tclTHEN
- (tclEVARS sigma)
- (h_reduce r_interp (interp_clause ist gl cl))
+ Proofview.V82.tactic begin fun gl ->
+ let (sigma,r_interp) = pf_interp_red_expr ist gl r in
+ tclTHEN
+ (tclEVARS sigma)
+ (h_reduce r_interp (interp_clause ist gl cl))
+ gl
+ end
| TacChange (None,c,cl) ->
- let is_onhyps = match cl.onhyps with
- | None | Some [] -> true
- | _ -> false
- in
- let is_onconcl = match cl.concl_occs with
- | AllOccurrences | NoOccurrences -> true
- | _ -> false
- in
- let (sigma,c_interp) =
- if is_onhyps && is_onconcl
- then pf_interp_type ist gl c
- else pf_interp_constr ist gl c
- in
- tclTHEN
- (tclEVARS sigma)
- (h_change None c_interp (interp_clause ist gl cl))
+ Proofview.V82.tactic begin fun gl ->
+ let is_onhyps = match cl.onhyps with
+ | None | Some [] -> true
+ | _ -> false
+ in
+ let is_onconcl = match cl.concl_occs with
+ | AllOccurrences | NoOccurrences -> true
+ | _ -> false
+ in
+ let (sigma,c_interp) =
+ if is_onhyps && is_onconcl
+ then pf_interp_type ist gl c
+ else pf_interp_constr ist gl c
+ in
+ tclTHEN
+ (tclEVARS sigma)
+ (h_change None c_interp (interp_clause ist gl cl))
+ gl
+ end
| TacChange (Some op,c,cl) ->
- let sign,op = interp_typed_pattern ist env sigma op in
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
+ Proofview.V82.tactic begin fun gl ->
+ let sign,op = interp_typed_pattern ist env sigma op in
(* spiwack: (2012/04/18) the evar_map output by pf_interp_constr
is dropped as the evar_map taken as input (from
extend_gl_hyps) is incorrect. This means that evar
instantiated by pf_interp_constr may be lost, there. *)
let to_catch = function Not_found -> true | e -> Errors.is_anomaly e in
let (_,c_interp) =
- try pf_interp_constr ist (extend_gl_hyps gl sign) c
- with e when to_catch e (* Hack *) ->
- errorlabstrm "" (strbrk "Failed to get enough information from the left-hand side to type the right-hand side.")
- in
- tclTHEN
- (tclEVARS sigma)
- (h_change (Some op) c_interp (interp_clause ist { gl with sigma=sigma } cl))
+ try pf_interp_constr ist (extend_gl_hyps gl sign) c
+ with e when to_catch e (* Hack *) ->
+ errorlabstrm "" (strbrk "Failed to get enough information from the left-hand side to type the right-hand side.")
+ in
+ tclTHEN
+ (tclEVARS sigma)
+ (h_change (Some op) c_interp (interp_clause ist { gl with sigma=sigma } cl))
+ gl
+ end
(* Equivalence relations *)
| TacReflexivity -> h_reflexivity
- | TacSymmetry c -> h_symmetry (interp_clause ist gl c)
+ | TacSymmetry c ->
+ Tacmach.New.of_old (fun gl -> interp_clause ist gl c) >>- fun cl ->
+ h_symmetry cl
| TacTransitivity c ->
- begin match c with
- | None -> h_transitivity None
- | Some c ->
- let (sigma,c_interp) = pf_interp_constr ist gl c in
- tclTHEN
- (tclEVARS sigma)
- (h_transitivity (Some c_interp))
- end
+ begin match c with
+ | None -> h_transitivity None
+ | Some c ->
+ Tacmach.New.of_old (fun gl -> pf_interp_constr ist gl c) >>- fun (sigma,c_interp) ->
+ Tacticals.New.tclTHEN
+ (Proofview.V82.tactic (tclEVARS sigma))
+ (h_transitivity (Some c_interp))
+ end
(* Equality and inversion *)
| TacRewrite (ev,l,cl,by) ->
let l = List.map (fun (b,m,c) ->
let f env sigma = interp_open_constr_with_bindings ist env sigma c in
(b,m,f)) l in
- let cl = interp_clause ist gl cl in
+ Tacmach.New.of_old (fun gl -> interp_clause ist gl cl) >>- fun cl ->
Equality.general_multi_multi_rewrite ev l cl
- (Option.map (fun by -> tclCOMPLETE (interp_tactic ist by), Equality.Naive) by)
+ (Option.map (fun by -> Tacticals.New.tclCOMPLETE (interp_tactic ist by), Equality.Naive) by)
| TacInversion (DepInversion (k,c,ids),hyp) ->
- let (sigma,c_interp) =
- match c with
- | None -> sigma , None
- | Some c ->
- let (sigma,c_interp) = pf_interp_constr ist gl c in
- sigma , Some c_interp
- in
+ Goal.defs >>- fun sigma ->
+ begin match c with
+ | None -> Goal.return (sigma , None)
+ | Some c ->
+ Tacmach.New.of_old (fun gl -> pf_interp_constr ist gl c) >- fun (sigma,c_interp) ->
+ Goal.return (sigma , Some c_interp)
+ end >>- fun (sigma,c_interp) ->
+ Tacmach.New.of_old (interp_intro_pattern ist) >>- fun interp_intro_pattern ->
+ Tacmach.New.of_old (fun gl -> interp_declared_or_quantified_hypothesis ist gl hyp) >>- fun dqhyps ->
Inv.dinv k c_interp
- (Option.map (interp_intro_pattern ist gl) ids)
- (interp_declared_or_quantified_hypothesis ist gl hyp)
+ (Option.map interp_intro_pattern ids)
+ dqhyps
| TacInversion (NonDepInversion (k,idl,ids),hyp) ->
+ Tacmach.New.of_old (interp_intro_pattern ist) >>- fun interp_intro_pattern ->
+ Tacmach.New.of_old (fun gl -> interp_hyp_list ist gl idl) >>- fun hyps ->
+ Tacmach.New.of_old (fun gl -> interp_declared_or_quantified_hypothesis ist gl hyp) >>- fun dqhyps ->
Inv.inv_clause k
- (Option.map (interp_intro_pattern ist gl) ids)
- (interp_hyp_list ist gl idl)
- (interp_declared_or_quantified_hypothesis ist gl hyp)
+ (Option.map interp_intro_pattern ids)
+ hyps
+ dqhyps
| TacInversion (InversionUsing (c,idl),hyp) ->
- let (sigma,c_interp) = pf_interp_constr ist gl c in
- Leminv.lemInv_clause (interp_declared_or_quantified_hypothesis ist gl hyp)
+ Tacmach.New.of_old (fun gl -> pf_interp_constr ist gl c) >>- fun (sigma,c_interp) ->
+ Tacmach.New.of_old (fun gl -> interp_declared_or_quantified_hypothesis ist gl hyp) >>- fun dqhyps ->
+ Tacmach.New.of_old (fun gl -> interp_hyp_list ist gl idl) >>- fun hyps ->
+ Leminv.lemInv_clause dqhyps
c_interp
- (interp_hyp_list ist gl idl)
+ hyps
(* For extensions *)
| TacExtend (loc,opn,l) ->
+ Goal.defs >>- fun goal_sigma ->
let tac = lookup_tactic opn in
- let (sigma,args) =
+ Tacmach.New.of_old begin fun gl ->
List.fold_right begin fun a (sigma,acc) ->
let (sigma,a_interp) = interp_genarg ist { gl with sigma=sigma } a in
sigma , a_interp::acc
- end l (project gl,[])
- in
+ end l (goal_sigma,[])
+ end >>- fun (sigma,args) ->
+ Proofview.V82.tactic (tclEVARS sigma) <*>
tac args ist
- | TacAlias (loc,s,l,(_,body)) -> fun gl ->
- let evdref = ref gl.sigma in
- let rec f x = match genarg_tag x with
- | IntOrVarArgType ->
- mk_int_or_var_value ist (out_gen (glbwit wit_int_or_var) x)
- | IdentArgType b ->
- value_of_ident (interp_fresh_ident ist env
- (out_gen (glbwit (wit_ident_gen b)) x))
- | VarArgType ->
- mk_hyp_value ist gl (out_gen (glbwit wit_var) x)
- | RefArgType ->
- Value.of_constr (constr_of_global
- (pf_interp_reference ist gl (out_gen (glbwit wit_ref) x)))
- | GenArgType -> f (out_gen (glbwit wit_genarg) x)
- | ConstrArgType ->
- let (sigma,v) = mk_constr_value ist gl (out_gen (glbwit wit_constr) x) in
- evdref := sigma;
- v
- | OpenConstrArgType false ->
- let (sigma,v) = mk_open_constr_value ist gl (snd (out_gen (glbwit wit_open_constr) x)) in
- evdref := sigma;
- v
- | ConstrMayEvalArgType ->
- let (sigma,c_interp) = interp_constr_may_eval ist gl (out_gen (glbwit wit_constr_may_eval) x) in
- evdref := sigma;
- Value.of_constr c_interp
- | ListArgType ConstrArgType ->
- let wit = glbwit (wit_list wit_constr) in
- let (sigma,l_interp) =
- List.fold_right begin fun c (sigma,acc) ->
- let (sigma,c_interp) = mk_constr_value ist { gl with sigma=sigma } c in
- sigma , c_interp::acc
- end (out_gen wit x) (project gl,[])
- in
- evdref := sigma;
- in_gen (topwit (wit_list wit_genarg)) l_interp
- | ListArgType VarArgType ->
- let wit = glbwit (wit_list wit_var) in
- let ans = List.map (mk_hyp_value ist gl) (out_gen wit x) in
- in_gen (topwit (wit_list wit_genarg)) ans
- | ListArgType IntOrVarArgType ->
- let wit = glbwit (wit_list wit_int_or_var) in
- let ans = List.map (mk_int_or_var_value ist) (out_gen wit x) in
- in_gen (topwit (wit_list wit_genarg)) ans
- | ListArgType (IdentArgType b) ->
- let wit = glbwit (wit_list (wit_ident_gen b)) in
- let mk_ident x = value_of_ident (interp_fresh_ident ist env x) in
- let ans = List.map mk_ident (out_gen wit x) in
- in_gen (topwit (wit_list wit_genarg)) ans
- | ListArgType _ -> app_list f x
- | ExtraArgType _ ->
- (** Special treatment of tactics *)
- let gl = { gl with sigma = !evdref } in
- if has_type x (glbwit wit_tactic) then
- let tac = out_gen (glbwit wit_tactic) x in
- let (sigma, v) = val_interp ist gl tac in
- let () = evdref := sigma in
- v
- else
- let (sigma, v) = Geninterp.generic_interp ist gl x in
- let () = evdref := sigma in
- v
- | QuantHypArgType | RedExprArgType
- | OpenConstrArgType _ | ConstrWithBindingsArgType
- | BindingsArgType
- | OptArgType _ | PairArgType _
- -> error "This argument type is not supported in tactic notations."
-
- in
- let addvar (x, v) accu = Id.Map.add x (f v) accu in
- let lfun = List.fold_right addvar l ist.lfun in
- let trace = push_trace (loc,LtacNotationCall s) ist in
- let gl = { gl with sigma = !evdref } in
- let ist = {
- lfun = lfun;
- extra = TacStore.set ist.extra f_trace trace; } in
- interp_tactic ist body gl
+ | TacAlias (loc,s,l,(_,body)) ->
+ Goal.env >>- fun env ->
+ let rec f x = match genarg_tag x with
+ | IntOrVarArgType ->
+ Proofview.tclUNIT (Goal.return (mk_int_or_var_value ist (out_gen (glbwit wit_int_or_var) x)))
+ | IdentArgType b ->
+ Proofview.tclUNIT begin
+ Goal.return (value_of_ident (interp_fresh_ident ist env
+ (out_gen (glbwit (wit_ident_gen b)) x)))
+ end
+ | VarArgType ->
+ Proofview.tclUNIT
+ (Tacmach.New.of_old (fun gl -> mk_hyp_value ist gl (out_gen (glbwit wit_var) x)))
+ | RefArgType ->
+ Proofview.tclUNIT begin
+ Tacmach.New.of_old (fun gl ->
+ Value.of_constr (constr_of_global
+ (pf_interp_reference ist gl (out_gen (glbwit wit_ref) x))))
+ end
+
+ | GenArgType -> f (out_gen (glbwit wit_genarg) x)
+ | ConstrArgType ->
+ Tacmach.New.of_old (fun gl -> mk_constr_value ist gl (out_gen (glbwit wit_constr) x)) >>-- fun (sigma,v) ->
+ Proofview.V82.tactic (tclEVARS sigma) <*>
+ Proofview.tclUNIT (Goal.return v)
+ | OpenConstrArgType false ->
+ Tacmach.New.of_old (fun gl -> mk_open_constr_value ist gl (snd (out_gen (glbwit wit_open_constr) x))) >>-- fun (sigma,v) ->
+ Proofview.V82.tactic (tclEVARS sigma) <*>
+ Proofview.tclUNIT (Goal.return v)
+ | ConstrMayEvalArgType ->
+ Tacmach.New.of_old (fun gl -> interp_constr_may_eval ist gl (out_gen (glbwit wit_constr_may_eval) x)) >>-- fun (sigma,c_interp) ->
+ Proofview.V82.tactic (tclEVARS sigma) <*>
+ Proofview.tclUNIT (Goal.return (Value.of_constr c_interp))
+ | ListArgType ConstrArgType ->
+ let wit = glbwit (wit_list wit_constr) in
+ Tacmach.New.of_old begin fun gl ->
+ List.fold_right begin fun c (sigma,acc) ->
+ let (sigma,c_interp) = mk_constr_value ist { gl with sigma=sigma } c in
+ sigma , c_interp::acc
+ end (out_gen wit x) (project gl,[])
+ end >>-- fun (sigma,l_interp) ->
+ Proofview.V82.tactic (tclEVARS sigma) <*>
+ Proofview.tclUNIT (Goal.return (in_gen (topwit (wit_list wit_genarg)) l_interp))
+ | ListArgType VarArgType ->
+ let wit = glbwit (wit_list wit_var) in
+ Proofview.tclUNIT (Tacmach.New.of_old (fun gl ->
+ let ans = List.map (mk_hyp_value ist gl) (out_gen wit x) in
+ in_gen (topwit (wit_list wit_genarg)) ans
+ ))
+ | ListArgType IntOrVarArgType ->
+ let wit = glbwit (wit_list wit_int_or_var) in
+ let ans = List.map (mk_int_or_var_value ist) (out_gen wit x) in
+ Proofview.tclUNIT (Goal.return (in_gen (topwit (wit_list wit_genarg)) ans))
+ | ListArgType (IdentArgType b) ->
+ let wit = glbwit (wit_list (wit_ident_gen b)) in
+ let mk_ident x = value_of_ident (interp_fresh_ident ist env x) in
+ let ans = List.map mk_ident (out_gen wit x) in
+ Proofview.tclUNIT (Goal.return (in_gen (topwit (wit_list wit_genarg)) ans))
+ | ListArgType t ->
+ (* arnaud: unsafe, faire avec des combinateurs. Dans la version originale c'était juste [Genarg.app_list f x] *)
+ fold_list begin fun a l ->
+ f a >>== fun a' ->
+ l >>== fun l ->
+ Proofview.tclUNIT (Goal.return (a'::l))
+ end x (Proofview.tclUNIT (Goal.return [])) >>== fun l ->
+ Proofview.tclUNIT (Goal.return (in_gen
+ (topwit (wit_list (Obj.magic t)))
+ l))
+ | ExtraArgType _ ->
+ (** Special treatment of tactics *)
+ if has_type x (glbwit wit_tactic) then
+ let tac = out_gen (glbwit wit_tactic) x in
+ val_interp ist tac >>== fun v ->
+ Proofview.tclUNIT (Goal.return v)
+ else
+ Tacmach.New.of_old (fun gl ->
+ Geninterp.generic_interp ist gl x) >>-- fun (newsigma,v) ->
+ Proofview.V82.tactic (tclEVARS newsigma) <*>
+ Proofview.tclUNIT (Goal.return v)
+ | QuantHypArgType | RedExprArgType
+ | OpenConstrArgType _ | ConstrWithBindingsArgType
+ | BindingsArgType
+ | OptArgType _ | PairArgType _
+ -> Proofview.tclZERO (UserError("" , str"This argument type is not supported in tactic notations."))
+ in
+ let addvar (x, v) accu =
+ accu >>== fun accu ->
+ f v >>== fun v ->
+ Proofview.tclUNIT (Goal.return (Id.Map.add x v accu))
+ in
+ List.fold_right addvar l (Proofview.tclUNIT (Goal.return ist.lfun)) >>= fun lfun ->
+ let trace = push_trace (loc,LtacNotationCall s) ist in
+ let ist = {
+ lfun = lfun;
+ extra = TacStore.set ist.extra f_trace trace; } in
+ interp_tactic ist body
(* Initial call for interpretation *)
@@ -1951,13 +2173,17 @@ let default_ist () =
let extra = TacStore.set TacStore.empty f_debug (get_debug ()) in
{ lfun = Id.Map.empty; extra = extra }
-let eval_tactic t gls =
+let eval_tactic t =
+ Proofview.tclUNIT () >= fun () -> (* delay for [db_initialize] and [default_ist] *)
db_initialize ();
- interp_tactic (default_ist ()) t gls
+ interp_tactic (default_ist ()) t
(* globalization + interpretation *)
-let interp_tac_gen lfun avoid_ids debug t gl =
+
+let interp_tac_gen lfun avoid_ids debug t =
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
let extra = TacStore.set TacStore.empty f_debug debug in
let extra = TacStore.set extra f_avoid_ids avoid_ids in
let ist = { lfun = lfun; extra = extra } in
@@ -1965,24 +2191,27 @@ let interp_tac_gen lfun avoid_ids debug t gl =
interp_tactic ist
(intern_pure_tactic {
ltacvars; ltacrecvars = Id.Map.empty;
- gsigma = project gl; genv = pf_env gl } t) gl
+ gsigma = sigma; genv = env } t)
let interp t = interp_tac_gen Id.Map.empty [] (get_debug()) t
let _ = Proof_global.set_interp_tac interp
-let eval_ltac_constr gl t =
- interp_ltac_constr (default_ist ()) gl
+let eval_ltac_constr t =
+ Proofview.tclUNIT () >= fun () -> (* delay for [make_empty_glob_sign] and [default_ist] *)
+ interp_ltac_constr (default_ist ())
(intern_tactic_or_tacarg (make_empty_glob_sign ()) t )
(* Used to hide interpretation for pretty-print, now just launch tactics *)
-let hide_interp t ot gl =
+let hide_interp t ot =
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
let ist = { ltacvars = Id.Set.empty; ltacrecvars = Id.Map.empty;
- gsigma = project gl; genv = pf_env gl } in
+ gsigma = sigma; genv = env } in
let te = intern_pure_tactic ist t in
let t = eval_tactic te in
match ot with
- | None -> t gl
- | Some t' -> (tclTHEN t t') gl
+ | None -> t
+ | Some t' -> Tacticals.New.tclTHEN t t'
(***************************************************************************)
(** Register standard arguments *)
diff --git a/tactics/tacinterp.mli b/tactics/tacinterp.mli
index 063104128..dffea3e65 100644
--- a/tactics/tacinterp.mli
+++ b/tactics/tacinterp.mli
@@ -51,11 +51,11 @@ val extract_ltac_constr_values : interp_sign -> Environ.env ->
(** Tactic extensions *)
val add_tactic :
- string -> (typed_generic_argument list -> interp_sign -> tactic) -> unit
+ string -> (typed_generic_argument list -> interp_sign -> unit Proofview.tactic) -> unit
val overwriting_add_tactic :
- string -> (typed_generic_argument list -> interp_sign -> tactic) -> unit
+ string -> (typed_generic_argument list -> interp_sign -> unit Proofview.tactic) -> unit
val lookup_tactic :
- string -> (typed_generic_argument list) -> interp_sign -> tactic
+ string -> (typed_generic_argument list) -> interp_sign -> unit Proofview.tactic
(** To embed several objects in Coqast.t *)
val tactic_in : (interp_sign -> glob_tactic_expr) -> Dyn.t
@@ -71,18 +71,16 @@ val set_debug : debug_info -> unit
(** Gives the state of debug *)
val get_debug : unit -> debug_info
-
(** Adds an interpretation function for extra generic arguments *)
val interp_genarg : interp_sign -> goal sigma -> glob_generic_argument ->
Evd.evar_map * typed_generic_argument
(** Interprets any expression *)
-val val_interp : interp_sign -> goal sigma -> glob_tactic_expr -> Evd.evar_map * value
+val val_interp : interp_sign -> glob_tactic_expr -> value Goal.sensitive Proofview.tactic
(** Interprets an expression that evaluates to a constr *)
-val interp_ltac_constr : interp_sign -> goal sigma -> glob_tactic_expr ->
- Evd.evar_map * constr
+val interp_ltac_constr : interp_sign -> glob_tactic_expr -> constr Goal.sensitive Proofview.tactic
(** Interprets redexp arguments *)
val interp_redexp : Environ.env -> Evd.evar_map -> raw_red_expr -> Evd.evar_map * red_expr
@@ -99,20 +97,20 @@ val interp_open_constr_with_bindings : interp_sign -> Environ.env -> Evd.evar_ma
(** Initial call for interpretation *)
-val eval_tactic : glob_tactic_expr -> tactic
+val eval_tactic : glob_tactic_expr -> unit Proofview.tactic
(** Globalization + interpretation *)
val interp_tac_gen : value Id.Map.t -> Id.t list ->
- debug_info -> raw_tactic_expr -> tactic
+ debug_info -> raw_tactic_expr -> unit Proofview.tactic
-val interp : raw_tactic_expr -> tactic
+val interp : raw_tactic_expr -> unit Proofview.tactic
-val eval_ltac_constr : goal sigma -> raw_tactic_expr -> Evd.evar_map * constr
+val eval_ltac_constr : raw_tactic_expr -> constr Goal.sensitive Proofview.tactic
(** Hides interpretation for pretty-print *)
-val hide_interp : raw_tactic_expr -> tactic option -> tactic
+val hide_interp : raw_tactic_expr -> unit Proofview.tactic option -> unit Proofview.tactic
(** Declare the xml printer *)
val declare_xml_printer :
diff --git a/tactics/tactic_option.mli b/tactics/tactic_option.mli
index 914408369..be36b3f82 100644
--- a/tactics/tactic_option.mli
+++ b/tactics/tactic_option.mli
@@ -12,5 +12,5 @@ open Vernacexpr
val declare_tactic_option : ?default:Tacexpr.glob_tactic_expr -> string ->
(* put *) (locality_flag -> glob_tactic_expr -> unit) *
- (* get *) (unit -> locality_flag * tactic) *
+ (* get *) (unit -> locality_flag * unit Proofview.tactic) *
(* print *) (unit -> Pp.std_ppcmds)
diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml
index 4625cc84b..ee89b55a8 100644
--- a/tactics/tacticals.ml
+++ b/tactics/tacticals.ml
@@ -50,7 +50,6 @@ let tclAT_LEAST_ONCE = Refiner.tclAT_LEAST_ONCE
let tclFAIL = Refiner.tclFAIL
let tclFAIL_lazy = Refiner.tclFAIL_lazy
let tclDO = Refiner.tclDO
-let tclTIMEOUT = Refiner.tclTIMEOUT
let tclWEAK_PROGRESS = Refiner.tclWEAK_PROGRESS
let tclPROGRESS = Refiner.tclPROGRESS
let tclSHOWHYPS = Refiner.tclSHOWHYPS
@@ -367,3 +366,258 @@ let make_case_branch_assumptions ba gl =
let case_on_ba tac ba gl = tac (make_case_branch_assumptions ba gl) gl
+
+(** Tacticals of Ltac defined directly in term of Proofview *)
+module New = struct
+ open Proofview
+ open Proofview.Notations
+
+ let tclTHEN t1 t2 =
+ t1 <*> tclEXTEND [] t2 []
+
+ let tclFAIL lvl msg =
+ tclZERO (Refiner.FailError (lvl,lazy msg))
+
+ let catch_failerror e =
+ try
+ Refiner.catch_failerror e;
+ tclUNIT ()
+ with e -> tclZERO e
+ let tclORELSE0 t1 t2 =
+ tclORELSE
+ t1
+ begin fun e ->
+ catch_failerror e <*> t2
+ end
+ let tclORELSE t1 t2 =
+ tclORELSE0 (tclPROGRESS t1) t2
+
+ let tclTHENS3PARTS t1 l1 repeat l2 =
+ t1 <*> tclEXTEND (Array.to_list l1) repeat (Array.to_list l2)
+ let tclTHENSFIRSTn t1 l repeat =
+ tclTHENS3PARTS t1 l repeat [||]
+ let tclTHENFIRSTn t1 l =
+ tclTHENSFIRSTn t1 l (tclUNIT())
+ let tclTHENFIRST t1 t2 =
+ t1 <*> tclFOCUS 1 1 t2
+ let tclTHENLASTn t1 l =
+ t1 <*> tclEXTEND [] (tclUNIT()) (Array.to_list l)
+ let tclTHENLAST t1 t2 = tclTHENLASTn t1 [|t2|]
+ let tclTHENS t l =
+ t <*> tclDISPATCH l
+ let tclTHENLIST l =
+ List.fold_left tclTHEN (tclUNIT()) l
+
+
+ (* [tclMAP f [x1..xn]] builds [(f x1);(f x2);...(f xn)] *)
+ let tclMAP tacfun l =
+ List.fold_right (fun x -> (tclTHEN (tacfun x))) l (tclUNIT())
+
+ let tclTRY t =
+ tclORELSE0 t (tclUNIT ())
+
+ let tclIFTHENELSE t1 t2 t3 =
+ tclIFCATCH t1
+ (fun () -> tclEXTEND [] t2 [])
+ (fun _ -> t3)
+ let tclIFTHENSVELSE t1 a t3 =
+ tclIFCATCH t1
+ (fun () -> tclDISPATCH (Array.to_list a))
+ (fun _ -> t3)
+ let tclIFTHENTRYELSEMUST t1 t2 =
+ tclIFTHENELSE t1 (tclTRY t2) t2
+
+ (* Try the first tactic that does not fail in a list of tactics *)
+ let rec tclFIRST = function
+ | [] -> tclZERO (Errors.UserError ("Tacticals.New.tclFIRST",str"No applicable tactic."))
+ | t::rest -> tclORELSE0 t (tclFIRST rest)
+
+ let rec tclFIRST_PROGRESS_ON tac = function
+ | [] -> tclFAIL 0 (str "No applicable tactic")
+ | [a] -> tac a (* so that returned failure is the one from last item *)
+ | a::tl -> tclORELSE (tac a) (tclFIRST_PROGRESS_ON tac tl)
+
+ let rec tclDO n t =
+ if n < 0 then
+ tclZERO (Errors.UserError (
+ "Refiner.tclDO",
+ str"Wrong argument : Do needs a positive integer.")
+ )
+ else if n = 0 then tclUNIT ()
+ else if n = 1 then t
+ else tclTHEN t (tclDO (n-1) t)
+
+ let rec tclREPEAT0 t =
+ tclIFCATCH t
+ (fun () -> tclEXTEND [] (tclREPEAT0 t) [])
+ (fun _ -> tclUNIT ())
+ let tclREPEAT t =
+ tclREPEAT0 (tclPROGRESS t)
+ let rec tclREPEAT_MAIN0 t =
+ tclIFCATCH t
+ (fun () -> tclFOCUS 1 1 (tclREPEAT_MAIN0 t))
+ (fun _ -> tclUNIT ())
+ let tclREPEAT_MAIN t =
+ tclREPEAT_MAIN0 (tclPROGRESS t)
+
+ let tclCOMPLETE t =
+ t >= fun res ->
+ (tclEXTEND
+ []
+ (tclZERO (Errors.UserError ("",str"Proof is not complete.")))
+ []
+ ) <*>
+ tclUNIT res
+
+ (* Try the first thats solves the current goal *)
+ let tclSOLVE tacl = tclFIRST (List.map tclCOMPLETE tacl)
+
+
+ let tclWITHHOLES accept_unresolved_holes tac sigma x =
+ tclEVARMAP >= fun sigma_initial ->
+ if sigma == sigma_initial then tac x
+ else
+ let check_evars env new_sigma sigma initial_sigma =
+ try
+ Refiner.check_evars env new_sigma sigma initial_sigma;
+ tclUNIT ()
+ with e ->
+ tclZERO e
+ in
+ let check_evars_if =
+ if not accept_unresolved_holes then
+ tclEVARMAP >= fun sigma_final ->
+ tclENV >= fun env ->
+ check_evars env sigma_final sigma sigma_initial
+ else
+ tclUNIT ()
+ in
+ Proofview.V82.tactic (Refiner.tclEVARS sigma) <*> tac x <*> check_evars_if
+
+ let nthDecl m =
+ Goal.hyps >- fun hyps ->
+ let hyps = Environ.named_context_of_val hyps in
+ try Goal.return (List.nth hyps (m-1))
+ with Failure _ -> error "No such assumption."
+
+ let nthHypId m = nthDecl m >- fun (id,_,_) -> Goal.return id
+ let nthHyp m = nthHypId m >- fun id -> Goal.return (mkVar id)
+
+ let onNthHypId m tac =
+ nthHypId m >>- tac
+ let onNthHyp m tac =
+ nthHyp m >>- tac
+
+ let onLastHypId = onNthHypId 1
+ let onLastHyp = onNthHyp 1
+
+ let onNthDecl m tac = nthDecl m >>- tac
+ let onLastDecl = onNthDecl 1
+
+ let ifOnHyp pred tac1 tac2 id =
+ Tacmach.New.pf_get_hyp_typ id >>- fun typ ->
+ if pred (id,typ) then
+ tac1 id
+ else
+ tac2 id
+
+ let fullGoal =
+ Tacmach.New.pf_ids_of_hyps >- fun hyps ->
+ Goal.return (None :: List.map Option.make hyps)
+
+ let tryAllHyps tac =
+ Tacmach.New.pf_ids_of_hyps >>- fun hyps ->
+ tclFIRST_PROGRESS_ON tac hyps
+ let tryAllHypsAndConcl tac =
+ fullGoal >>- fun fullGoal ->
+ tclFIRST_PROGRESS_ON tac fullGoal
+
+ let onClause tac cl =
+ Tacmach.New.pf_ids_of_hyps >>- fun hyps ->
+ tclMAP tac (Locusops.simple_clause_of (fun () -> hyps) cl)
+
+ (* Find the right elimination suffix corresponding to the sort of the goal *)
+ (* c should be of type A1->.. An->B with B an inductive definition *)
+ let general_elim_then_using mk_elim
+ isrec allnames tac predicate (indbindings,elimbindings)
+ ind indclause =
+ Tacmach.New.of_old (mk_elim ind) >>- fun elim ->
+ (* applying elimination_scheme just a little modified *)
+ let indclause' = clenv_match_args indbindings indclause in
+ Tacmach.New.pf_apply Typing.type_of >>- fun type_of ->
+ Tacmach.New.of_old (fun gl -> mk_clenv_from gl (elim,type_of elim)) >>- fun elimclause ->
+ let indmv =
+ match kind_of_term (last_arg elimclause.templval.Evd.rebus) with
+ | Meta mv -> mv
+ | _ -> anomaly (str"elimination")
+ in
+ let pmv =
+ let p, _ = decompose_app elimclause.templtyp.Evd.rebus in
+ match kind_of_term p with
+ | Meta p -> p
+ | _ ->
+ let name_elim =
+ match kind_of_term elim with
+ | Const kn -> string_of_con kn
+ | Var id -> string_of_id id
+ | _ -> "\b"
+ in
+ error ("The elimination combinator " ^ name_elim ^ " is unknown.")
+ in
+ let elimclause' = clenv_fchain indmv elimclause indclause' in
+ let elimclause' = clenv_match_args elimbindings elimclause' in
+ let branchsigns = compute_construtor_signatures isrec ind in
+ let brnames = compute_induction_names (Array.length branchsigns) allnames in
+ let after_tac ce i =
+ let (hd,largs) = decompose_app ce.templtyp.Evd.rebus in
+ let ba = { branchsign = branchsigns.(i);
+ branchnames = brnames.(i);
+ nassums =
+ List.fold_left
+ (fun acc b -> if b then acc+2 else acc+1)
+ 0 branchsigns.(i);
+ branchnum = i+1;
+ ity = ind;
+ largs = List.map (clenv_nf_meta ce) largs;
+ pred = clenv_nf_meta ce hd }
+ in
+ tac ba
+ in
+ let branchtacs ce = Array.init (Array.length branchsigns) (after_tac ce) in
+ let elimclause' =
+ match predicate with
+ | None -> elimclause'
+ | Some p ->
+ clenv_unify ~flags:Unification.elim_flags
+ Reduction.CONV (mkMeta pmv) p elimclause'
+ in
+ new_elim_res_pf_THEN_i elimclause' branchtacs
+
+ let elimination_then_using tac predicate bindings c =
+ Tacmach.New.of_old (fun gl -> pf_reduce_to_quantified_ind gl (pf_type_of gl c)) >>- fun (ind,t) ->
+ Tacmach.New.of_old (fun gl -> mk_clenv_from gl (c,t)) >>- fun indclause ->
+ let isrec,mkelim =
+ if (Global.lookup_mind (fst ind)).mind_record
+ then false,gl_make_case_dep
+ else true,gl_make_elim
+ in
+ general_elim_then_using mkelim isrec
+ None tac predicate bindings ind indclause
+
+ let case_then_using =
+ general_elim_then_using gl_make_case_dep false
+
+ let case_nodep_then_using =
+ general_elim_then_using gl_make_case_nodep false
+
+
+ let elimination_then tac = elimination_then_using tac None
+
+ let elim_on_ba tac ba =
+ Tacmach.New.of_old (make_elim_branch_assumptions ba) >>- fun branches ->
+ tac branches
+
+ let case_on_ba tac ba =
+ Tacmach.New.of_old (make_case_branch_assumptions ba) >>- fun branches ->
+ tac branches
+end
diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli
index 13eaaff5c..58aff5fdb 100644
--- a/tactics/tacticals.mli
+++ b/tactics/tacticals.mli
@@ -52,7 +52,6 @@ val tclAT_LEAST_ONCE : tactic -> tactic
val tclFAIL : int -> std_ppcmds -> tactic
val tclFAIL_lazy : int -> std_ppcmds Lazy.t -> tactic
val tclDO : int -> tactic -> tactic
-val tclTIMEOUT : int -> tactic -> tactic
val tclWEAK_PROGRESS : tactic -> tactic
val tclPROGRESS : tactic -> tactic
val tclSHOWHYPS : tactic -> tactic
@@ -173,3 +172,85 @@ val simple_elimination_then :
val elim_on_ba : (branch_assumptions -> tactic) -> branch_args -> tactic
val case_on_ba : (branch_assumptions -> tactic) -> branch_args -> tactic
+
+(** Tacticals defined directly in term of Proofview *)
+module New : sig
+ open Proofview
+
+
+ val tclTHEN : unit tactic -> unit tactic -> unit tactic
+ (* [tclFAIL n msg] fails with [msg] as an error message at level [n]
+ (meaning that it will jump over [n] error catching tacticals FROM
+ THIS MODULE. *)
+ val tclFAIL : int -> Pp.std_ppcmds -> 'a tactic
+ val tclORELSE0 : 'a tactic -> 'a tactic -> 'a tactic
+ val tclORELSE : 'a tactic -> 'a tactic -> 'a 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] subgoals and [tac2] to the rest of the subgoals in the
+ middle. Raises an error if the number of resulting subgoals is
+ strictly less than [n+m] *)
+ val tclTHENS3PARTS : unit tactic -> unit tactic array -> unit tactic -> unit tactic array -> unit tactic
+ val tclTHENSFIRSTn : unit tactic -> unit tactic array -> unit tactic -> unit tactic
+ val tclTHENFIRSTn : unit tactic -> unit tactic array -> unit tactic
+ (** [tclTHENFIRST tac1 tac2 gls] applies the tactic [tac1] to [gls]
+ and [tac2] to the first resulting subgoal *)
+ val tclTHENFIRST : unit tactic -> unit tactic -> unit tactic
+ val tclTHENLASTn : unit tactic -> unit tactic array -> unit tactic
+ val tclTHENLAST : unit tactic -> unit tactic -> unit tactic
+ (* [tclTHENS t l = t <*> tclDISPATCH l] *)
+ val tclTHENS : unit tactic -> unit tactic list -> unit tactic
+ (* [tclTHENLIST [t1;…;tn]] is [t1<*>…<*>tn] *)
+ val tclTHENLIST : unit tactic list -> unit tactic
+
+ (** [tclMAP f [x1..xn]] builds [(f x1);(f x2);...(f xn)] *)
+ val tclMAP : ('a -> unit tactic) -> 'a list -> unit tactic
+
+ val tclTRY : unit tactic -> unit tactic
+ val tclFIRST : 'a tactic list -> 'a tactic
+ val tclFIRST_PROGRESS_ON : ('a -> unit tactic) -> 'a list -> unit tactic
+ val tclIFTHENELSE : unit tactic -> unit tactic -> unit tactic -> unit tactic
+ val tclIFTHENSVELSE : unit tactic -> unit tactic array -> unit tactic -> unit tactic
+ val tclIFTHENTRYELSEMUST : unit tactic -> unit tactic -> unit tactic
+
+ val tclDO : int -> unit tactic -> unit tactic
+ val tclREPEAT : unit tactic -> unit tactic
+ (* Repeat on the first subgoal (no failure if no more subgoal) *)
+ val tclREPEAT_MAIN : unit tactic -> unit tactic
+ val tclCOMPLETE : 'a tactic -> 'a tactic
+ val tclSOLVE : unit tactic list -> unit tactic
+ val tclWITHHOLES : bool -> ('a -> unit tactic) -> Evd.evar_map -> 'a -> unit tactic
+
+
+ val ifOnHyp : (identifier * types -> bool) ->
+ (identifier -> unit Proofview.tactic) -> (identifier -> unit Proofview.tactic) ->
+ identifier -> unit Proofview.tactic
+
+ val onNthHypId : int -> (identifier -> unit tactic) -> unit tactic
+ val onLastHypId : (identifier -> unit tactic) -> unit tactic
+ val onLastHyp : (constr -> unit tactic) -> unit tactic
+ val onLastDecl : (named_declaration -> unit tactic) -> unit tactic
+
+ val tryAllHyps : (identifier -> unit tactic) -> unit tactic
+ val tryAllHypsAndConcl : (identifier option -> unit tactic) -> unit tactic
+ val onClause : (identifier option -> unit tactic) -> clause -> unit tactic
+
+ val elimination_then :
+ (branch_args -> unit Proofview.tactic) ->
+ (arg_bindings * arg_bindings) -> constr -> unit Proofview.tactic
+
+ val case_then_using :
+ intro_pattern_expr located option -> (branch_args -> unit Proofview.tactic) ->
+ constr option -> (arg_bindings * arg_bindings) ->
+ inductive -> clausenv -> unit Proofview.tactic
+
+ val case_nodep_then_using :
+ intro_pattern_expr located option -> (branch_args -> unit Proofview.tactic) ->
+ constr option -> (arg_bindings * arg_bindings) ->
+ inductive -> clausenv -> unit Proofview.tactic
+
+ val elim_on_ba : (branch_assumptions -> unit Proofview.tactic) -> branch_args -> unit Proofview.tactic
+ val case_on_ba : (branch_assumptions -> unit Proofview.tactic) -> branch_args -> unit Proofview.tactic
+end
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index 004a93ed1..c67e4b8d2 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -44,6 +44,7 @@ open Locus
open Locusops
open Misctypes
open Miscops
+open Proofview.Notations
exception Bound
@@ -420,16 +421,19 @@ type intro_name_flag =
| IntroBasedOn of Id.t * Id.t list
| IntroMustBe of Id.t
-let find_name loc decl gl = function
+let find_name loc decl = function
| IntroAvoid idl ->
(* this case must be compatible with [find_intro_names] below. *)
- let id = fresh_id idl (default_id (pf_env gl) gl.sigma decl) gl in id
- | IntroBasedOn (id,idl) -> fresh_id idl id gl
+ Goal.env >- fun env ->
+ Goal.defs >- fun sigma ->
+ Tacmach.New.of_old (fresh_id idl (default_id env sigma decl))
+ | IntroBasedOn (id,idl) -> Tacmach.New.of_old (fresh_id idl id)
| IntroMustBe id ->
(* When name is given, we allow to hide a global name *)
- let id' = next_ident_away id (pf_ids_of_hyps gl) in
+ Tacmach.New.pf_ids_of_hyps >- fun ids_of_hyps ->
+ let id' = next_ident_away id ids_of_hyps in
if not (Id.equal id' id) then user_err_loc (loc,"",pr_id id ++ str" is already used.");
- id'
+ Goal.return id'
(* Returns the names that would be created by intros, without doing
intros. This function is supposed to be compatible with an
@@ -448,29 +452,36 @@ let find_intro_names ctxt gl =
List.rev res
let build_intro_tac id dest tac = match dest with
- | MoveLast -> tclTHEN (introduction id) (tac id)
- | dest -> tclTHENLIST [introduction id; move_hyp true id dest; tac id]
+ | MoveLast -> Tacticals.New.tclTHEN (Proofview.V82.tactic (introduction id)) (tac id)
+ | dest -> Tacticals.New.tclTHENLIST [Proofview.V82.tactic (introduction id); Proofview.V82.tactic (move_hyp true id dest); tac id]
-let rec intro_then_gen loc name_flag move_flag force_flag dep_flag tac gl =
- match kind_of_term (pf_concl gl) with
+let rec intro_then_gen loc name_flag move_flag force_flag dep_flag tac =
+ Goal.concl >>- fun concl ->
+ match kind_of_term concl with
| Prod (name,t,u) when not dep_flag || (dependent (mkRel 1) u) ->
- build_intro_tac (find_name loc (name,None,t) gl name_flag) move_flag tac gl
+ find_name loc (name,None,t) name_flag >>- fun name ->
+ build_intro_tac name move_flag tac
| LetIn (name,b,t,u) when not dep_flag || (dependent (mkRel 1) u) ->
- build_intro_tac (find_name loc (name,Some b,t) gl name_flag) move_flag tac
- gl
+ find_name loc (name,Some b,t) name_flag >>- fun name ->
+ build_intro_tac name move_flag tac
| _ ->
- if not force_flag then raise (RefinerError IntroNeedsProduct);
- (* Note: red_in_concl includes betaiotazeta and this was like *)
- (* this since at least V6.3 (a pity *)
- (* that intro do betaiotazeta only when reduction is needed; and *)
- (* probably also a pity that intro does zeta *)
- try
- tclTHEN hnf_in_concl
- (intro_then_gen loc name_flag move_flag false dep_flag tac) gl
- with RefinerError IntroNeedsProduct ->
- user_err_loc(loc,"Intro",str "No product even after head-reduction.")
-
-let intro_gen loc n m f d = intro_then_gen loc n m f d (fun _ -> tclIDTAC)
+ begin if not force_flag then Proofview.tclZERO (RefinerError IntroNeedsProduct)
+ (* Note: red_in_concl includes betaiotazeta and this was like *)
+ (* this since at least V6.3 (a pity *)
+ (* that intro do betaiotazeta only when reduction is needed; and *)
+ (* probably also a pity that intro does zeta *)
+ else Proofview.tclUNIT ()
+ end <*>
+ Proofview.tclORELSE
+ (Tacticals.New.tclTHEN (Proofview.V82.tactic hnf_in_concl)
+ (intro_then_gen loc name_flag move_flag false dep_flag tac))
+ begin function
+ | RefinerError IntroNeedsProduct ->
+ Proofview.tclZERO (Loc.add_loc (Errors.UserError("Intro",str "No product even after head-reduction.")) loc)
+ | e -> Proofview.tclZERO e
+ end
+
+let intro_gen loc n m f d = intro_then_gen loc n m f d (fun _ -> Proofview.tclUNIT ())
let intro_mustbe_force id = intro_gen dloc (IntroMustBe id) MoveLast true false
let intro_using id = intro_gen dloc (IntroBasedOn (id,[])) MoveLast false false
let intro_then = intro_then_gen dloc (IntroAvoid []) MoveLast false false
@@ -483,20 +494,26 @@ let intro_then_force = intro_then_gen dloc (IntroAvoid []) MoveLast true false
(**** Multiple introduction tactics ****)
let rec intros_using = function
- | [] -> tclIDTAC
- | str::l -> tclTHEN (intro_using str) (intros_using l)
+ | [] -> Proofview.tclUNIT()
+ | str::l -> Tacticals.New.tclTHEN (intro_using str) (intros_using l)
-let intros = tclREPEAT intro
+let intros = Tacticals.New.tclREPEAT intro
let intro_erasing id = tclTHEN (thin [id]) (introduction id)
let intro_forthcoming_then_gen loc name_flag move_flag dep_flag tac =
- let rec aux ids gl =
- try
+ let rec aux ids =
+ Proofview.tclORELSE
+ begin
intro_then_gen loc name_flag move_flag false dep_flag
- (fun id -> aux (id::ids)) gl
- with RefinerError IntroNeedsProduct ->
- tac ids gl in
+ (fun id -> aux (id::ids))
+ end
+ begin function
+ | RefinerError IntroNeedsProduct ->
+ tac ids
+ | e -> Proofview.tclZERO e
+ end
+ in
aux []
let rec get_next_hyp_position id = function
@@ -528,14 +545,14 @@ let intro_replacing id gl =
tclTHENLIST
[thin_for_replacing [id]; introduction id; move_hyp true id next_hyp] gl
-let intros_replacing ids gl =
+let intros_replacing ids =
let rec introrec = function
- | [] -> tclIDTAC
+ | [] -> Proofview.tclUNIT()
| id::tl ->
- tclTHEN (tclORELSE (intro_replacing id) (intro_using id))
+ Tacticals.New.tclTHEN (Tacticals.New.tclORELSE (Proofview.V82.tactic (intro_replacing id)) (intro_using id))
(introrec tl)
in
- introrec ids gl
+ introrec ids
(* User-level introduction tactics *)
@@ -582,8 +599,9 @@ let depth_of_quantified_hypothesis red h gl =
(if red then strbrk " even after head-reduction" else mt ()) ++
str".")
-let intros_until_gen red h g =
- tclDO (depth_of_quantified_hypothesis red h g) (if red then introf else intro) g
+let intros_until_gen red h =
+ Tacmach.New.of_old (depth_of_quantified_hypothesis red h) >>- fun n ->
+ Tacticals.New.tclDO n (if red then introf else intro)
let intros_until_id id = intros_until_gen true (NamedHyp id)
let intros_until_n_gen red n = intros_until_gen red (AnonHyp n)
@@ -595,16 +613,16 @@ let intros_until_n_wored = intros_until_n_gen false
let tclCHECKVAR id gl = ignore (pf_get_hyp gl id); tclIDTAC gl
let try_intros_until_id_check id =
- tclORELSE (intros_until_id id) (tclCHECKVAR id)
+ Tacticals.New.tclORELSE (intros_until_id id) (Proofview.V82.tactic (tclCHECKVAR id))
let try_intros_until tac = function
- | NamedHyp id -> tclTHEN (try_intros_until_id_check id) (tac id)
- | AnonHyp n -> tclTHEN (intros_until_n n) (onLastHypId tac)
+ | NamedHyp id -> Tacticals.New.tclTHEN (try_intros_until_id_check id) (tac id)
+ | AnonHyp n -> Tacticals.New.tclTHEN (intros_until_n n) (Tacticals.New.onLastHypId tac)
let rec intros_move = function
- | [] -> tclIDTAC
+ | [] -> Proofview.tclUNIT ()
| (hyp,destopt) :: rest ->
- tclTHEN (intro_gen dloc (IntroMustBe hyp) destopt false false)
+ Tacticals.New.tclTHEN (intro_gen dloc (IntroMustBe hyp) destopt false false)
(intros_move rest)
let dependent_in_decl a (_,c,t) =
@@ -619,12 +637,12 @@ let onOpenInductionArg tac = function
| ElimOnConstr cbl ->
tac cbl
| ElimOnAnonHyp n ->
- tclTHEN
+ Tacticals.New.tclTHEN
(intros_until_n n)
- (onLastHyp (fun c -> tac (Evd.empty,(c,NoBindings))))
+ (Tacticals.New.onLastHyp (fun c -> tac (Evd.empty,(c,NoBindings))))
| ElimOnIdent (_,id) ->
(* A quantified hypothesis *)
- tclTHEN
+ Tacticals.New.tclTHEN
(try_intros_until_id_check id)
(tac (Evd.empty,(mkVar id,NoBindings)))
@@ -632,10 +650,10 @@ let onInductionArg tac = function
| ElimOnConstr cbl ->
tac cbl
| ElimOnAnonHyp n ->
- tclTHEN (intros_until_n n) (onLastHyp (fun c -> tac (c,NoBindings)))
+ Tacticals.New.tclTHEN (intros_until_n n) (Tacticals.New.onLastHyp (fun c -> tac (c,NoBindings)))
| ElimOnIdent (_,id) ->
(* A quantified hypothesis *)
- tclTHEN (try_intros_until_id_check id) (tac (mkVar id,NoBindings))
+ Tacticals.New.tclTHEN (try_intros_until_id_check id) (tac (mkVar id,NoBindings))
let map_induction_arg f = function
| ElimOnConstr (sigma,(c,bl)) -> ElimOnConstr (f (sigma,c),bl)
@@ -683,7 +701,7 @@ let cut c gl =
gl
| _ -> error "Not a proposition or a type."
-let cut_intro t = tclTHENFIRST (cut t) intro
+let cut_intro t = Tacticals.New.tclTHENFIRST (Proofview.V82.tactic (cut t)) intro
(* [assert_replacing id T tac] adds the subgoals of the proof of [T]
before the current goal
@@ -827,10 +845,10 @@ let general_case_analysis_in_context with_evars (c,lbindc) gl =
let general_case_analysis with_evars (c,lbindc as cx) =
match kind_of_term c with
| Var id when lbindc == NoBindings ->
- tclTHEN (try_intros_until_id_check id)
- (general_case_analysis_in_context with_evars cx)
+ Tacticals.New.tclTHEN (try_intros_until_id_check id)
+ (Proofview.V82.tactic (general_case_analysis_in_context with_evars cx))
| _ ->
- general_case_analysis_in_context with_evars cx
+ Proofview.V82.tactic (general_case_analysis_in_context with_evars cx)
let simplest_case c = general_case_analysis false (c,NoBindings)
@@ -847,22 +865,28 @@ let find_eliminator c gl =
let c = lookup_eliminator ind (elimination_sort_of_goal gl) in
{elimindex = None; elimbody = (c,NoBindings)}
-let default_elim with_evars (c,_ as cx) gl =
- try general_elim with_evars cx (find_eliminator c gl) gl
- with IsRecord ->
- (* For records, induction principles aren't there by default anymore.
- Instead, we do a case analysis instead. *)
- general_case_analysis with_evars cx gl
+(* arnaud: probable bug avec le try *)
+let default_elim with_evars (c,_ as cx) =
+ Proofview.tclORELSE
+ (Tacmach.New.of_old (find_eliminator c) >>- fun elim ->
+ Proofview.V82.tactic (general_elim with_evars cx elim))
+ begin function
+ | IsRecord ->
+ (* For records, induction principles aren't there by default
+ anymore. Instead, we do a case analysis instead. *)
+ general_case_analysis with_evars cx
+ | e -> Proofview.tclZERO e
+ end
let elim_in_context with_evars c = function
| Some elim ->
- general_elim with_evars c {elimindex = Some (-1); elimbody = elim}
+ Proofview.V82.tactic (general_elim with_evars c {elimindex = Some (-1); elimbody = elim})
| None -> default_elim with_evars c
let elim with_evars (c,lbindc as cx) elim =
match kind_of_term c with
| Var id when lbindc == NoBindings ->
- tclTHEN (try_intros_until_id_check id)
+ Tacticals.New.tclTHEN (try_intros_until_id_check id)
(elim_in_context with_evars cx elim)
| _ ->
elim_in_context with_evars cx elim
@@ -1178,11 +1202,11 @@ let clear_wildcards ids =
* true in the boolean list. *)
let rec intros_clearing = function
- | [] -> tclIDTAC
- | (false::tl) -> tclTHEN intro (intros_clearing tl)
+ | [] -> Proofview.tclUNIT ()
+ | (false::tl) -> Tacticals.New.tclTHEN intro (intros_clearing tl)
| (true::tl) ->
- tclTHENLIST
- [ intro; onLastHypId (fun id -> clear [id]); intros_clearing tl]
+ Tacticals.New.tclTHENLIST
+ [ intro; Tacticals.New.onLastHypId (fun id -> Proofview.V82.tactic (clear [id])); intros_clearing tl]
(* Modifying/Adding an hypothesis *)
@@ -1253,16 +1277,17 @@ let check_number_of_constructors expctdnumopt i nconstr =
end;
if i > nconstr then error "Not enough constructors."
-let constructor_tac with_evars expctdnumopt i lbind gl =
- let cl = pf_concl gl in
- let (mind,redcl) = pf_reduce_to_quantified_ind gl cl in
+let constructor_tac with_evars expctdnumopt i lbind =
+ Goal.concl >>- fun cl ->
+ Tacmach.New.pf_apply Tacred.reduce_to_quantified_ind >>- fun reduce_to_quantified_ind ->
+ let (mind,redcl) = reduce_to_quantified_ind cl in
let nconstr =
Array.length (snd (Global.lookup_inductive mind)).mind_consnames in
check_number_of_constructors expctdnumopt i nconstr;
let cons = mkConstruct (ith_constructor_of_inductive mind i) in
- let apply_tac = general_apply true false with_evars (dloc,(cons,lbind)) in
- (tclTHENLIST
- [convert_concl_no_check redcl DEFAULTcast; intros; apply_tac]) gl
+ let apply_tac = Proofview.V82.tactic (general_apply true false with_evars (dloc,(cons,lbind))) in
+ (Tacticals.New.tclTHENLIST
+ [Proofview.V82.tactic (convert_concl_no_check redcl DEFAULTcast); intros; apply_tac])
let one_constructor i lbind = constructor_tac false None i lbind
@@ -1271,21 +1296,23 @@ let one_constructor i lbind = constructor_tac false None i lbind
Should be generalize in Constructor (Fun c : I -> tactic)
*)
-let any_constructor with_evars tacopt gl =
- let t = match tacopt with None -> tclIDTAC | Some t -> t in
- let mind = fst (pf_reduce_to_quantified_ind gl (pf_concl gl)) in
+let any_constructor with_evars tacopt =
+ let t = match tacopt with None -> Proofview.tclUNIT () | Some t -> t in
+ Goal.concl >>- fun cl ->
+ Tacmach.New.pf_apply Tacred.reduce_to_quantified_ind >>- fun reduce_to_quantified_ind ->
+ let mind = fst (reduce_to_quantified_ind cl) in
let nconstr =
Array.length (snd (Global.lookup_inductive mind)).mind_consnames in
if Int.equal nconstr 0 then error "The type has no constructors.";
- tclFIRST
+ Tacticals.New.tclFIRST
(List.map
- (fun i -> tclTHEN (constructor_tac with_evars None i NoBindings) t)
- (List.interval 1 nconstr)) gl
+ (fun i -> Tacticals.New.tclTHEN (constructor_tac with_evars None i NoBindings) t)
+ (List.interval 1 nconstr))
let left_with_bindings with_evars = constructor_tac with_evars (Some 2) 1
let right_with_bindings with_evars = constructor_tac with_evars (Some 2) 2
let split_with_bindings with_evars l =
- tclMAP (constructor_tac with_evars (Some 1) 1) l
+ Tacticals.New.tclMAP (constructor_tac with_evars (Some 1) 1) l
let left = left_with_bindings false
let simplest_left = left NoBindings
@@ -1341,49 +1368,52 @@ let intro_decomp_eq loc b l l' thin tac id gl =
let _,t = pf_reduce_to_quantified_ind gl (pf_type_of gl c) in
let eq,eq_args = my_find_eq_data_decompose gl t in
let eq_clause = make_clenv_binding gl (c,t) NoBindings in
- !intro_decomp_eq_function
+ Proofview.V82.of_tactic (!intro_decomp_eq_function
(fun n -> tac ((dloc,id)::thin) (adjust_intro_patterns n l @ l'))
- (eq,t,eq_args) eq_clause gl
+ (eq,t,eq_args) eq_clause) gl
-let intro_or_and_pattern loc b ll l' thin tac id gl =
+let intro_or_and_pattern loc b ll l' thin tac id =
let c = mkVar id in
- let ind,t = pf_reduce_to_quantified_ind gl (pf_type_of gl c) in
+ Tacmach.New.of_old (fun gl ->
+ pf_reduce_to_quantified_ind gl (pf_type_of gl c)) >>- fun (ind,t) ->
let nv = mis_constr_nargs ind in
let bracketed = b || not (List.is_empty l') in
let adjust n l = if bracketed then adjust_intro_patterns n l else l in
let ll = fix_empty_or_and_pattern (Array.length nv) ll in
check_or_and_pattern_size loc ll (Array.length nv);
- tclTHENLASTn
- (tclTHEN (simplest_case c) (clear [id]))
+ Tacticals.New.tclTHENLASTn
+ (Tacticals.New.tclTHEN (simplest_case c) (Proofview.V82.tactic (clear [id])))
(Array.map2 (fun n l -> tac thin ((adjust n l)@l'))
nv (Array.of_list ll))
- gl
-let rewrite_hyp l2r id gl =
+let rewrite_hyp l2r id =
let rew_on l2r =
Hook.get forward_general_multi_rewrite l2r false (mkVar id,NoBindings) in
let subst_on l2r x rhs =
Hook.get forward_subst_one true x (id,rhs,l2r) in
let clear_var_and_eq c =
tclTRY (tclTHEN (clear [id]) (tclTRY (clear [destVar c]))) in
- let t = pf_whd_betadeltaiota gl (pf_type_of gl (mkVar id)) in
+ Goal.env >>- fun env ->
+ Tacmach.New.pf_apply Typing.type_of >>- fun type_of ->
+ Tacmach.New.pf_apply whd_betadeltaiota >>- fun whd_betadeltaiota ->
+ let t = whd_betadeltaiota (type_of (mkVar id)) in
(* TODO: detect setoid equality? better detect the different equalities *)
match match_with_equality_type t with
| Some (hdcncl,[_;lhs;rhs]) ->
- if l2r && isVar lhs && not (occur_var (pf_env gl) (destVar lhs) rhs) then
- subst_on l2r (destVar lhs) rhs gl
- else if not l2r && isVar rhs && not (occur_var (pf_env gl) (destVar rhs) lhs) then
- subst_on l2r (destVar rhs) lhs gl
+ if l2r && isVar lhs && not (occur_var env (destVar lhs) rhs) then
+ subst_on l2r (destVar lhs) rhs
+ else if not l2r && isVar rhs && not (occur_var env (destVar rhs) lhs) then
+ subst_on l2r (destVar rhs) lhs
else
- tclTHEN (rew_on l2r onConcl) (tclTRY (clear [id])) gl
+ Tacticals.New.tclTHEN (rew_on l2r onConcl) (Proofview.V82.tactic (tclTRY (clear [id])))
| Some (hdcncl,[c]) ->
let l2r = not l2r in (* equality of the form eq_true *)
if isVar c then
- tclTHEN (rew_on l2r allHypsAndConcl) (clear_var_and_eq c) gl
+ Tacticals.New.tclTHEN (rew_on l2r allHypsAndConcl) (Proofview.V82.tactic (clear_var_and_eq c))
else
- tclTHEN (rew_on l2r onConcl) (tclTRY (clear [id])) gl
+ Tacticals.New.tclTHEN (rew_on l2r onConcl) (Proofview.V82.tactic (tclTRY (clear [id])))
| _ ->
- error "Cannot find a known equation."
+ Proofview.tclZERO (Errors.UserError ("",Pp.str"Cannot find a known equation."))
let rec explicit_intro_names = function
| (_, IntroIdentifier id) :: l ->
@@ -1408,7 +1438,7 @@ let check_thin_clash_then id thin avoid tac =
let newid = next_ident_away (add_suffix id "'") avoid in
let thin =
List.map (on_snd (fun id' -> if Id.equal id id' then newid else id')) thin in
- tclTHEN (rename_hyp [id,newid]) (tac thin)
+ Tacticals.New.tclTHEN (Proofview.V82.tactic (rename_hyp [id,newid])) (tac thin)
else
tac thin
@@ -1443,26 +1473,26 @@ let rec intros_patterns b avoid ids thin destopt tac = function
(intro_or_and_pattern loc b ll l' thin
(fun thin -> intros_patterns b avoid ids thin destopt tac))
| (loc, IntroInjection l) :: l' ->
- intro_then_force
- (intro_decomp_eq loc b l l' thin
- (fun thin -> intros_patterns b avoid ids thin destopt tac))
+ intro_then_force (fun id ->
+ (Proofview.V82.tactic (intro_decomp_eq loc b l l' thin
+ (fun thin -> intros_patterns b avoid ids thin destopt tac) id)))
| (loc, IntroRewrite l2r) :: l ->
intro_then_gen loc (IntroAvoid(avoid@explicit_intro_names l))
MoveLast true false
(fun id ->
- tclTHENLAST (* Skip the side conditions of the rewriting step *)
+ Tacticals.New.tclTHENLAST (* Skip the side conditions of the rewriting step *)
(rewrite_hyp l2r id)
(intros_patterns b avoid ids thin destopt tac l))
| [] -> tac ids thin
let intros_pattern destopt =
- intros_patterns false [] [] [] destopt (fun _ -> clear_wildcards)
+ intros_patterns false [] [] [] destopt (fun _ l -> Proofview.V82.tactic (clear_wildcards l))
let intro_pattern destopt pat =
intros_pattern destopt [dloc,pat]
let intro_patterns = function
- | [] -> tclREPEAT intro
+ | [] -> Tacticals.New.tclREPEAT intro
| l -> intros_pattern MoveLast l
(**************************)
@@ -1471,24 +1501,42 @@ let intro_patterns = function
let make_id s = fresh_id [] (default_id_of_sort s)
-let prepare_intros s ipat gl = match ipat with
- | None -> make_id s gl, tclIDTAC
+let prepare_intros s ipat =
+ let make_id s = Tacmach.New.of_old (make_id s) in
+ let fresh_id l id = Tacmach.New.of_old (fresh_id l id) in
+ match ipat with
+ | None ->
+ make_id s >- fun id ->
+ Goal.return (id , Proofview.tclUNIT ())
| Some (loc,ipat) -> match ipat with
- | IntroIdentifier id -> id, tclIDTAC
- | IntroAnonymous -> make_id s gl, tclIDTAC
- | IntroFresh id -> fresh_id [] id gl, tclIDTAC
- | IntroWildcard -> let id = make_id s gl in id, clear_wildcards [dloc,id]
+ | IntroIdentifier id ->
+ Goal.return (id, Proofview.tclUNIT ())
+ | IntroAnonymous ->
+ make_id s >- fun id ->
+ Goal.return (id , Proofview.tclUNIT ())
+ | IntroFresh id ->
+ fresh_id [] id >- fun id ->
+ Goal.return (id , Proofview.tclUNIT ())
+ | IntroWildcard ->
+ make_id s >- fun id ->
+ Goal.return (id , Proofview.V82.tactic (clear_wildcards [dloc,id]))
| IntroRewrite l2r ->
- let id = make_id s gl in
- id, Hook.get forward_general_multi_rewrite l2r false (mkVar id,NoBindings) allHypsAndConcl
- | IntroOrAndPattern ll -> make_id s gl,
- onLastHypId
+ make_id s >- fun id ->
+ Goal.return (id, Hook.get forward_general_multi_rewrite l2r false (mkVar id,NoBindings) allHypsAndConcl)
+ | IntroOrAndPattern ll ->
+ make_id s >- fun id ->
+ Goal.return (id ,
+ Tacticals.New.onLastHypId
(intro_or_and_pattern loc true ll [] []
- (fun thin -> intros_patterns true [] [] thin MoveLast (fun _ -> clear_wildcards)))
- | IntroInjection l -> make_id s gl,
- onLastHypId
+ (fun thin -> intros_patterns true [] [] thin MoveLast (fun _ l -> Proofview.V82.tactic (clear_wildcards l))))
+ )
+ | IntroInjection l ->
+ make_id s >- fun id ->
+ Goal.return (id ,
+ Proofview.V82.tactic (onLastHypId
(intro_decomp_eq loc true l [] []
- (fun thin -> intros_patterns true [] [] thin MoveLast (fun _ -> clear_wildcards)))
+ (fun thin -> intros_patterns true [] [] thin MoveLast (fun _ l -> Proofview.V82.tactic (clear_wildcards l))))
+ ))
| IntroForthcoming _ -> user_err_loc
(loc,"",str "Introduction pattern for one hypothesis expected")
@@ -1508,15 +1556,16 @@ let clear_if_overwritten c ipats =
| [id] -> thin [id]
| _ -> tclIDTAC
-let assert_as first ipat c gl =
- match kind_of_term (pf_hnf_type_of gl c) with
+let assert_as first ipat c =
+ Tacmach.New.of_old pf_hnf_type_of >>- fun hnf_type_of ->
+ match kind_of_term (hnf_type_of c) with
| Sort s ->
- let id,tac = prepare_intros s ipat gl in
+ prepare_intros s ipat >>- fun (id,tac) ->
let repl = not (Option.is_empty (allow_replace c ipat)) in
- tclTHENS
- ((if first then internal_cut_gen else internal_cut_rev_gen) repl id c)
- (if first then [tclIDTAC; tac] else [tac; tclIDTAC]) gl
- | _ -> error "Not a proposition or a type."
+ Tacticals.New.tclTHENS
+ (Proofview.V82.tactic ((if first then internal_cut_gen else internal_cut_rev_gen) repl id c))
+ (if first then [Proofview.tclUNIT (); tac] else [tac; Proofview.tclUNIT ()])
+ | _ -> Proofview.tclZERO (Errors.UserError ("",str"Not a proposition or a type."))
let assert_tac na = assert_as true (ipat_of_name na)
@@ -1527,35 +1576,37 @@ let as_tac id ipat = match ipat with
Hook.get forward_general_multi_rewrite l2r false (mkVar id,NoBindings) allHypsAndConcl
| Some (loc,IntroOrAndPattern ll) ->
intro_or_and_pattern loc true ll [] []
- (fun thin -> intros_patterns true [] [] thin MoveLast (fun _ -> clear_wildcards))
+ (fun thin -> intros_patterns true [] [] thin MoveLast (fun _ l -> Proofview.V82.tactic (clear_wildcards l)))
id
| Some (loc,IntroInjection l) ->
- intro_decomp_eq loc true l [] []
- (fun thin -> intros_patterns true [] [] thin MoveLast (fun _ -> clear_wildcards))
- id
+ Proofview.V82.tactic (intro_decomp_eq loc true l [] []
+ (fun thin -> intros_patterns true [] [] thin MoveLast (fun _ l -> Proofview.V82.tactic (clear_wildcards l)))
+ id)
| Some (loc,
(IntroIdentifier _ | IntroAnonymous | IntroFresh _ |
IntroWildcard | IntroForthcoming _)) ->
user_err_loc (loc,"", str "Disjunctive/conjunctive pattern expected")
- | None -> tclIDTAC
+ | None -> Proofview.tclUNIT ()
let tclMAPLAST tacfun l =
- List.fold_right (fun x -> tclTHENLAST (tacfun x)) l tclIDTAC
+ let tacfun x = Proofview.V82.tactic (tacfun x) in
+ List.fold_right (fun x -> Tacticals.New.tclTHENLAST (tacfun x)) l (Proofview.tclUNIT())
let tclMAPFIRST tacfun l =
- List.fold_right (fun x -> tclTHENFIRST (tacfun x)) l tclIDTAC
+ let tacfun x = Proofview.V82.tactic (tacfun x) in
+ List.fold_right (fun x -> Tacticals.New.tclTHENFIRST (tacfun x)) l (Proofview.tclUNIT())
let general_apply_in sidecond_first with_delta with_destruct with_evars
id lemmas ipat =
if sidecond_first then
(* Skip the side conditions of the applied lemma *)
- tclTHENLAST
+ Tacticals.New.tclTHENLAST
(tclMAPLAST
(apply_in_once sidecond_first with_delta with_destruct with_evars id)
lemmas)
(as_tac id ipat)
else
- tclTHENFIRST
+ Tacticals.New.tclTHENFIRST
(tclMAPFIRST
(apply_in_once sidecond_first with_delta with_destruct with_evars id)
lemmas)
@@ -1829,56 +1880,69 @@ let letin_abstract id c (test,out) (occs,check_occs) gl =
if List.is_empty depdecls then MoveLast else MoveAfter(pi1(List.last depdecls)) in
(depdecls,lastlhyp,ccl,out test)
-let letin_tac_gen with_eq name (sigmac,c) test ty occs gl =
+let letin_tac_gen with_eq name (sigmac,c) test ty occs =
+ Goal.env >>- fun env ->
+ Goal.hyps >>- fun hyps ->
+ let hyps = Environ.named_context_of_val hyps in
let id =
- let t = match ty with Some t -> t | None -> typ_of (pf_env gl) sigmac c in
+ let t = match ty with Some t -> t | None -> typ_of env sigmac c in
let x = id_of_name_using_hdchar (Global.env()) t name in
- if name == Anonymous then fresh_id [] x gl else
- if not (mem_named_context x (pf_hyps gl)) then x else
+ if name == Anonymous then Tacmach.New.of_old (fresh_id [] x) else
+ if not (mem_named_context x hyps) then Goal.return x else
error ("The variable "^(Id.to_string x)^" is already declared.") in
- let (depdecls,lastlhyp,ccl,c) = letin_abstract id c test occs gl in
- let t = match ty with Some t -> t | None -> pf_apply typ_of gl c in
- let newcl,eq_tac = match with_eq with
+ id >>- fun id ->
+ Tacmach.New.of_old (letin_abstract id c test occs) >>- fun (depdecls,lastlhyp,ccl,c) ->
+ let t = match ty with Some t -> Goal.return t | None -> Tacmach.New.pf_apply (fun e s -> typ_of e s c) in
+ t >>- fun t ->
+ let newcl = match with_eq with
| Some (lr,(loc,ido)) ->
let heq = match ido with
- | IntroAnonymous -> fresh_id [id] (add_prefix "Heq" id) gl
- | IntroFresh heq_base -> fresh_id [id] heq_base gl
- | IntroIdentifier id -> id
+ | IntroAnonymous -> Tacmach.New.of_old (fresh_id [id] (add_prefix "Heq" id))
+ | IntroFresh heq_base -> Tacmach.New.of_old (fresh_id [id] heq_base)
+ | IntroIdentifier id -> Goal.return id
| _ -> error"Expect an introduction pattern naming one hypothesis." in
+ heq >- fun heq ->
let eqdata = build_coq_eq_data () in
let args = if lr then [t;mkVar id;c] else [t;c;mkVar id]in
let eq = applist (eqdata.eq,args) in
let refl = applist (eqdata.refl, [t;mkVar id]) in
- mkNamedLetIn id c t (mkLetIn (Name heq, refl, eq, ccl)),
- tclTHEN
- (intro_gen loc (IntroMustBe heq) lastlhyp true false)
- (thin_body [heq;id])
+ Goal.return begin
+ mkNamedLetIn id c t (mkLetIn (Name heq, refl, eq, ccl)),
+ Tacticals.New.tclTHEN
+ (intro_gen loc (IntroMustBe heq) lastlhyp true false)
+ (Proofview.V82.tactic (thin_body [heq;id]))
+ end
| None ->
- mkNamedLetIn id c t ccl, tclIDTAC in
- tclTHENLIST
- [ convert_concl_no_check newcl DEFAULTcast;
+ Goal.return (mkNamedLetIn id c t ccl, Proofview.tclUNIT ()) in
+ newcl >>- fun (newcl,eq_tac) ->
+ Tacticals.New.tclTHENLIST
+ [ Proofview.V82.tactic (convert_concl_no_check newcl DEFAULTcast);
intro_gen dloc (IntroMustBe id) lastlhyp true false;
- tclMAP convert_hyp_no_check depdecls;
- eq_tac ] gl
+ Proofview.V82.tactic (tclMAP convert_hyp_no_check depdecls);
+ eq_tac ]
let make_eq_test c = (make_eq_test c,fun _ -> c)
-let letin_tac with_eq name c ty occs gl =
- letin_tac_gen with_eq name (project gl,c) (make_eq_test c) ty (occs,true) gl
+let letin_tac with_eq name c ty occs =
+ Goal.defs >>- fun sigma ->
+ letin_tac_gen with_eq name (sigma,c) (make_eq_test c) ty (occs,true)
-let letin_pat_tac with_eq name c ty occs gl =
+let letin_pat_tac with_eq name c ty occs =
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun sigma ->
letin_tac_gen with_eq name c
- (make_pattern_test (pf_env gl) (project gl) c)
- ty (occs,true) gl
+ (make_pattern_test env sigma c)
+ ty (occs,true)
(* Tactics "pose proof" (usetac=None) and "assert" (otherwise) *)
-let forward usetac ipat c gl =
+let forward usetac ipat c =
match usetac with
| None ->
- let t = pf_type_of gl c in
- tclTHENFIRST (assert_as true ipat t) (exact_no_check c) gl
+ Tacmach.New.pf_apply Typing.type_of >>- fun type_of ->
+ let t = type_of c in
+ Tacticals.New.tclTHENFIRST (assert_as true ipat t) (Proofview.V82.tactic (exact_no_check c))
| Some tac ->
- tclTHENFIRST (assert_as true ipat c) tac gl
+ Tacticals.New.tclTHENFIRST (assert_as true ipat c) tac
let pose_proof na c = forward None (ipat_of_name na) c
let assert_by na t tac = forward (Some tac) (ipat_of_name na) t
@@ -1977,18 +2041,22 @@ let re_intro_dependent_hypotheses (lstatus,rstatus) (_,tophyp) =
let newlstatus = (* if some IH has taken place at the top of hyps *)
List.map (function (hyp,MoveLast) -> (hyp,tophyp) | x -> x) lstatus
in
- tclTHEN
+ Tacticals.New.tclTHEN
(intros_move rstatus)
(intros_move newlstatus)
-let safe_dest_intros_patterns avoid thin dest pat tac gl =
- try intros_patterns true avoid [] thin dest tac pat gl
- with UserError ("move_hyp",_) ->
+let safe_dest_intros_patterns avoid thin dest pat tac =
+ Proofview.tclORELSE
+ (intros_patterns true avoid [] thin dest tac pat)
+ begin function
+ | UserError ("move_hyp",_) ->
(* May happen if the lemma has dependent arguments that are resolved
only after cook_sign is called, e.g. as in "destruct dec" in context
"dec:forall x, {x=0}+{x<>0}; a:A |- if dec a then True else False"
where argument a of dec will be found only lately *)
- intros_patterns true avoid [] [] MoveLast tac pat gl
+ intros_patterns true avoid [] [] MoveLast tac pat
+ | e -> Proofview.tclZERO e
+ end
type elim_arg_kind = RecArg | IndArg | OtherArg
@@ -2018,45 +2086,48 @@ let get_recarg_dest (recargdests,tophyp) =
had to be introduced at the top of the context).
*)
-let induct_discharge dests avoid' tac (avoid,ra) names gl =
+let induct_discharge dests avoid' tac (avoid,ra) names =
let avoid = avoid @ avoid' in
- let rec peel_tac ra dests names thin gl =
+ let rec peel_tac ra dests names thin =
match ra with
| (RecArg,deprec,recvarname) ::
(IndArg,depind,hyprecname) :: ra' ->
- let recpat,names = match names with
+ let recpat = match names with
| [loc,IntroIdentifier id as pat] ->
let id' = next_ident_away (add_prefix "IH" id) avoid in
- (pat, [dloc, IntroIdentifier id'])
- | _ -> consume_pattern avoid recvarname deprec gl names in
+ Goal.return (pat, [dloc, IntroIdentifier id'])
+ | _ -> Tacmach.New.of_old (fun gl -> consume_pattern avoid recvarname deprec gl names) in
+ recpat >>- fun (recpat,names) ->
let dest = get_recarg_dest dests in
- safe_dest_intros_patterns avoid thin dest [recpat] (fun ids thin gl ->
- let hyprec,names = consume_pattern avoid hyprecname depind gl names in
+ safe_dest_intros_patterns avoid thin dest [recpat] (fun ids thin ->
+ let hyprec = Tacmach.New.of_old (fun gl -> consume_pattern avoid hyprecname depind gl names) in
+ hyprec >>- fun (hyprec,names) ->
safe_dest_intros_patterns avoid thin MoveLast [hyprec] (fun ids' thin ->
- peel_tac ra' (update_dest dests ids') names thin) gl)
- gl
+ peel_tac ra' (update_dest dests ids') names thin))
| (IndArg,dep,hyprecname) :: ra' ->
(* Rem: does not happen in Coq schemes, only in user-defined schemes *)
- let pat,names = consume_pattern avoid hyprecname dep gl names in
+ let pat = Tacmach.New.of_old (fun gl -> consume_pattern avoid hyprecname dep gl names) in
+ pat >>- fun (pat,names) ->
safe_dest_intros_patterns avoid thin MoveLast [pat] (fun ids thin ->
- peel_tac ra' (update_dest dests ids) names thin) gl
+ peel_tac ra' (update_dest dests ids) names thin)
| (RecArg,dep,recvarname) :: ra' ->
- let pat,names = consume_pattern avoid recvarname dep gl names in
+ let pat = Tacmach.New.of_old (fun gl -> consume_pattern avoid recvarname dep gl names) in
+ pat >>- fun (pat,names) ->
let dest = get_recarg_dest dests in
safe_dest_intros_patterns avoid thin dest [pat] (fun ids thin ->
- peel_tac ra' dests names thin) gl
+ peel_tac ra' dests names thin)
| (OtherArg,_,_) :: ra' ->
let pat,names = match names with
| [] -> (dloc, IntroAnonymous), []
| pat::names -> pat,names in
let dest = get_recarg_dest dests in
safe_dest_intros_patterns avoid thin dest [pat] (fun ids thin ->
- peel_tac ra' dests names thin) gl
+ peel_tac ra' dests names thin)
| [] ->
check_unused_names names;
- tclTHEN (clear_wildcards thin) (tac dests) gl
+ Tacticals.New.tclTHEN (Proofview.V82.tactic (clear_wildcards thin)) (tac dests)
in
- peel_tac ra dests names [] gl
+ peel_tac ra dests names []
(* - le recalcul de indtyp à chaque itération de atomize_one est pour ne pas
s'embêter à regarder si un letin_tac ne fait pas des
@@ -2064,40 +2135,44 @@ let induct_discharge dests avoid' tac (avoid,ra) names gl =
(* Marche pas... faut prendre en compte l'occurrence précise... *)
-let atomize_param_of_ind (indref,nparams,_) hyp0 gl =
- let tmptyp0 = pf_get_hyp_typ gl hyp0 in
- let typ0 = pf_apply reduce_to_quantified_ref gl indref tmptyp0 in
+let atomize_param_of_ind (indref,nparams,_) hyp0 =
+ Tacmach.New.pf_get_hyp_typ hyp0 >>- fun tmptyp0 ->
+ Tacmach.New.pf_apply reduce_to_quantified_ref >>- fun reduce_to_quantified_ref ->
+ let typ0 = reduce_to_quantified_ref indref tmptyp0 in
let prods, indtyp = decompose_prod typ0 in
let argl = snd (decompose_app indtyp) in
let params = List.firstn nparams argl in
(* le gl est important pour ne pas préévaluer *)
- let rec atomize_one i avoid gl =
+ let rec atomize_one i avoid =
if not (Int.equal i nparams) then
- let tmptyp0 = pf_get_hyp_typ gl hyp0 in
+ Tacmach.New.pf_get_hyp_typ hyp0 >>- fun tmptyp0 ->
(* If argl <> [], we expect typ0 not to be quantified, in order to
avoid bound parameters... then we call pf_reduce_to_atomic_ind *)
- let indtyp = pf_apply reduce_to_atomic_ref gl indref tmptyp0 in
+ Tacmach.New.pf_apply reduce_to_atomic_ref >>- fun reduce_to_atomic_ref ->
+ let indtyp = reduce_to_atomic_ref indref tmptyp0 in
let argl = snd (decompose_app indtyp) in
let c = List.nth argl (i-1) in
+ Goal.env >>- fun env ->
match kind_of_term c with
- | Var id when not (List.exists (occur_var (pf_env gl) id) avoid) ->
- atomize_one (i-1) ((mkVar id)::avoid) gl
+ | Var id when not (List.exists (occur_var env id) avoid) ->
+ atomize_one (i-1) ((mkVar id)::avoid)
| Var id ->
- let x = fresh_id [] id gl in
- tclTHEN
+ Tacmach.New.of_old (fresh_id [] id) >>- fun x ->
+ Tacticals.New.tclTHEN
(letin_tac None (Name x) (mkVar id) None allHypsAndConcl)
- (atomize_one (i-1) ((mkVar x)::avoid)) gl
+ (atomize_one (i-1) ((mkVar x)::avoid))
| _ ->
- let id = id_of_name_using_hdchar (Global.env()) (pf_type_of gl c)
+ Tacmach.New.pf_apply Typing.type_of >>- fun type_of ->
+ let id = id_of_name_using_hdchar (Global.env()) (type_of c)
Anonymous in
- let x = fresh_id [] id gl in
- tclTHEN
+ Tacmach.New.of_old (fresh_id [] id) >>- fun x ->
+ Tacticals.New.tclTHEN
(letin_tac None (Name x) c None allHypsAndConcl)
- (atomize_one (i-1) ((mkVar x)::avoid)) gl
+ (atomize_one (i-1) ((mkVar x)::avoid))
else
- tclIDTAC gl
+ Proofview.tclUNIT ()
in
- atomize_one (List.length argl) params gl
+ atomize_one (List.length argl) params
let find_atomic_param_of_ind nparams indtyp =
let argl = snd (decompose_app indtyp) in
@@ -2553,38 +2628,42 @@ let abstract_args gl generalize_vars dep id defined f args =
Some (make_abstract_generalize gl id concl dep ctx body c' eqs args refls,
dep, succ (List.length ctx), vars)
else None
-
-let abstract_generalize ?(generalize_vars=true) ?(force_dep=false) id gl =
+
+let abstract_generalize ?(generalize_vars=true) ?(force_dep=false) id =
+ Proofview.tclUNIT () >= fun () -> (* delay for [check_required_library] *)
Coqlib.check_required_library Coqlib.jmeq_module_name;
- let f, args, def, id, oldid =
- let oldid = pf_get_new_id id gl in
- let (_, b, t) = pf_get_hyp gl id in
- match b with
+ let args =
+ Tacmach.New.pf_get_new_id id >>-- fun oldid ->
+ Tacmach.New.pf_get_hyp id >>-- fun (_, b, t) ->
+ Proofview.tclUNIT
+ begin match b with
| None -> let f, args = decompose_app t in
- f, args, false, id, oldid
+ Goal.return (f, args, false, id, oldid)
| Some t ->
let f, args = decompose_app t in
- f, args, true, id, oldid
+ Goal.return (f, args, true, id, oldid)
+ end
in
- if List.is_empty args then tclIDTAC gl
+ args >>= fun (f, args, def, id, oldid) ->
+ if List.is_empty args then Proofview.tclUNIT ()
else
let args = Array.of_list args in
- let newc = abstract_args gl generalize_vars force_dep id def f args in
+ Tacmach.New.of_old (fun gl -> abstract_args gl generalize_vars force_dep id def f args) >>- fun newc ->
match newc with
- | None -> tclIDTAC gl
+ | None -> Proofview.tclUNIT ()
| Some (newc, dep, n, vars) ->
let tac =
if dep then
- tclTHENLIST [refine newc; rename_hyp [(id, oldid)]; tclDO n intro;
- generalize_dep ~with_let:true (mkVar oldid)]
+ Tacticals.New.tclTHENLIST [Proofview.V82.tactic (refine newc); Proofview.V82.tactic (rename_hyp [(id, oldid)]); Tacticals.New.tclDO n intro;
+ Proofview.V82.tactic (generalize_dep ~with_let:true (mkVar oldid))]
else
- tclTHENLIST [refine newc; clear [id]; tclDO n intro]
- in
- if List.is_empty vars then tac gl
- else tclTHEN tac
- (fun gl -> tclFIRST [revert vars ;
+ Tacticals.New.tclTHENLIST [Proofview.V82.tactic (refine newc); Proofview.V82.tactic (clear [id]); Tacticals.New.tclDO n intro]
+ in
+ if List.is_empty vars then tac
+ else Tacticals.New.tclTHEN tac
+ (Proofview.V82.tactic (fun gl -> tclFIRST [revert vars ;
tclMAP (fun id ->
- tclTRY (generalize_dep ~with_let:true (mkVar id))) vars] gl) gl
+ tclTRY (generalize_dep ~with_let:true (mkVar id))) vars] gl))
let rec compare_upto_variables x y =
if (isVar x || isRel x) && (isVar y || isRel y) then true
@@ -3010,39 +3089,39 @@ let induction_tac_felim with_evars indvars nparams elim gl =
(* Apply induction "in place" replacing the hypothesis on which
induction applies with the induction hypotheses *)
-let apply_induction_with_discharge induct_tac elim indhyps destopt avoid names tac gl =
- let isrec, elim, indsign = get_eliminator elim gl in
+let apply_induction_with_discharge induct_tac elim indhyps destopt avoid names tac =
+ Tacmach.New.of_old (get_eliminator elim) >>- fun (isrec, elim, indsign) ->
let names = compute_induction_names (Array.length indsign) names in
- (if isrec then tclTHENFIRSTn else tclTHENLASTn)
- (tclTHEN
+ (if isrec then Tacticals.New.tclTHENFIRSTn else Tacticals.New.tclTHENLASTn)
+ (Tacticals.New.tclTHEN
(induct_tac elim)
- (tclMAP (fun id -> tclTRY (expand_hyp id)) (List.rev indhyps)))
- (Array.map2 (induct_discharge destopt avoid tac) indsign names) gl
+ (Proofview.V82.tactic (tclMAP (fun id -> tclTRY (expand_hyp id)) (List.rev indhyps))))
+ (Array.map2 (induct_discharge destopt avoid tac) indsign names)
(* Apply induction "in place" taking into account dependent
hypotheses from the context *)
-let apply_induction_in_context hyp0 elim indvars names induct_tac gl =
- let env = pf_env gl in
+let apply_induction_in_context hyp0 elim indvars names induct_tac =
+ Goal.env >>- fun env ->
+ Goal.concl >>- fun concl ->
let statuslists,lhyp0,indhyps,deps = cook_sign hyp0 indvars env in
let deps = List.map (on_pi3 refresh_universes_strict) deps in
- let tmpcl = it_mkNamedProd_or_LetIn (pf_concl gl) deps in
+ let tmpcl = it_mkNamedProd_or_LetIn concl deps in
let dephyps = List.map (fun (id,_,_) -> id) deps in
let deps_cstr =
List.fold_left
(fun a (id,b,_) -> if Option.is_empty b then (mkVar id)::a else a) [] deps in
- tclTHENLIST
+ Tacticals.New.tclTHENLIST
[
(* Generalize dependent hyps (but not args) *)
- if List.is_empty deps then tclIDTAC else apply_type tmpcl deps_cstr;
+ if deps = [] then Proofview.tclUNIT () else Proofview.V82.tactic (apply_type tmpcl deps_cstr);
(* clear dependent hyps *)
- thin dephyps;
+ Proofview.V82.tactic (thin dephyps);
(* side-conditions in elim (resp case) schemes come last (resp first) *)
apply_induction_with_discharge
induct_tac elim (List.rev indhyps) lhyp0 (List.rev dephyps) names
(re_intro_dependent_hypotheses statuslists)
]
- gl
(* Induction with several induction arguments, main differences with
induction_from_context is that there is no main induction argument,
@@ -3050,7 +3129,7 @@ let apply_induction_in_context hyp0 elim indvars names induct_tac gl =
all args and params must be given, so we help a bit the unifier by
making the "pattern" by hand before calling induction_tac_felim
FIXME: REUNIF AVEC induction_tac_felim? *)
-let induction_from_context_l with_evars elim_info lid names gl =
+let induction_from_context_l with_evars elim_info lid names =
let indsign,scheme = elim_info in
(* number of all args, counting farg and indarg if present. *)
let nargs_indarg_farg = scheme.nargs
@@ -3078,7 +3157,7 @@ let induction_from_context_l with_evars elim_info lid names gl =
let realindvars = (* hyp0 is a real induction arg if it is not the
farg in the conclusion of the induction scheme *)
List.rev ((if scheme.farg_in_concl then indvars else hyp0::indvars) @ lid_params) in
- let induct_tac elim = tclTHENLIST [
+ let induct_tac elim = Proofview.V82.tactic (tclTHENLIST [
(* pattern to make the predicate appear. *)
reduce (Pattern (List.map inj_with_occurrences lidcstr)) onConcl;
(* Induction by "refine (indscheme ?i ?j ?k...)" + resolution of all
@@ -3086,10 +3165,10 @@ let induction_from_context_l with_evars elim_info lid names gl =
functional one). *)
(* FIXME: Tester ca avec un principe dependant et non-dependant *)
induction_tac_felim with_evars realindvars scheme.nparams elim
- ] in
+ ]) in
let elim = ElimUsing (({elimindex = Some scheme.index; elimbody = Option.get scheme.elimc}, scheme.elimt), indsign) in
apply_induction_in_context
- None elim (hyp0::indvars) names induct_tac gl
+ None elim (hyp0::indvars) names induct_tac
(* Unification between ((elimc:elimt) ?i ?j ?k ?l ... ?m) and the
hypothesis on which the induction is made *)
@@ -3103,31 +3182,31 @@ let induction_tac with_evars elim (varname,lbind) typ gl =
elimination_clause_scheme with_evars i elimclause indclause gl
let induction_from_context isrec with_evars (indref,nparams,elim) (hyp0,lbind) names
- inhyps gl =
- let tmptyp0 = pf_get_hyp_typ gl hyp0 in
- let typ0 = pf_apply reduce_to_quantified_ref gl indref tmptyp0 in
+ inhyps =
+ Tacmach.New.pf_get_hyp_typ hyp0 >>- fun tmptyp0 ->
+ Tacmach.New.pf_apply reduce_to_quantified_ref >>- fun reduce_to_quantified_ref ->
+ let typ0 = reduce_to_quantified_ref indref tmptyp0 in
let indvars = find_atomic_param_of_ind nparams ((strip_prod typ0)) in
- let induct_tac elim = tclTHENLIST [
+ let induct_tac elim = Proofview.V82.tactic (tclTHENLIST [
induction_tac with_evars elim (hyp0,lbind) typ0;
tclTRY (unfold_body hyp0);
thin [hyp0]
- ] in
+ ]) in
apply_induction_in_context
- (Some (hyp0,inhyps)) elim indvars names induct_tac gl
+ (Some (hyp0,inhyps)) elim indvars names induct_tac
-let induction_with_atomization_of_ind_arg isrec with_evars elim names (hyp0,lbind) inhyps gl =
- let elim_info = find_induction_type isrec elim hyp0 gl in
- tclTHEN
+let induction_with_atomization_of_ind_arg isrec with_evars elim names (hyp0,lbind) inhyps =
+ Tacmach.New.of_old (find_induction_type isrec elim hyp0) >>- fun elim_info ->
+ Tacticals.New.tclTHEN
(atomize_param_of_ind elim_info hyp0)
(induction_from_context isrec with_evars elim_info
- (hyp0,lbind) names inhyps) gl
+ (hyp0,lbind) names inhyps)
(* Induction on a list of induction arguments. Analyse the elim
scheme (which is mandatory for multiple ind args), check that all
parameters and arguments are given (mandatory too). *)
-let induction_without_atomization isrec with_evars elim names lid gl =
- let (indsign,scheme as elim_info) =
- find_elim_signature isrec elim (List.hd lid) gl in
+let induction_without_atomization isrec with_evars elim names lid =
+ Tacmach.New.of_old (find_elim_signature isrec elim (List.hd lid)) >>- fun (indsign,scheme as elim_info) ->
let awaited_nargs =
scheme.nparams + scheme.nargs
+ (if scheme.farg_in_concl then 1 else 0)
@@ -3135,8 +3214,8 @@ let induction_without_atomization isrec with_evars elim names lid gl =
in
let nlid = List.length lid in
if not (Int.equal nlid awaited_nargs)
- then error "Not the right number of induction arguments."
- else induction_from_context_l with_evars elim_info lid names gl
+ then Proofview.tclZERO (Errors.UserError ("", str"Not the right number of induction arguments."))
+ else induction_from_context_l with_evars elim_info lid names
let has_selected_occurrences = function
| None -> false
@@ -3167,89 +3246,93 @@ let clear_unselected_context id inhyps cls gl =
thin ids gl
| None -> tclIDTAC gl
-let new_induct_gen isrec with_evars elim (eqname,names) (sigma,(c,lbind)) cls gl =
+let new_induct_gen isrec with_evars elim (eqname,names) (sigma,(c,lbind)) cls =
let inhyps = match cls with
| Some {onhyps=Some hyps} -> List.map (fun ((_,id),_) -> id) hyps
| _ -> [] in
match kind_of_term c with
| Var id when not (mem_named_context id (Global.named_context()))
- && lbind == NoBindings && not with_evars && Option.is_empty eqname
- && not (has_selected_occurrences cls) ->
- tclTHEN
- (clear_unselected_context id inhyps cls)
+ & lbind == NoBindings & not with_evars & Option.is_empty eqname
+ & not (has_selected_occurrences cls) ->
+ Tacticals.New.tclTHEN
+ (Proofview.V82.tactic (clear_unselected_context id inhyps cls))
(induction_with_atomization_of_ind_arg
- isrec with_evars elim names (id,lbind) inhyps) gl
+ isrec with_evars elim names (id,lbind) inhyps)
| _ ->
- let x = id_of_name_using_hdchar (Global.env()) (typ_of (pf_env gl) sigma c)
+ Goal.env >>- fun env ->
+ let x = id_of_name_using_hdchar (Global.env()) (typ_of env sigma c)
Anonymous in
- let id = fresh_id [] x gl in
+ Tacmach.New.of_old (fresh_id [] x) >>- fun id ->
(* We need the equality name now *)
let with_eq = Option.map (fun eq -> (false,eq)) eqname in
(* TODO: if ind has predicate parameters, use JMeq instead of eq *)
- tclTHEN
+ Goal.env >>- fun env ->
+ Goal.defs >>- fun defs ->
+ Tacticals.New.tclTHEN
(* Warning: letin is buggy when c is not of inductive type *)
(letin_tac_gen with_eq (Name id) (sigma,c)
- (make_pattern_test (pf_env gl) (project gl) (sigma,c))
+ (make_pattern_test env defs (sigma,c))
None (Option.default allHypsAndConcl cls,false))
(induction_with_atomization_of_ind_arg
- isrec with_evars elim names (id,lbind) inhyps) gl
+ isrec with_evars elim names (id,lbind) inhyps)
(* Induction on a list of arguments. First make induction arguments
atomic (using letins), then do induction. The specificity here is
that all arguments and parameters of the scheme are given
(mandatory for the moment), so we don't need to deal with
parameters of the inductive type as in new_induct_gen. *)
-let new_induct_gen_l isrec with_evars elim (eqname,names) lc gl =
+let new_induct_gen_l isrec with_evars elim (eqname,names) lc =
if not (Option.is_empty eqname) then
errorlabstrm "" (str "Do not know what to do with " ++
pr_intro_pattern (Option.get eqname));
let newlc = ref [] in
let letids = ref [] in
- let rec atomize_list l gl =
+ let rec atomize_list l =
match l with
- | [] -> tclIDTAC gl
+ | [] -> Proofview.tclUNIT ()
| c::l' ->
match kind_of_term c with
| Var id when not (mem_named_context id (Global.named_context()))
&& not with_evars ->
let _ = newlc:= id::!newlc in
- atomize_list l' gl
+ atomize_list l'
| _ ->
+ Tacmach.New.pf_apply Typing.type_of >>- fun type_of ->
let x =
- id_of_name_using_hdchar (Global.env()) (pf_type_of gl c) Anonymous in
+ id_of_name_using_hdchar (Global.env()) (type_of c) Anonymous in
- let id = fresh_id [] x gl in
+ Tacmach.New.of_old (fresh_id [] x) >>- fun id ->
let newl' = List.map (replace_term c (mkVar id)) l' in
let _ = newlc:=id::!newlc in
let _ = letids:=id::!letids in
- tclTHEN
+ Tacticals.New.tclTHEN
(letin_tac None (Name id) c None allHypsAndConcl)
- (atomize_list newl') gl in
- tclTHENLIST
+ (atomize_list newl') in
+ Tacticals.New.tclTHENLIST
[
(atomize_list lc);
- (fun gl' -> (* recompute each time to have the new value of newlc *)
- induction_without_atomization isrec with_evars elim names !newlc gl') ;
+ (Proofview.tclUNIT () >= fun () -> (* recompute each time to have the new value of newlc *)
+ induction_without_atomization isrec with_evars elim names !newlc) ;
(* after induction, try to unfold all letins created by atomize_list
FIXME: unfold_all does not exist anywhere else? *)
- (fun gl' -> (* recompute each time to have the new value of letids *)
- tclMAP (fun x -> tclTRY (unfold_all x)) !letids gl')
+ (Proofview.V82.tactic( fun gl' -> (* recompute each time to have the new value of letids *)
+ tclMAP (fun x -> tclTRY (unfold_all x)) !letids gl'))
]
- gl
(* Induction either over a term, over a quantified premisse, or over
several quantified premisses (like with functional induction
principles).
TODO: really unify induction with one and induction with several
args *)
-let induct_destruct isrec with_evars (lc,elim,names,cls) gl =
+let induct_destruct isrec with_evars (lc,elim,names,cls) =
assert (List.length lc > 0); (* ensured by syntax, but if called inside caml? *)
- if Int.equal (List.length lc) 1 && not (is_functional_induction elim gl) then
+ Tacmach.New.of_old (is_functional_induction elim) >>- fun ifi ->
+ if Int.equal (List.length lc) 1 && not ifi then
(* standard induction *)
onOpenInductionArg
(fun c -> new_induct_gen isrec with_evars elim names c cls)
- (List.hd lc) gl
+ (List.hd lc)
else begin
(* functional induction *)
(* Several induction hyps: induction scheme is mandatory *)
@@ -3259,8 +3342,9 @@ let induct_destruct isrec with_evars (lc,elim,names,cls) gl =
str "Example: induction x1 x2 x3 using my_scheme.");
if not (Option.is_empty cls) then
error "'in' clause not supported here.";
+ Tacmach.New.pf_apply finish_evar_resolution >>- fun finish_evar_resolution ->
let lc = List.map
- (map_induction_arg (pf_apply finish_evar_resolution gl)) lc in
+ (map_induction_arg finish_evar_resolution) lc in
begin match lc with
| [_] ->
(* Hook to recover standard induction on non-standard induction schemes *)
@@ -3269,7 +3353,7 @@ let induct_destruct isrec with_evars (lc,elim,names,cls) gl =
(fun (c,lbind) ->
if lbind != NoBindings then
error "'with' clause not supported here.";
- new_induct_gen_l isrec with_evars elim names [c]) (List.hd lc) gl
+ new_induct_gen_l isrec with_evars elim names [c]) (List.hd lc)
| _ ->
let newlc =
List.map (fun x ->
@@ -3277,17 +3361,17 @@ let induct_destruct isrec with_evars (lc,elim,names,cls) gl =
| ElimOnConstr (x,NoBindings) -> x
| _ -> error "Don't know where to find some argument.")
lc in
- new_induct_gen_l isrec with_evars elim names newlc gl
+ new_induct_gen_l isrec with_evars elim names newlc
end
end
let induction_destruct isrec with_evars = function
- | [],_,_ -> tclIDTAC
+ | [],_,_ -> Proofview.tclUNIT ()
| [a,b],el,cl -> induct_destruct isrec with_evars ([a],el,b,cl)
| (a,b)::l,None,cl ->
- tclTHEN
+ Tacticals.New.tclTHEN
(induct_destruct isrec with_evars ([a],None,b,cl))
- (tclMAP (fun (a,b) -> induct_destruct false with_evars ([a],None,b,cl)) l)
+ (Tacticals.New.tclMAP (fun (a,b) -> induct_destruct false with_evars ([a],None,b,cl)) l)
| l,Some el,cl ->
let check_basic_using = function
| a,(None,None) -> a
@@ -3305,8 +3389,8 @@ let new_destruct ev lc e idl cls = induct_destruct false ev (lc,e,idl,cls)
(* Induction tactics *)
(* This was Induction before 6.3 (induction only in quantified premisses) *)
-let simple_induct_id s = tclTHEN (intros_until_id s) (onLastHyp simplest_elim)
-let simple_induct_nodep n = tclTHEN (intros_until_n n) (onLastHyp simplest_elim)
+let simple_induct_id s = Tacticals.New.tclTHEN (intros_until_id s) (Tacticals.New.onLastHyp simplest_elim)
+let simple_induct_nodep n = Tacticals.New.tclTHEN (intros_until_n n) (Tacticals.New.onLastHyp simplest_elim)
let simple_induct = function
| NamedHyp id -> simple_induct_id id
@@ -3315,9 +3399,9 @@ let simple_induct = function
(* Destruction tactics *)
let simple_destruct_id s =
- (tclTHEN (intros_until_id s) (onLastHyp simplest_case))
+ (Tacticals.New.tclTHEN (intros_until_id s) (Tacticals.New.onLastHyp simplest_case))
let simple_destruct_nodep n =
- (tclTHEN (intros_until_n n) (onLastHyp simplest_case))
+ (Tacticals.New.tclTHEN (intros_until_n n) (Tacticals.New.onLastHyp simplest_case))
let simple_destruct = function
| NamedHyp id -> simple_destruct_id id
@@ -3366,50 +3450,53 @@ let case_type t gl =
HH (29/5/99) replaces failures by specific error messages
*)
-let andE id gl =
- let t = pf_get_hyp_typ gl id in
- if is_conjunction (pf_hnf_constr gl t) then
- (tclTHEN (simplest_elim (mkVar id)) (tclDO 2 intro)) gl
+let andE id =
+ Tacmach.New.pf_get_hyp_typ id >>- fun t ->
+ Tacmach.New.of_old pf_hnf_constr >>- fun hnf_constr ->
+ if is_conjunction (hnf_constr t) then
+ (Tacticals.New.tclTHEN (simplest_elim (mkVar id)) (Tacticals.New.tclDO 2 intro))
else
- errorlabstrm "andE"
- (str("Tactic andE expects "^(Id.to_string id)^" is a conjunction."))
+ Proofview.tclZERO (Errors.UserError (
+ "andE" , (str("Tactic andE expects "^(Id.to_string id)^" is a conjunction."))))
let dAnd cls =
- onClause
+ Tacticals.New.onClause
(function
| None -> simplest_split
| Some id -> andE id)
cls
-let orE id gl =
- let t = pf_get_hyp_typ gl id in
- if is_disjunction (pf_hnf_constr gl t) then
- (tclTHEN (simplest_elim (mkVar id)) intro) gl
+let orE id =
+ Tacmach.New.pf_get_hyp_typ id >>- fun t ->
+ Tacmach.New.of_old pf_hnf_constr >>- fun hnf_constr ->
+ if is_disjunction (hnf_constr t) then
+ (Tacticals.New.tclTHEN (simplest_elim (mkVar id)) intro)
else
- errorlabstrm "orE"
- (str("Tactic orE expects "^(Id.to_string id)^" is a disjunction."))
+ Proofview.tclZERO (Errors.UserError (
+ "orE" , (str("Tactic orE expects "^(Id.to_string id)^" is a disjunction."))))
let dorE b cls =
- onClause
+ Tacticals.New.onClause
(function
| Some id -> orE id
| None -> (if b then right else left) NoBindings)
cls
-let impE id gl =
- let t = pf_get_hyp_typ gl id in
- if is_imp_term (pf_hnf_constr gl t) then
- let (dom, _, rng) = destProd (pf_hnf_constr gl t) in
- tclTHENLAST
+let impE id =
+ Tacmach.New.pf_get_hyp_typ id >>- fun t ->
+ Tacmach.New.of_old pf_hnf_constr >>- fun hnf_constr ->
+ if is_imp_term (hnf_constr t) then
+ let (dom, _, rng) = destProd (hnf_constr t) in
+ Tacticals.New.tclTHENLAST
(cut_intro rng)
- (apply_term (mkVar id) [mkMeta (new_meta())]) gl
+ (Proofview.V82.tactic (apply_term (mkVar id) [mkMeta (new_meta())]))
else
- errorlabstrm "impE"
- (str("Tactic impE expects "^(Id.to_string id)^
- " is a an implication."))
+ Proofview.tclZERO (Errors.UserError (
+ "impE" , (str("Tactic impE expects "^(Id.to_string id)^
+ " is a an implication."))))
let dImp cls =
- onClause
+ Tacticals.New.onClause
(function
| None -> intro
| Some id -> impE id)
@@ -3423,22 +3510,29 @@ let dImp cls =
let (forward_setoid_reflexivity, setoid_reflexivity) = Hook.make ()
-let reflexivity_red allowred gl =
+let reflexivity_red allowred =
(* PL: usual reflexivity don't perform any reduction when searching
for an equality, but we may need to do some when called back from
inside setoid_reflexivity (see Optimize cases in setoid_replace.ml). *)
- let concl = if not allowred then pf_concl gl
- else whd_betadeltaiota (pf_env gl) (project gl) (pf_concl gl)
+ let concl = if not allowred then Goal.concl
+ else
+ Goal.concl >- fun c ->
+ Tacmach.New.pf_apply (fun env sigma ->whd_betadeltaiota env sigma c)
in
+ concl >>- fun concl ->
match match_with_equality_type concl with
- | None -> raise NoEquationFound
- | Some _ -> one_constructor 1 NoBindings gl
-
-let reflexivity gl =
- try reflexivity_red false gl
- with NoEquationFound -> Hook.get forward_setoid_reflexivity gl
+ | None -> Proofview.tclZERO NoEquationFound
+ | Some _ -> one_constructor 1 NoBindings
+
+let reflexivity =
+ Proofview.tclORELSE
+ (reflexivity_red false)
+ begin function
+ | NoEquationFound -> Hook.get forward_setoid_reflexivity
+ | e -> Proofview.tclZERO e
+ end
-let intros_reflexivity = (tclTHEN intros reflexivity)
+let intros_reflexivity = (Tacticals.New.tclTHEN intros reflexivity)
(* Symmetry tactics *)
@@ -3456,50 +3550,76 @@ let prove_symmetry hdcncl eq_kind =
| MonomorphicLeibnizEq (c1,c2) -> mkApp(hdcncl,[|c2;c1|])
| PolymorphicLeibnizEq (typ,c1,c2) -> mkApp(hdcncl,[|typ;c2;c1|])
| HeterogenousEq (t1,c1,t2,c2) -> mkApp(hdcncl,[|t2;c2;t1;c1|]) in
- tclTHENFIRST (cut symc)
- (tclTHENLIST
+ Tacticals.New.tclTHENFIRST (Proofview.V82.tactic (cut symc))
+ (Tacticals.New.tclTHENLIST
[ intro;
- onLastHyp simplest_case;
+ Tacticals.New.onLastHyp simplest_case;
one_constructor 1 NoBindings ])
-let symmetry_red allowred gl =
+let match_with_equation c =
+ try
+ let res = match_with_equation c in
+ Proofview.tclUNIT res
+ with NoEquationFound ->
+ Proofview.tclZERO NoEquationFound
+
+let symmetry_red allowred =
(* PL: usual symmetry don't perform any reduction when searching
for an equality, but we may need to do some when called back from
inside setoid_reflexivity (see Optimize cases in setoid_replace.ml). *)
let concl =
- if not allowred then pf_concl gl else pf_whd_betadeltaiota gl (pf_concl gl)
+ if not allowred then
+ Goal.concl
+ else
+ Goal.concl >- fun c ->
+ Tacmach.New.pf_apply (fun env sigma -> whd_betadeltaiota env sigma c)
in
- match match_with_equation concl with
+ concl >>- fun concl ->
+ match_with_equation concl >= fun with_eqn ->
+ match with_eqn with
| Some eq_data,_,_ ->
- tclTHEN
- (convert_concl_no_check concl DEFAULTcast)
- (apply eq_data.sym) gl
- | None,eq,eq_kind -> prove_symmetry eq eq_kind gl
-
-let symmetry gl =
- try symmetry_red false gl with NoEquationFound -> Hook.get forward_setoid_symmetry gl
+ Proofview.V82.tactic begin
+ tclTHEN
+ (convert_concl_no_check concl DEFAULTcast)
+ (apply eq_data.sym)
+ end
+ | None,eq,eq_kind -> prove_symmetry eq eq_kind
+
+let symmetry =
+ Proofview.tclORELSE
+ (symmetry_red false)
+ begin function
+ | NoEquationFound -> Hook.get forward_setoid_symmetry
+ | e -> Proofview.tclZERO e
+ end
let (forward_setoid_symmetry_in, setoid_symmetry_in) = Hook.make ()
-let symmetry_in id gl =
- let ctype = pf_type_of gl (mkVar id) in
+
+let symmetry_in id =
+ Tacmach.New.pf_apply Typing.type_of >>- fun type_of ->
+ let ctype = type_of (mkVar id) in
let sign,t = decompose_prod_assum ctype in
- try
- let _,hdcncl,eq = match_with_equation t in
- let symccl = match eq with
- | MonomorphicLeibnizEq (c1,c2) -> mkApp (hdcncl, [| c2; c1 |])
- | PolymorphicLeibnizEq (typ,c1,c2) -> mkApp (hdcncl, [| typ; c2; c1 |])
- | HeterogenousEq (t1,c1,t2,c2) -> mkApp (hdcncl, [| t2; c2; t1; c1 |]) in
- tclTHENS (cut (it_mkProd_or_LetIn symccl sign))
- [ intro_replacing id;
- tclTHENLIST [ intros; symmetry; apply (mkVar id); assumption ] ]
- gl
- with NoEquationFound -> Hook.get forward_setoid_symmetry_in id gl
+ Proofview.tclORELSE
+ begin
+ match_with_equation t >= fun (_,hdcncl,eq) ->
+ let symccl = match eq with
+ | MonomorphicLeibnizEq (c1,c2) -> mkApp (hdcncl, [| c2; c1 |])
+ | PolymorphicLeibnizEq (typ,c1,c2) -> mkApp (hdcncl, [| typ; c2; c1 |])
+ | HeterogenousEq (t1,c1,t2,c2) -> mkApp (hdcncl, [| t2; c2; t1; c1 |]) in
+ Tacticals.New.tclTHENS (Proofview.V82.tactic (cut (it_mkProd_or_LetIn symccl sign)))
+ [ Proofview.V82.tactic (intro_replacing id);
+ Tacticals.New.tclTHENLIST [ intros; symmetry; Proofview.V82.tactic (apply (mkVar id)); Proofview.V82.tactic assumption ] ]
+ end
+ begin function
+ | NoEquationFound -> Hook.get forward_setoid_symmetry_in id
+ | e -> Proofview.tclZERO e
+ end
let intros_symmetry =
- onClause
+ Tacticals.New.onClause
(function
- | None -> tclTHEN intros symmetry
+ | None -> Tacticals.New.tclTHEN intros symmetry
| Some id -> symmetry_in id)
(* Transitivity tactics *)
@@ -3516,52 +3636,67 @@ let intros_symmetry =
let (forward_setoid_transitivity, setoid_transitivity) = Hook.make ()
+
(* This is probably not very useful any longer *)
-let prove_transitivity hdcncl eq_kind t gl =
- let eq1,eq2 =
- match eq_kind with
- | MonomorphicLeibnizEq (c1,c2) ->
- (mkApp (hdcncl, [| c1; t|]), mkApp (hdcncl, [| t; c2 |]))
+let prove_transitivity hdcncl eq_kind t =
+ begin match eq_kind with
+ | MonomorphicLeibnizEq (c1,c2) ->
+ Goal.return (mkApp (hdcncl, [| c1; t|]), mkApp (hdcncl, [| t; c2 |]))
| PolymorphicLeibnizEq (typ,c1,c2) ->
- (mkApp (hdcncl, [| typ; c1; t |]), mkApp (hdcncl, [| typ; t; c2 |]))
+ Goal.return (mkApp (hdcncl, [| typ; c1; t |]), mkApp (hdcncl, [| typ; t; c2 |]))
| HeterogenousEq (typ1,c1,typ2,c2) ->
- let typt = pf_type_of gl t in
- (mkApp(hdcncl, [| typ1; c1; typt ;t |]),
- mkApp(hdcncl, [| typt; t; typ2; c2 |])) in
- tclTHENFIRST (cut eq2)
- (tclTHENFIRST (cut eq1)
- (tclTHENLIST
- [ tclDO 2 intro;
- onLastHyp simplest_case;
- assumption ])) gl
-
-let transitivity_red allowred t gl =
+ Tacmach.New.pf_apply Typing.type_of >- fun type_of ->
+ let typt = type_of t in
+ Goal.return
+ (mkApp(hdcncl, [| typ1; c1; typt ;t |]),
+ mkApp(hdcncl, [| typt; t; typ2; c2 |]))
+ end >>- fun (eq1,eq2) ->
+ Tacticals.New.tclTHENFIRST (Proofview.V82.tactic (cut eq2))
+ (Tacticals.New.tclTHENFIRST (Proofview.V82.tactic (cut eq1))
+ (Tacticals.New.tclTHENLIST
+ [ Tacticals.New.tclDO 2 intro;
+ Tacticals.New.onLastHyp simplest_case;
+ Proofview.V82.tactic assumption ]))
+
+let transitivity_red allowred t =
(* PL: usual transitivity don't perform any reduction when searching
for an equality, but we may need to do some when called back from
inside setoid_reflexivity (see Optimize cases in setoid_replace.ml). *)
let concl =
- if not allowred then pf_concl gl else pf_whd_betadeltaiota gl (pf_concl gl)
+ if not allowred then
+ Goal.concl
+ else
+ Goal.concl >- fun c ->
+ Tacmach.New.pf_apply (fun env sigma -> whd_betadeltaiota env sigma c)
in
- match match_with_equation concl with
+ concl >>- fun concl ->
+ match_with_equation concl >= fun with_eqn ->
+ match with_eqn with
| Some eq_data,_,_ ->
- tclTHEN
- (convert_concl_no_check concl DEFAULTcast)
- (match t with
- | None -> eapply eq_data.trans
- | Some t -> apply_list [eq_data.trans;t]) gl
+ Proofview.V82.tactic begin
+ tclTHEN
+ (convert_concl_no_check concl DEFAULTcast)
+ (match t with
+ | None -> eapply eq_data.trans
+ | Some t -> apply_list [eq_data.trans;t])
+ end
| None,eq,eq_kind ->
match t with
- | None -> error "etransitivity not supported for this relation."
- | Some t -> prove_transitivity eq eq_kind t gl
-
-let transitivity_gen t gl =
- try transitivity_red false t gl
- with NoEquationFound -> Hook.get forward_setoid_transitivity t gl
+ | None -> Proofview.tclZERO (Errors.UserError ("",str"etransitivity not supported for this relation."))
+ | Some t -> prove_transitivity eq eq_kind t
+
+let transitivity_gen t =
+ Proofview.tclORELSE
+ (transitivity_red false t)
+ begin function
+ | NoEquationFound -> Hook.get forward_setoid_transitivity t
+ | e -> Proofview.tclZERO e
+ end
let etransitivity = transitivity_gen None
let transitivity t = transitivity_gen (Some t)
-let intros_transitivity n = tclTHEN intros (transitivity_gen n)
+let intros_transitivity n = Tacticals.New.tclTHEN intros (transitivity_gen n)
(* tactical to save as name a subproof such that the generalisation of
the current goal, abstracted with respect to the local signature,
@@ -3590,7 +3725,7 @@ let abstract_subproof id tac gl =
with Uninstantiated_evar _ ->
error "\"abstract\" cannot handle existentials." in
let const = Pfedit.build_constant_by_tactic id secsign concl
- (tclCOMPLETE (tclTHEN (tclDO (List.length sign) intro) tac)) in
+ (Tacticals.New.tclCOMPLETE (Tacticals.New.tclTHEN (Tacticals.New.tclDO (List.length sign) intro) tac)) in
let cd = Entries.DefinitionEntry const in
let decl = (cd, IsProof Lemma) in
(** ppedrot: seems legit to have abstracted subproofs as local*)
@@ -3612,10 +3747,10 @@ let tclABSTRACT name_op tac gl =
abstract_subproof s tac gl
-let admit_as_an_axiom gl =
- let gl = simplest_case (Coqlib.build_coq_proof_admitted ()) gl in
- Pp.feedback Interface.AddedAxiom;
- gl
+let admit_as_an_axiom =
+ (* spiwack: admit temporarily won't report an unsafe status *)
+ Proofview.tclUNIT () >= fun () -> (* delay for Coqlib.build_coq_proof_admitted *)
+ simplest_case (Coqlib.build_coq_proof_admitted ())
let unify ?(state=full_transparent_state) x y gl =
try
@@ -3633,3 +3768,11 @@ let emit_side_effects eff gl =
prerr_endline ("emitting: " ^ Declareops.string_of_side_effect e))
eff;
{ it = [gl.it] ; sigma = Evd.emit_side_effects eff gl.sigma; }
+
+(** Tacticals defined directly in term of Proofview *)
+module New = struct
+ open Proofview
+
+ let exact_proof c = Proofview.V82.tactic (exact_proof c)
+
+end
diff --git a/tactics/tactics.mli b/tactics/tactics.mli
index 7acd803bd..dff606fe1 100644
--- a/tactics/tactics.mli
+++ b/tactics/tactics.mli
@@ -61,49 +61,49 @@ val fresh_id_in_env : Id.t list -> Id.t -> env -> Id.t
val fresh_id : Id.t list -> Id.t -> goal sigma -> Id.t
val find_intro_names : rel_context -> goal sigma -> Id.t list
-val intro : tactic
-val introf : tactic
-val intro_move : Id.t option -> Id.t move_location -> tactic
+val intro : unit Proofview.tactic
+val introf : unit Proofview.tactic
+val intro_move : Id.t option -> Id.t move_location -> unit Proofview.tactic
(** [intro_avoiding idl] acts as intro but prevents the new Id.t
to belong to [idl] *)
-val intro_avoiding : Id.t list -> tactic
+val intro_avoiding : Id.t list -> unit Proofview.tactic
val intro_replacing : Id.t -> tactic
-val intro_using : Id.t -> tactic
-val intro_mustbe_force : Id.t -> tactic
-val intro_then : (Id.t -> tactic) -> tactic
-val intros_using : Id.t list -> tactic
+val intro_using : Id.t -> unit Proofview.tactic
+val intro_mustbe_force : Id.t -> unit Proofview.tactic
+val intro_then : (Id.t -> unit Proofview.tactic) -> unit Proofview.tactic
+val intros_using : Id.t list -> unit Proofview.tactic
val intro_erasing : Id.t -> tactic
-val intros_replacing : Id.t list -> tactic
+val intros_replacing : Id.t list -> unit Proofview.tactic
-val intros : tactic
+val intros : unit Proofview.tactic
(** [depth_of_quantified_hypothesis b h g] returns the index of [h] in
the conclusion of goal [g], up to head-reduction if [b] is [true] *)
val depth_of_quantified_hypothesis :
bool -> quantified_hypothesis -> goal sigma -> int
-val intros_until_n_wored : int -> tactic
-val intros_until : quantified_hypothesis -> tactic
+val intros_until_n_wored : int -> unit Proofview.tactic
+val intros_until : quantified_hypothesis -> unit Proofview.tactic
-val intros_clearing : bool list -> tactic
+val intros_clearing : bool list -> unit Proofview.tactic
-(** Assuming a tactic [tac] depending on an hypothesis identifier,
+(** Assuming a tactic [tac] depending on an hypothesis Id.t,
[try_intros_until tac arg] first assumes that arg denotes a
quantified hypothesis (denoted by name or by index) and try to
introduce it in context before to apply [tac], otherwise assume the
hypothesis is already in context and directly apply [tac] *)
val try_intros_until :
- (Id.t -> tactic) -> quantified_hypothesis -> tactic
+ (Id.t -> unit Proofview.tactic) -> quantified_hypothesis -> unit Proofview.tactic
(** Apply a tactic on a quantified hypothesis, an hypothesis in context
or a term with bindings *)
val onInductionArg :
- (constr with_bindings -> tactic) ->
- constr with_bindings induction_arg -> tactic
+ (constr with_bindings -> unit Proofview.tactic) ->
+ constr with_bindings induction_arg -> unit Proofview.tactic
(** Complete intro_patterns up to some length; fails if more patterns
than required *)
@@ -113,10 +113,10 @@ val adjust_intro_patterns : int -> intro_pattern_expr located list ->
(** {6 Introduction tactics with eliminations. } *)
-val intro_pattern : Id.t move_location -> intro_pattern_expr -> tactic
-val intro_patterns : intro_pattern_expr located list -> tactic
+val intro_pattern : Id.t move_location -> intro_pattern_expr -> unit Proofview.tactic
+val intro_patterns : intro_pattern_expr located list -> unit Proofview.tactic
val intros_pattern :
- Id.t move_location -> intro_pattern_expr located list -> tactic
+ Id.t move_location -> intro_pattern_expr located list -> unit Proofview.tactic
(** {6 Exact tactics. } *)
@@ -197,9 +197,9 @@ val cut_and_apply : constr -> tactic
val apply_in :
advanced_flag -> evars_flag -> Id.t ->
constr with_bindings located list ->
- intro_pattern_expr located option -> tactic
+ intro_pattern_expr located option -> unit Proofview.tactic
-val simple_apply_in : Id.t -> constr -> tactic
+val simple_apply_in : Id.t -> constr -> unit Proofview.tactic
(** {6 Elimination tactics. } *)
@@ -274,30 +274,30 @@ val general_elim : evars_flag ->
val general_elim_in : evars_flag -> Id.t ->
constr with_bindings -> eliminator -> tactic
-val default_elim : evars_flag -> constr with_bindings -> tactic
-val simplest_elim : constr -> tactic
+val default_elim : evars_flag -> constr with_bindings -> unit Proofview.tactic
+val simplest_elim : constr -> unit Proofview.tactic
val elim :
- evars_flag -> constr with_bindings -> constr with_bindings option -> tactic
+ evars_flag -> constr with_bindings -> constr with_bindings option -> unit Proofview.tactic
-val simple_induct : quantified_hypothesis -> tactic
+val simple_induct : quantified_hypothesis -> unit Proofview.tactic
val new_induct : evars_flag ->
(evar_map * constr with_bindings) induction_arg list ->
constr with_bindings option ->
intro_pattern_expr located option * intro_pattern_expr located option ->
- clause option -> tactic
+ clause option -> unit Proofview.tactic
(** {6 Case analysis tactics. } *)
-val general_case_analysis : evars_flag -> constr with_bindings -> tactic
-val simplest_case : constr -> tactic
+val general_case_analysis : evars_flag -> constr with_bindings -> unit Proofview.tactic
+val simplest_case : constr -> unit Proofview.tactic
-val simple_destruct : quantified_hypothesis -> tactic
+val simple_destruct : quantified_hypothesis -> unit Proofview.tactic
val new_destruct : evars_flag ->
(evar_map * constr with_bindings) induction_arg list ->
constr with_bindings option ->
intro_pattern_expr located option * intro_pattern_expr located option ->
- clause option -> tactic
+ clause option -> unit Proofview.tactic
(** {6 Generic case analysis / induction tactics. } *)
@@ -306,7 +306,7 @@ val induction_destruct : rec_flag -> evars_flag ->
(intro_pattern_expr located option * intro_pattern_expr located option))
list *
constr with_bindings option *
- clause option -> tactic
+ clause option -> unit Proofview.tactic
(** {6 Eliminations giving the type instead of the proof. } *)
@@ -315,68 +315,68 @@ val elim_type : constr -> tactic
(** {6 Some eliminations which are frequently used. } *)
-val impE : Id.t -> tactic
-val andE : Id.t -> tactic
-val orE : Id.t -> tactic
-val dImp : clause -> tactic
-val dAnd : clause -> tactic
-val dorE : bool -> clause ->tactic
+val impE : Id.t -> unit Proofview.tactic
+val andE : Id.t -> unit Proofview.tactic
+val orE : Id.t -> unit Proofview.tactic
+val dImp : clause -> unit Proofview.tactic
+val dAnd : clause -> unit Proofview.tactic
+val dorE : bool -> clause -> unit Proofview.tactic
(** {6 Introduction tactics. } *)
val constructor_tac : evars_flag -> int option -> int ->
- constr bindings -> tactic
-val any_constructor : evars_flag -> tactic option -> tactic
-val one_constructor : int -> constr bindings -> tactic
+ constr bindings -> unit Proofview.tactic
+val any_constructor : evars_flag -> unit Proofview.tactic option -> unit Proofview.tactic
+val one_constructor : int -> constr bindings -> unit Proofview.tactic
-val left : constr bindings -> tactic
-val right : constr bindings -> tactic
-val split : constr bindings -> tactic
+val left : constr bindings -> unit Proofview.tactic
+val right : constr bindings -> unit Proofview.tactic
+val split : constr bindings -> unit Proofview.tactic
-val left_with_bindings : evars_flag -> constr bindings -> tactic
-val right_with_bindings : evars_flag -> constr bindings -> tactic
-val split_with_bindings : evars_flag -> constr bindings list -> tactic
+val left_with_bindings : evars_flag -> constr bindings -> unit Proofview.tactic
+val right_with_bindings : evars_flag -> constr bindings -> unit Proofview.tactic
+val split_with_bindings : evars_flag -> constr bindings list -> unit Proofview.tactic
-val simplest_left : tactic
-val simplest_right : tactic
-val simplest_split : tactic
+val simplest_left : unit Proofview.tactic
+val simplest_right : unit Proofview.tactic
+val simplest_split : unit Proofview.tactic
(** {6 Logical connective tactics. } *)
-val setoid_reflexivity : tactic Hook.t
-val reflexivity_red : bool -> tactic
-val reflexivity : tactic
-val intros_reflexivity : tactic
+val setoid_reflexivity : unit Proofview.tactic Hook.t
+val reflexivity_red : bool -> unit Proofview.tactic
+val reflexivity : unit Proofview.tactic
+val intros_reflexivity : unit Proofview.tactic
-val setoid_symmetry : tactic Hook.t
-val symmetry_red : bool -> tactic
-val symmetry : tactic
-val setoid_symmetry_in : (Id.t -> tactic) Hook.t
-val symmetry_in : Id.t -> tactic
-val intros_symmetry : clause -> tactic
+val setoid_symmetry : unit Proofview.tactic Hook.t
+val symmetry_red : bool -> unit Proofview.tactic
+val symmetry : unit Proofview.tactic
+val setoid_symmetry_in : (Id.t -> unit Proofview.tactic) Hook.t
+val symmetry_in : Id.t -> unit Proofview.tactic
+val intros_symmetry : clause -> unit Proofview.tactic
-val setoid_transitivity : (constr option -> tactic) Hook.t
-val transitivity_red : bool -> constr option -> tactic
-val transitivity : constr -> tactic
-val etransitivity : tactic
-val intros_transitivity : constr option -> tactic
+val setoid_transitivity : (constr option -> unit Proofview.tactic) Hook.t
+val transitivity_red : bool -> constr option -> unit Proofview.tactic
+val transitivity : constr -> unit Proofview.tactic
+val etransitivity : unit Proofview.tactic
+val intros_transitivity : constr option -> unit Proofview.tactic
val cut : constr -> tactic
-val cut_intro : constr -> tactic
+val cut_intro : constr -> unit Proofview.tactic
val assert_replacing : Id.t -> types -> tactic -> tactic
val cut_replacing : Id.t -> types -> tactic -> tactic
val cut_in_parallel : constr list -> tactic
-val assert_as : bool -> intro_pattern_expr located option -> constr -> tactic
-val forward : tactic option -> intro_pattern_expr located option -> constr -> tactic
+val assert_as : bool -> intro_pattern_expr located option -> constr -> unit Proofview.tactic
+val forward : unit Proofview.tactic option -> intro_pattern_expr located option -> constr -> unit Proofview.tactic
val letin_tac : (bool * intro_pattern_expr located) option -> Name.t ->
- constr -> types option -> clause -> tactic
+ constr -> types option -> clause -> unit Proofview.tactic
val letin_pat_tac : (bool * intro_pattern_expr located) option -> Name.t ->
- evar_map * constr -> types option -> clause -> tactic
-val assert_tac : Name.t -> types -> tactic
-val assert_by : Name.t -> types -> tactic -> tactic
-val pose_proof : Name.t -> constr -> tactic
+ evar_map * constr -> types option -> clause -> unit Proofview.tactic
+val assert_tac : Name.t -> types -> unit Proofview.tactic
+val assert_by : Name.t -> types -> unit Proofview.tactic -> unit Proofview.tactic
+val pose_proof : Name.t -> constr -> unit Proofview.tactic
val generalize : constr list -> tactic
val generalize_gen : ((occurrences * constr) * Name.t) list -> tactic
@@ -385,23 +385,30 @@ val generalize_dep : ?with_let:bool (** Don't lose let bindings *) -> constr -
val unify : ?state:Names.transparent_state -> constr -> constr -> tactic
val resolve_classes : tactic
-val tclABSTRACT : Id.t option -> tactic -> tactic
+val tclABSTRACT : Id.t option -> unit Proofview.tactic -> tactic
-val admit_as_an_axiom : tactic
+val admit_as_an_axiom : unit Proofview.tactic
-val abstract_generalize : ?generalize_vars:bool -> ?force_dep:bool -> Id.t -> tactic
+val abstract_generalize : ?generalize_vars:bool -> ?force_dep:bool -> Id.t -> unit Proofview.tactic
val specialize_eqs : Id.t -> tactic
val general_multi_rewrite :
- (bool -> evars_flag -> constr with_bindings -> clause -> tactic) Hook.t
+ (bool -> evars_flag -> constr with_bindings -> clause -> unit Proofview.tactic) Hook.t
val subst_one :
- (bool -> Id.t -> Id.t * constr * bool -> tactic) Hook.t
+ (bool -> Id.t -> Id.t * constr * bool -> unit Proofview.tactic) Hook.t
val declare_intro_decomp_eq :
- ((int -> tactic) -> Coqlib.coq_eq_data * types *
+ ((int -> unit Proofview.tactic) -> Coqlib.coq_eq_data * types *
(types * constr * constr) ->
- clausenv -> tactic) -> unit
+ clausenv -> unit Proofview.tactic) -> unit
val emit_side_effects : Declareops.side_effects -> tactic
+
+(** Tacticals defined directly in term of Proofview *)
+module New : sig
+ open Proofview
+
+ val exact_proof : Constrexpr.constr_expr -> unit tactic
+end
diff --git a/tactics/tauto.ml4 b/tactics/tauto.ml4
index 587ea3311..05a08c825 100644
--- a/tactics/tauto.ml4
+++ b/tactics/tauto.ml4
@@ -293,27 +293,39 @@ let t_reduction_not_iff = tacticIn reduction_not_iff
let intuition_gen flags tac =
interp (tacticIn (tauto_intuit flags t_reduction_not_iff tac))
-let tauto_intuitionistic flags g =
- try intuition_gen flags <:tactic<fail>> g
- with
- Refiner.FailError _ | UserError _ ->
- errorlabstrm "tauto" (str "tauto failed.")
+let tauto_intuitionistic flags =
+ Proofview.tclORELSE
+ (intuition_gen flags <:tactic<fail>>)
+ begin function
+ | Refiner.FailError _ | UserError _ ->
+ Proofview.tclZERO (UserError ("tauto" , str "tauto failed."))
+ | e -> Proofview.tclZERO e
+ end
let coq_nnpp_path =
let dir = List.map Id.of_string ["Classical_Prop";"Logic";"Coq"] in
Libnames.make_path (DirPath.make dir) (Id.of_string "NNPP")
-let tauto_classical flags nnpp g =
- try tclTHEN (apply nnpp) (tauto_intuitionistic flags) g
- with UserError _ -> errorlabstrm "tauto" (str "Classical tauto failed.")
+let tauto_classical flags nnpp =
+ Proofview.tclORELSE
+ (Tacticals.New.tclTHEN (Proofview.V82.tactic (apply nnpp)) (tauto_intuitionistic flags))
+ begin function
+ | UserError _ -> Proofview.tclZERO (UserError ("tauto" , str "Classical tauto failed."))
+ | e -> Proofview.tclZERO e
+ end
-let tauto_gen flags g =
- try
- let nnpp = constr_of_global (Nametab.global_of_path coq_nnpp_path) in
+let tauto_gen flags =
+ (* spiwack: I use [tclBIND (tclUNIT ())] as a way to delay the effect
+ (in [constr_of_global]) to the application of the tactic. *)
+ Proofview.tclBIND
+ (Proofview.tclUNIT ())
+ begin fun () -> try
+ let nnpp = constr_of_global (Nametab.global_of_path coq_nnpp_path) in
(* try intuitionistic version first to avoid an axiom if possible *)
- tclORELSE (tauto_intuitionistic flags) (tauto_classical flags nnpp) g
- with Not_found ->
- tauto_intuitionistic flags g
+ Tacticals.New.tclORELSE (tauto_intuitionistic flags) (tauto_classical flags nnpp)
+ with Not_found ->
+ tauto_intuitionistic flags
+ end
let default_intuition_tac = <:tactic< auto with * >>