diff options
author | Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr> | 2014-06-23 17:48:13 +0200 |
---|---|---|
committer | Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr> | 2014-06-24 13:52:34 +0200 |
commit | 89dc208ff140ba6ecd7b2c931401f9c58fb2985e (patch) | |
tree | dbc00054387bce1ab4c7452ea13d0900afe7201a | |
parent | efa3add0c03b70ecda3890cc6c69e66850605e7d (diff) |
Clenvtac.res_pf is in the new tactic monad.
-rw-r--r-- | proofs/clenvtac.ml | 6 | ||||
-rw-r--r-- | proofs/clenvtac.mli | 2 | ||||
-rw-r--r-- | tactics/leminv.ml | 2 | ||||
-rw-r--r-- | tactics/tactics.ml | 6 |
4 files changed, 9 insertions, 7 deletions
diff --git a/proofs/clenvtac.ml b/proofs/clenvtac.ml index 941d789dd..e8af6ffbd 100644 --- a/proofs/clenvtac.ml +++ b/proofs/clenvtac.ml @@ -78,8 +78,10 @@ open Unification let dft = default_unify_flags -let res_pf clenv ?(with_evars=false) ?(flags=dft ()) gls = - clenv_refine with_evars (clenv_unique_resolver ~flags clenv gls) gls +let res_pf clenv ?(with_evars=false) ?(flags=dft ()) = + Proofview.V82.tactic begin fun gl -> + clenv_refine with_evars (clenv_unique_resolver ~flags clenv gl) gl + end (* [unifyTerms] et [unify] ne semble pas gérer les Meta, en particulier ne semblent pas vérifier que des instances différentes diff --git a/proofs/clenvtac.mli b/proofs/clenvtac.mli index 3cfe1f315..eea32019e 100644 --- a/proofs/clenvtac.mli +++ b/proofs/clenvtac.mli @@ -15,7 +15,7 @@ open Unification (** Tactics *) val unify : ?flags:unify_flags -> constr -> unit Proofview.tactic val clenv_refine : evars_flag -> ?with_classes:bool -> clausenv -> tactic -val res_pf : clausenv -> ?with_evars:evars_flag -> ?flags:unify_flags -> tactic +val res_pf : clausenv -> ?with_evars:evars_flag -> ?flags:unify_flags -> unit Proofview.tactic val clenv_pose_dependent_evars : evars_flag -> clausenv -> clausenv val clenv_value_cast_meta : clausenv -> constr diff --git a/tactics/leminv.ml b/tactics/leminv.ml index d4f521054..237ce38b8 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -253,7 +253,7 @@ let lemInv id c gls = try let clause = mk_clenv_type_of gls c in let clause = clenv_constrain_last_binding (mkVar id) clause in - Clenvtac.res_pf clause ~flags:(Unification.elim_flags ()) gls + Proofview.V82.of_tactic (Clenvtac.res_pf clause ~flags:(Unification.elim_flags ()) ~with_evars:false) gls with | NoSuchBinding -> errorlabstrm "" diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 1a3e84a1f..5931257d2 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -843,7 +843,7 @@ let elimination_clause_scheme with_evars ?(flags=elim_flags ()) i (elim, elimty, (str "The type of elimination clause is not well-formed.")) in let elimclause' = clenv_fchain ~flags indmv elimclause indclause in - Clenvtac.res_pf elimclause' ~with_evars:with_evars ~flags gl + Proofview.V82.of_tactic (Clenvtac.res_pf elimclause' ~with_evars:with_evars ~flags) gl (* * Elimination tactic with bindings and using an arbitrary @@ -1080,7 +1080,7 @@ let general_apply with_delta with_destruct with_evars (loc,(c,lbind)) gl0 = let n = nb_prod thm_ty - nprod in if n<0 then error "Applied theorem has not enough premisses."; let clause = pf_apply make_clenv_binding_apply gl (Some n) (c,thm_ty) lbind in - Clenvtac.res_pf clause ~with_evars:with_evars ~flags:flags gl + Proofview.V82.of_tactic (Clenvtac.res_pf clause ~with_evars:with_evars ~flags:flags) gl in try try_apply thm_ty0 concl_nprod with PretypeError _|RefinerError _|UserError _|Failure _ as exn -> @@ -3586,7 +3586,7 @@ let elim_scheme_type elim t gl = (* t is inductive, then CUMUL or CONV is irrelevant *) clenv_unify ~flags:(elim_flags ()) Reduction.CUMUL t (clenv_meta_type clause mv) clause in - Clenvtac.res_pf clause' ~flags:(elim_flags ()) gl + Proofview.V82.of_tactic (Clenvtac.res_pf clause' ~flags:(elim_flags ()) ~with_evars:false) gl | _ -> anomaly (Pp.str "elim_scheme_type") let elim_type t gl = |