From 0437339ccac602d692b5b8c071b77ac756805520 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 15 Jun 2017 16:41:09 +0200 Subject: Removing Proof_type from the API. Unluckily, this forces replacing a lot of code in plugins, because the API defined the type of goals and tactics in Proof_type, and by the no-alias rule, this was the only one. But Proof_type was already implicitly deprecated, so that the API should have relied on Tacmach instead. --- API/API.ml | 13 --- API/API.mli | 101 ++++++++++------------ plugins/cc/ccalgo.ml | 2 +- plugins/cc/ccalgo.mli | 2 +- plugins/funind/functional_principles_proofs.ml | 3 +- plugins/funind/functional_principles_proofs.mli | 4 +- plugins/funind/functional_principles_types.mli | 2 +- plugins/funind/indfun.ml | 4 +- plugins/funind/indfun.mli | 2 +- plugins/funind/indfun_common.mli | 4 +- plugins/funind/invfun.ml | 10 +-- plugins/funind/recdef.ml | 2 +- plugins/funind/recdef.mli | 4 +- plugins/rtauto/refl_tauto.mli | 6 +- plugins/ssr/ssrast.mli | 4 +- plugins/ssr/ssrbwd.mli | 4 +- plugins/ssr/ssrcommon.ml | 2 +- plugins/ssr/ssrcommon.mli | 106 ++++++++++++------------ plugins/ssr/ssrelim.mli | 16 ++-- plugins/ssr/ssrequality.mli | 12 +-- plugins/ssr/ssrfwd.mli | 6 +- plugins/ssr/ssripats.ml | 6 +- plugins/ssr/ssripats.mli | 16 ++-- plugins/ssr/ssrprinters.mli | 2 +- plugins/ssr/ssrtacticals.mli | 6 +- plugins/ssr/ssrview.ml | 2 +- plugins/ssrmatching/ssrmatching.mli | 2 +- 27 files changed, 158 insertions(+), 185 deletions(-) diff --git a/API/API.ml b/API/API.ml index 2b7bbd561..60703e1fa 100644 --- a/API/API.ml +++ b/API/API.ml @@ -200,16 +200,3 @@ module Entries = | ParameterEntry of parameter_entry | ProjectionEntry of projection_entry end - -(* NOTE: It does not make sense to replace the following "module expression" - simply with "module Proof_type = Proof_type" because - there is only "proofs/proof_type.mli"; - there is no "proofs/proof_type.ml" file *) -module Proof_type = - struct - type goal = Goal.goal - type tactic = goal Evd.sigma -> goal list Evd.sigma - type rule = Proof_type.prim_rule = - | Cut of bool * bool * Names.Id.t * Term.types - | Refine of Term.constr - end diff --git a/API/API.mli b/API/API.mli index 69278e7c9..9ed952dc2 100644 --- a/API/API.mli +++ b/API/API.mli @@ -3310,17 +3310,6 @@ sig val module_is_known : string -> bool end -(* All items in the Proof_type module are deprecated. *) -module Proof_type : -sig - type goal = Evar.t - type rule = Proof_type.prim_rule = - | Cut of bool * bool * Names.Id.t * Term.types - | Refine of Term.constr - - type tactic = goal Evd.sigma -> goal list Evd.sigma -end - module Redexpr : sig type red_expr = @@ -3332,61 +3321,61 @@ end module Tacmach : sig - type tactic = Proof_type.tactic - [@@ocaml.deprecated "alias for API.Proof_type.tactic"] + type tactic = Goal.goal Evd.sigma -> Goal.goal list Evd.sigma type 'a sigma = 'a Evd.sigma [@@ocaml.deprecated "alias of API.Evd.sigma"] val re_sig : 'a -> Evd.evar_map -> 'a Evd.sigma - val pf_reduction_of_red_expr : Proof_type.goal Evd.sigma -> Redexpr.red_expr -> EConstr.constr -> Evd.evar_map * EConstr.constr + val pf_reduction_of_red_expr : Goal.goal Evd.sigma -> Redexpr.red_expr -> EConstr.constr -> Evd.evar_map * EConstr.constr - val pf_unsafe_type_of : Proof_type.goal Evd.sigma -> EConstr.constr -> EConstr.types + val pf_unsafe_type_of : Goal.goal Evd.sigma -> EConstr.constr -> EConstr.types - val pf_get_new_id : Names.Id.t -> Proof_type.goal Evd.sigma -> Names.Id.t + val pf_get_new_id : Names.Id.t -> Goal.goal Evd.sigma -> Names.Id.t - val pf_env : Proof_type.goal Evd.sigma -> Environ.env + val pf_env : Goal.goal Evd.sigma -> Environ.env - val pf_concl : Proof_type.goal Evd.sigma -> EConstr.types + val pf_concl : Goal.goal Evd.sigma -> EConstr.types - val pf_apply : (Environ.env -> Evd.evar_map -> 'a) -> Proof_type.goal Evd.sigma -> 'a + val pf_apply : (Environ.env -> Evd.evar_map -> 'a) -> Goal.goal Evd.sigma -> 'a - val pf_get_hyp : Proof_type.goal Evd.sigma -> Names.Id.t -> EConstr.named_declaration - val pf_get_hyp_typ : Proof_type.goal Evd.sigma -> Names.Id.t -> EConstr.types - val project : Proof_type.goal Evd.sigma -> Evd.evar_map - val refine : EConstr.constr -> Proof_type.tactic - val pf_type_of : Proof_type.goal Evd.sigma -> EConstr.constr -> Evd.evar_map * EConstr.types + val pf_get_hyp : Goal.goal Evd.sigma -> Names.Id.t -> EConstr.named_declaration + val pf_get_hyp_typ : Goal.goal Evd.sigma -> Names.Id.t -> EConstr.types + val project : Goal.goal Evd.sigma -> Evd.evar_map + val refine : EConstr.constr -> tactic + val refine_no_check : EConstr.constr -> tactic + val pf_type_of : Goal.goal Evd.sigma -> EConstr.constr -> Evd.evar_map * EConstr.types - val pf_hyps : Proof_type.goal Evd.sigma -> EConstr.named_context + val pf_hyps : Goal.goal Evd.sigma -> EConstr.named_context - val pf_ids_of_hyps : Proof_type.goal Evd.sigma -> Names.Id.t list + val pf_ids_of_hyps : Goal.goal Evd.sigma -> Names.Id.t list - val pf_reduce_to_atomic_ind : Proof_type.goal Evd.sigma -> EConstr.types -> (Names.inductive * EConstr.EInstance.t) * EConstr.types + val pf_reduce_to_atomic_ind : Goal.goal Evd.sigma -> EConstr.types -> (Names.inductive * EConstr.EInstance.t) * EConstr.types - val pf_reduce_to_quantified_ind : Proof_type.goal Evd.sigma -> EConstr.types -> (Names.inductive * EConstr.EInstance.t) * EConstr.types + val pf_reduce_to_quantified_ind : Goal.goal Evd.sigma -> EConstr.types -> (Names.inductive * EConstr.EInstance.t) * EConstr.types val pf_eapply : (Environ.env -> Evd.evar_map -> 'a -> Evd.evar_map * 'b) -> Evar.t Evd.sigma -> 'a -> Evar.t Evd.sigma * 'b val pf_unfoldn : (Locus.occurrences * Names.evaluable_global_reference) list - -> Proof_type.goal Evd.sigma -> EConstr.constr -> EConstr.constr + -> Goal.goal Evd.sigma -> EConstr.constr -> EConstr.constr - val pf_reduce : (Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr) -> Proof_type.goal Evd.sigma -> EConstr.constr -> EConstr.constr + val pf_reduce : (Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr) -> Goal.goal Evd.sigma -> EConstr.constr -> EConstr.constr - val pf_conv_x : Proof_type.goal Evd.sigma -> EConstr.constr -> EConstr.constr -> bool + val pf_conv_x : Goal.goal Evd.sigma -> EConstr.constr -> EConstr.constr -> bool - val pf_is_matching : Proof_type.goal Evd.sigma -> Pattern.constr_pattern -> EConstr.constr -> bool + val pf_is_matching : Goal.goal Evd.sigma -> Pattern.constr_pattern -> EConstr.constr -> bool - val pf_hyps_types : Proof_type.goal Evd.sigma -> (Names.Id.t * EConstr.types) list + val pf_hyps_types : Goal.goal Evd.sigma -> (Names.Id.t * EConstr.types) list - val pr_gls : Proof_type.goal Evd.sigma -> Pp.std_ppcmds + val pr_gls : Goal.goal Evd.sigma -> Pp.std_ppcmds - val pf_nf_betaiota : Proof_type.goal Evd.sigma -> EConstr.constr -> EConstr.constr + val pf_nf_betaiota : Goal.goal Evd.sigma -> EConstr.constr -> EConstr.constr - val pf_last_hyp : Proof_type.goal Evd.sigma -> EConstr.named_declaration + val pf_last_hyp : Goal.goal Evd.sigma -> EConstr.named_declaration - val pf_nth_hyp_id : Proof_type.goal Evd.sigma -> int -> Names.Id.t + val pf_nth_hyp_id : Goal.goal Evd.sigma -> int -> Names.Id.t val sig_it : 'a Evd.sigma -> 'a @@ -3395,7 +3384,7 @@ sig val pf_apply : (Environ.env -> Evd.evar_map -> 'a) -> 'b Proofview.Goal.t -> 'a val project : 'a Proofview.Goal.t -> Evd.evar_map val pf_unsafe_type_of : 'a Proofview.Goal.t -> EConstr.constr -> EConstr.types - val of_old : (Proof_type.goal Evd.sigma -> 'a) -> [ `NF ] Proofview.Goal.t -> 'a + val of_old : (Goal.goal Evd.sigma -> 'a) -> [ `NF ] Proofview.Goal.t -> 'a val pf_env : 'a Proofview.Goal.t -> Environ.env val pf_ids_of_hyps : 'a Proofview.Goal.t -> Names.Id.t list @@ -3516,21 +3505,19 @@ sig val repackage : Evd.evar_map ref -> 'a -> 'a Evd.sigma - val refiner : Proof_type.rule -> Proof_type.tactic - - val tclSHOWHYPS : Proof_type.tactic -> Proof_type.tactic + val tclSHOWHYPS : Tacmach.tactic -> Tacmach.tactic exception FailError of int * Pp.std_ppcmds Lazy.t - val tclEVARS : Evd.evar_map -> Proof_type.tactic - val tclMAP : ('a -> Proof_type.tactic) -> 'a list -> Proof_type.tactic - val tclREPEAT : Proof_type.tactic -> Proof_type.tactic - val tclORELSE : Proof_type.tactic -> Proof_type.tactic -> Proof_type.tactic - val tclFAIL : int -> Pp.std_ppcmds -> Proof_type.tactic - val tclIDTAC : Proof_type.tactic - val tclTHEN : Proof_type.tactic -> Proof_type.tactic -> Proof_type.tactic - val tclTHENLIST : Proof_type.tactic list -> Proof_type.tactic - val tclTRY : Proof_type.tactic -> Proof_type.tactic - val tclAT_LEAST_ONCE : Proof_type.tactic -> Proof_type.tactic + val tclEVARS : Evd.evar_map -> Tacmach.tactic + val tclMAP : ('a -> Tacmach.tactic) -> 'a list -> Tacmach.tactic + val tclREPEAT : Tacmach.tactic -> Tacmach.tactic + val tclORELSE : Tacmach.tactic -> Tacmach.tactic -> Tacmach.tactic + val tclFAIL : int -> Pp.std_ppcmds -> Tacmach.tactic + val tclIDTAC : Tacmach.tactic + val tclTHEN : Tacmach.tactic -> Tacmach.tactic -> Tacmach.tactic + val tclTHENLIST : Tacmach.tactic list -> Tacmach.tactic + val tclTRY : Tacmach.tactic -> Tacmach.tactic + val tclAT_LEAST_ONCE : Tacmach.tactic -> Tacmach.tactic end module Termops : @@ -3668,7 +3655,7 @@ module Printer : sig val pr_named_context : Environ.env -> Evd.evar_map -> Context.Named.t -> Pp.std_ppcmds val pr_rel_context : Environ.env -> Evd.evar_map -> Context.Rel.t -> Pp.std_ppcmds - val pr_goal : Proof_type.goal Evd.sigma -> Pp.std_ppcmds + val pr_goal : Goal.goal Evd.sigma -> Pp.std_ppcmds val pr_constr_env : Prelude.env -> Prelude.evar_map -> Term.constr -> Pp.std_ppcmds val pr_lconstr_env : Prelude.env -> Prelude.evar_map -> Term.constr -> Pp.std_ppcmds @@ -4124,7 +4111,7 @@ end module Tacticals : sig - open Proof_type + open Tacmach val tclORELSE : tactic -> tactic -> tactic val tclDO : int -> tactic -> tactic val tclIDTAC : tactic @@ -4132,7 +4119,7 @@ sig val tclTHEN : tactic -> tactic -> tactic val tclTHENLIST : tactic list -> tactic val pf_constr_of_global : - Globnames.global_reference -> (EConstr.constr -> Proof_type.tactic) -> Proof_type.tactic + Globnames.global_reference -> (EConstr.constr -> Tacmach.tactic) -> Tacmach.tactic val tclMAP : ('a -> tactic) -> 'a list -> tactic val tclTRY : tactic -> tactic val tclCOMPLETE : tactic -> tactic @@ -4153,13 +4140,13 @@ sig val tclTHENSEQ : tactic list -> tactic [@@ocaml.deprecated "alias of API.Tacticals.tclTHENLIST"] - val nLastDecls : int -> Proof_type.goal Evd.sigma -> EConstr.named_context + val nLastDecls : int -> Goal.goal Evd.sigma -> EConstr.named_context val tclTHEN_i : tactic -> (int -> tactic) -> tactic val tclPROGRESS : tactic -> tactic - val elimination_sort_of_goal : Proof_type.goal Evd.sigma -> Sorts.family + val elimination_sort_of_goal : Goal.goal Evd.sigma -> Sorts.family module New : sig @@ -4539,7 +4526,7 @@ sig val autounfold_tac : Hints.hint_db_name list option -> Locus.clause -> unit Proofview.tactic val autounfold_one : Hints.hint_db_name list -> Locus.hyp_location option -> unit Proofview.tactic val eauto_with_bases : - ?debug:Hints.debug -> bool * int -> Tactypes.delayed_open_constr list -> Hints.hint_db list -> Proof_type.tactic + ?debug:Hints.debug -> bool * int -> Tactypes.delayed_open_constr list -> Hints.hint_db list -> Tacmach.tactic end module Class_tactics : diff --git a/plugins/cc/ccalgo.ml b/plugins/cc/ccalgo.ml index 5c7cad7ff..39fb8feeb 100644 --- a/plugins/cc/ccalgo.ml +++ b/plugins/cc/ccalgo.ml @@ -270,7 +270,7 @@ type state = mutable rew_depth:int; mutable changed:bool; by_type: Int.Set.t Typehash.t; - mutable gls:Proof_type.goal Evd.sigma} + mutable gls:Goal.goal Evd.sigma} let dummy_node = { diff --git a/plugins/cc/ccalgo.mli b/plugins/cc/ccalgo.mli index 505029992..51e2301fe 100644 --- a/plugins/cc/ccalgo.mli +++ b/plugins/cc/ccalgo.mli @@ -129,7 +129,7 @@ val axioms : forest -> (term * term) Constrhash.t val epsilons : forest -> pa_constructor list -val empty : int -> Proof_type.goal Evd.sigma -> state +val empty : int -> Goal.goal Evd.sigma -> state val add_term : state -> term -> int diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index ef894b239..ba46f78aa 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -10,7 +10,6 @@ open Names open Pp open Tacmach open Termops -open Proof_type open Tacticals open Tactics open Indfun_common @@ -106,7 +105,7 @@ let make_refl_eq constructor type_of_t t = type pte_info = { - proving_tac : (Id.t list -> Proof_type.tactic); + proving_tac : (Id.t list -> Tacmach.tactic); is_valid : constr -> bool } diff --git a/plugins/funind/functional_principles_proofs.mli b/plugins/funind/functional_principles_proofs.mli index 5bb288678..d03fc475e 100644 --- a/plugins/funind/functional_principles_proofs.mli +++ b/plugins/funind/functional_principles_proofs.mli @@ -4,7 +4,7 @@ open Names val prove_princ_for_struct : Evd.evar_map ref -> bool -> - int -> Constant.t array -> EConstr.constr array -> int -> Proof_type.tactic + int -> Constant.t array -> EConstr.constr array -> int -> Tacmach.tactic val prove_principle_for_gen : @@ -14,7 +14,7 @@ val prove_principle_for_gen : int -> (* the number of recursive argument *) EConstr.types -> (* the type of the recursive argument *) EConstr.constr -> (* the wf relation used to prove the function *) - Proof_type.tactic + Tacmach.tactic (* val is_pte : rel_declaration -> bool *) diff --git a/plugins/funind/functional_principles_types.mli b/plugins/funind/functional_principles_types.mli index bb2b2d918..e70ef2365 100644 --- a/plugins/funind/functional_principles_types.mli +++ b/plugins/funind/functional_principles_types.mli @@ -28,7 +28,7 @@ val generate_functional_principle : (* The tactic to use to make the proof w.r the number of params *) - (EConstr.constr array -> int -> Proof_type.tactic) -> + (EConstr.constr array -> int -> Tacmach.tactic) -> unit val compute_new_princ_type_from_rel : constr array -> Sorts.t array -> diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index d12aa7f42..ad04e430c 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -343,7 +343,7 @@ let error_error names e = let generate_principle (evd:Evd.evar_map ref) pconstants on_error is_general do_built (fix_rec_l:(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list) recdefs interactive_proof (continue_proof : int -> Names.Constant.t array -> EConstr.constr array -> int -> - Proof_type.tactic) : unit = + Tacmach.tactic) : unit = let names = List.map (function (((_, name),_),_,_,_,_),_ -> name) fix_rec_l in let fun_bodies = List.map2 prepare_body fix_rec_l recdefs in let funs_args = List.map fst fun_bodies in @@ -446,7 +446,7 @@ let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexp let generate_correction_proof_wf f_ref tcc_lemma_ref is_mes functional_ref eq_ref rec_arg_num rec_arg_type nb_args relation - (_: int) (_:Names.Constant.t array) (_:EConstr.constr array) (_:int) : Proof_type.tactic = + (_: int) (_:Names.Constant.t array) (_:EConstr.constr array) (_:int) : Tacmach.tactic = Functional_principles_proofs.prove_principle_for_gen (f_ref,functional_ref,eq_ref) tcc_lemma_ref is_mes rec_arg_num rec_arg_type relation diff --git a/plugins/funind/indfun.mli b/plugins/funind/indfun.mli index 33420d813..fc7da6a33 100644 --- a/plugins/funind/indfun.mli +++ b/plugins/funind/indfun.mli @@ -16,7 +16,7 @@ val functional_induction : EConstr.constr -> (EConstr.constr * EConstr.constr bindings) option -> Tacexpr.or_and_intro_pattern option -> - Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma + Goal.goal Evd.sigma -> Goal.goal list Evd.sigma val make_graph : Globnames.global_reference -> unit diff --git a/plugins/funind/indfun_common.mli b/plugins/funind/indfun_common.mli index 6b40c9171..f7a9cedd7 100644 --- a/plugins/funind/indfun_common.mli +++ b/plugins/funind/indfun_common.mli @@ -105,7 +105,7 @@ exception ToShow of exn val is_strict_tcc : unit -> bool -val h_intros: Names.Id.t list -> Proof_type.tactic +val h_intros: Names.Id.t list -> Tacmach.tactic val h_id : Names.Id.t val hrec_id : Names.Id.t val acc_inv_id : EConstr.constr Util.delayed @@ -114,7 +114,7 @@ val well_founded_ltof : EConstr.constr Util.delayed val acc_rel : EConstr.constr Util.delayed val well_founded : EConstr.constr Util.delayed val evaluable_of_global_reference : Globnames.global_reference -> Names.evaluable_global_reference -val list_rewrite : bool -> (EConstr.constr*bool) list -> Proof_type.tactic +val list_rewrite : bool -> (EConstr.constr*bool) list -> Tacmach.tactic val decompose_lam_n : Evd.evar_map -> int -> EConstr.t -> (Names.Name.t * EConstr.t) list * EConstr.t diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index ebdb490e3..94ef2590c 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -218,7 +218,7 @@ let rec generate_fresh_id x avoid i = \end{enumerate} *) -let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes lemmas_types_infos i : Proof_type.tactic = +let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes lemmas_types_infos i : Tacmach.tactic = fun g -> (* first of all we recreate the lemmas types to be used as predicates of the induction principle that is~: @@ -468,7 +468,7 @@ let tauto = let rec intros_with_rewrite g = observe_tac "intros_with_rewrite" intros_with_rewrite_aux g -and intros_with_rewrite_aux : Proof_type.tactic = +and intros_with_rewrite_aux : Tacmach.tactic = fun g -> let eq_ind = make_eq () in let sigma = project g in @@ -629,7 +629,7 @@ let rec reflexivity_with_destruct_cases g = *) -let prove_fun_complete funcs graphs schemes lemmas_types_infos i : Proof_type.tactic = +let prove_fun_complete funcs graphs schemes lemmas_types_infos i : Tacmach.tactic = fun g -> (* We compute the types of the different mutually recursive lemmas in $\zeta$ normal form @@ -673,7 +673,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : Proof_type.ta using [f_equation] if it is recursive (that is the graph is infinite or unfold if the graph is finite *) - let rewrite_tac j ids : Proof_type.tactic = + let rewrite_tac j ids : Tacmach.tactic = let graph_def = graphs.(j) in let infos = try find_Function_infos (fst (destConst (project g) funcs.(j))) @@ -953,7 +953,7 @@ let revert_graph kn post_tac hid g = \end{enumerate} *) -let functional_inversion kn hid fconst f_correct : Proof_type.tactic = +let functional_inversion kn hid fconst f_correct : Tacmach.tactic = fun g -> let old_ids = List.fold_right Id.Set.add (pf_ids_of_hyps g) Id.Set.empty in let sigma = project g in diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 3cd20a950..8e12b239e 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -30,7 +30,7 @@ open Nametab open Declare open Decl_kinds open Tacred -open Proof_type +open Goal open Pfedit open Glob_term open Pretyping diff --git a/plugins/funind/recdef.mli b/plugins/funind/recdef.mli index e1a072799..f3d5e7332 100644 --- a/plugins/funind/recdef.mli +++ b/plugins/funind/recdef.mli @@ -2,10 +2,10 @@ open API (* val evaluable_of_global_reference : Libnames.global_reference -> Names.evaluable_global_reference *) val tclUSER_if_not_mes : - Proof_type.tactic -> + Tacmach.tactic -> bool -> Names.Id.t list option -> - Proof_type.tactic + Tacmach.tactic val recursive_definition : bool -> Names.Id.t -> diff --git a/plugins/rtauto/refl_tauto.mli b/plugins/rtauto/refl_tauto.mli index ac260e51a..801fc4606 100644 --- a/plugins/rtauto/refl_tauto.mli +++ b/plugins/rtauto/refl_tauto.mli @@ -14,13 +14,13 @@ type atom_env= mutable env:(Term.constr*int) list} val make_form : atom_env -> - Proof_type.goal Evd.sigma -> EConstr.types -> Proof_search.form + Goal.goal Evd.sigma -> EConstr.types -> Proof_search.form val make_hyps : atom_env -> - Proof_type.goal Evd.sigma -> + Goal.goal Evd.sigma -> EConstr.types list -> EConstr.named_context -> (Names.Id.t * Proof_search.form) list -val rtauto_tac : Proof_type.tactic +val rtauto_tac : Tacmach.tactic diff --git a/plugins/ssr/ssrast.mli b/plugins/ssr/ssrast.mli index 0f4b86d10..94eaa1d6a 100644 --- a/plugins/ssr/ssrast.mli +++ b/plugins/ssr/ssrast.mli @@ -145,6 +145,6 @@ type 'a ssrseqarg = ssrindex * ('a ssrhint * 'a option) (* OOP : these are general shortcuts *) type gist = Tacintern.glob_sign type ist = Tacinterp.interp_sign -type goal = Proof_type.goal +type goal = Goal.goal type 'a sigma = 'a Evd.sigma -type v82tac = Proof_type.tactic +type v82tac = Tacmach.tactic diff --git a/plugins/ssr/ssrbwd.mli b/plugins/ssr/ssrbwd.mli index b0e98bdb4..20a1263d2 100644 --- a/plugins/ssr/ssrbwd.mli +++ b/plugins/ssr/ssrbwd.mli @@ -10,7 +10,7 @@ open API -val apply_top_tac : Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma +val apply_top_tac : Goal.goal Evd.sigma -> Goal.goal list Evd.sigma val inner_ssrapplytac : Ssrast.ssrterm list -> @@ -19,4 +19,4 @@ val inner_ssrapplytac : list list -> Ssrast.ssrhyps -> Ssrast.ist -> - Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma + Goal.goal Evd.sigma -> Goal.goal list Evd.sigma diff --git a/plugins/ssr/ssrcommon.ml b/plugins/ssr/ssrcommon.ml index ded9fa66b..411ce6853 100644 --- a/plugins/ssr/ssrcommon.ml +++ b/plugins/ssr/ssrcommon.ml @@ -960,7 +960,7 @@ let applyn ~with_evars ?beta ?(with_shelve=false) n t gl = | _ -> assert false in loop sigma t [] n in pp(lazy(str"Refiner.refiner " ++ Printer.pr_econstr t)); - Tacmach.refine_no_check (EConstr.Unsafe.to_constr t) gl + Tacmach.refine_no_check t gl let refine_with ?(first_goes_last=false) ?beta ?(with_evars=true) oc gl = let rec mkRels = function 1 -> [] | n -> mkRel n :: mkRels (n-1) in diff --git a/plugins/ssr/ssrcommon.mli b/plugins/ssr/ssrcommon.mli index 7a4b47a46..f61168576 100644 --- a/plugins/ssr/ssrcommon.mli +++ b/plugins/ssr/ssrcommon.mli @@ -9,9 +9,9 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) open API +open Tacmach open Names open Environ -open Proof_type open Evd open Constrexpr open Ssrast @@ -122,11 +122,11 @@ val intern_term : ssrterm -> Glob_term.glob_constr val pf_intern_term : - Tacinterp.interp_sign -> Proof_type.goal Evd.sigma -> + Tacinterp.interp_sign -> Goal.goal Evd.sigma -> ssrterm -> Glob_term.glob_constr val interp_term : - Tacinterp.interp_sign -> Proof_type.goal Evd.sigma -> + Tacinterp.interp_sign -> Goal.goal Evd.sigma -> ssrterm -> evar_map * EConstr.t val interp_wit : @@ -136,28 +136,28 @@ val interp_hyp : ist -> goal sigma -> ssrhyp -> evar_map * ssrhyp val interp_hyps : ist -> goal sigma -> ssrhyps -> evar_map * ssrhyps val interp_refine : - Tacinterp.interp_sign -> Proof_type.goal Evd.sigma -> + Tacinterp.interp_sign -> Goal.goal Evd.sigma -> Glob_term.glob_constr -> evar_map * (evar_map * EConstr.constr) val interp_open_constr : - Tacinterp.interp_sign -> Proof_type.goal Evd.sigma -> + Tacinterp.interp_sign -> Goal.goal Evd.sigma -> Tacexpr.glob_constr_and_expr -> evar_map * (evar_map * EConstr.t) val pf_e_type_of : - Proof_type.goal Evd.sigma -> - EConstr.constr -> Proof_type.goal Evd.sigma * EConstr.types + Goal.goal Evd.sigma -> + EConstr.constr -> Goal.goal Evd.sigma * EConstr.types val splay_open_constr : - Proof_type.goal Evd.sigma -> + Goal.goal Evd.sigma -> evar_map * EConstr.t -> (Names.Name.t * EConstr.t) list * EConstr.t -val isAppInd : Proof_type.goal Evd.sigma -> EConstr.types -> bool +val isAppInd : Goal.goal Evd.sigma -> EConstr.types -> bool val interp_view_nbimps : Tacinterp.interp_sign -> - Proof_type.goal Evd.sigma -> Glob_term.glob_constr -> int + Goal.goal Evd.sigma -> Glob_term.glob_constr -> int val interp_nbargs : Tacinterp.interp_sign -> - Proof_type.goal Evd.sigma -> Glob_term.glob_constr -> int + Goal.goal Evd.sigma -> Glob_term.glob_constr -> int val mk_term : ssrtermkind -> 'b -> ssrtermkind * (Glob_term.glob_constr * 'b option) @@ -169,20 +169,20 @@ val mk_internal_id : string -> Id.t val mk_tagged_id : string -> int -> Id.t val mk_evar_name : int -> Name.t val ssr_anon_hyp : string -val pf_type_id : Proof_type.goal Evd.sigma -> EConstr.types -> Id.t +val pf_type_id : Goal.goal Evd.sigma -> EConstr.types -> Id.t val pf_abs_evars : - Proof_type.goal Evd.sigma -> + Goal.goal Evd.sigma -> evar_map * EConstr.t -> int * EConstr.t * Evar.t list * UState.t val pf_abs_evars2 : (* ssr2 *) - Proof_type.goal Evd.sigma -> Evar.t list -> + Goal.goal Evd.sigma -> Evar.t list -> evar_map * EConstr.t -> int * EConstr.t * Evar.t list * UState.t val pf_abs_cterm : - Proof_type.goal Evd.sigma -> int -> EConstr.t -> EConstr.t + Goal.goal Evd.sigma -> int -> EConstr.t -> EConstr.t val pf_merge_uc : UState.t -> 'a Evd.sigma -> 'a Evd.sigma @@ -190,21 +190,21 @@ val pf_merge_uc_of : evar_map -> 'a Evd.sigma -> 'a Evd.sigma val constr_name : evar_map -> EConstr.t -> Name.t val pf_type_of : - Proof_type.goal Evd.sigma -> - Term.constr -> Proof_type.goal Evd.sigma * Term.types + Goal.goal Evd.sigma -> + Term.constr -> Goal.goal Evd.sigma * Term.types val pfe_type_of : - Proof_type.goal Evd.sigma -> - EConstr.t -> Proof_type.goal Evd.sigma * EConstr.types + Goal.goal Evd.sigma -> + EConstr.t -> Goal.goal Evd.sigma * EConstr.types val pf_abs_prod : Name.t -> - Proof_type.goal Evd.sigma -> + Goal.goal Evd.sigma -> EConstr.t -> - EConstr.t -> Proof_type.goal Evd.sigma * EConstr.types + EConstr.t -> Goal.goal Evd.sigma * EConstr.types val pf_mkprod : - Proof_type.goal Evd.sigma -> + Goal.goal Evd.sigma -> EConstr.t -> ?name:Name.t -> - EConstr.t -> Proof_type.goal Evd.sigma * EConstr.types + EConstr.t -> Goal.goal Evd.sigma * EConstr.types val mkSsrRRef : string -> Glob_term.glob_constr * 'a option val mkSsrRef : string -> Globnames.global_reference @@ -213,15 +213,15 @@ val mkSsrConst : env -> evar_map -> evar_map * EConstr.t val pf_mkSsrConst : string -> - Proof_type.goal Evd.sigma -> - EConstr.t * Proof_type.goal Evd.sigma + Goal.goal Evd.sigma -> + EConstr.t * Goal.goal Evd.sigma val new_wild_id : tac_ctx -> Names.Id.t * tac_ctx val pf_fresh_global : Globnames.global_reference -> - Proof_type.goal Evd.sigma -> - Term.constr * Proof_type.goal Evd.sigma + Goal.goal Evd.sigma -> + Term.constr * Goal.goal Evd.sigma val is_discharged_id : Id.t -> bool val mk_discharged_id : Id.t -> Id.t @@ -230,15 +230,15 @@ val has_discharged_tag : string -> bool val ssrqid : string -> Libnames.qualid val new_tmp_id : tac_ctx -> (Names.Id.t * Name.t ref) * tac_ctx -val mk_anon_id : string -> Proof_type.goal Evd.sigma -> Id.t +val mk_anon_id : string -> Goal.goal Evd.sigma -> Id.t val pf_abs_evars_pirrel : - Proof_type.goal Evd.sigma -> + Goal.goal Evd.sigma -> evar_map * Term.constr -> int * Term.constr -val pf_nbargs : Proof_type.goal Evd.sigma -> EConstr.t -> int +val pf_nbargs : Goal.goal Evd.sigma -> EConstr.t -> int val gen_tmp_ids : ?ist:Geninterp.interp_sign -> - (Proof_type.goal * tac_ctx) Evd.sigma -> - (Proof_type.goal * tac_ctx) list Evd.sigma + (Goal.goal * tac_ctx) Evd.sigma -> + (Goal.goal * tac_ctx) list Evd.sigma val ssrevaltac : Tacinterp.interp_sign -> Tacinterp.Value.t -> Proofview.V82.tac @@ -258,23 +258,23 @@ val ssrautoprop_tac : val mkProt : EConstr.t -> EConstr.t -> - Proof_type.goal Evd.sigma -> - EConstr.t * Proof_type.goal Evd.sigma + Goal.goal Evd.sigma -> + EConstr.t * Goal.goal Evd.sigma val mkEtaApp : EConstr.t -> int -> int -> EConstr.t val mkRefl : EConstr.t -> EConstr.t -> - Proof_type.goal Evd.sigma -> EConstr.t * Proof_type.goal Evd.sigma + Goal.goal Evd.sigma -> EConstr.t * Goal.goal Evd.sigma val discharge_hyp : Id.t * (Id.t * string) -> - Proof_type.goal Evd.sigma -> Evar.t list Evd.sigma + Goal.goal Evd.sigma -> Evar.t list Evd.sigma val clear_wilds_and_tmp_and_delayed_ids : - (Proof_type.goal * tac_ctx) Evd.sigma -> - (Proof_type.goal * tac_ctx) list Evd.sigma + (Goal.goal * tac_ctx) Evd.sigma -> + (Goal.goal * tac_ctx) list Evd.sigma val view_error : string -> ssrterm -> 'a @@ -284,14 +284,14 @@ val top_id : Id.t val pf_abs_ssrterm : ?resolve_typeclasses:bool -> ist -> - Proof_type.goal Evd.sigma -> + Goal.goal Evd.sigma -> ssrterm -> evar_map * EConstr.t * UState.t * int val pf_interp_ty : ?resolve_typeclasses:bool -> Tacinterp.interp_sign -> - Proof_type.goal Evd.sigma -> + Goal.goal Evd.sigma -> Ssrast.ssrtermkind * (Glob_term.glob_constr * Constrexpr.constr_expr option) -> int * EConstr.t * EConstr.t * UState.t @@ -309,12 +309,12 @@ exception NotEnoughProducts val pf_saturate : ?beta:bool -> ?bi_types:bool -> - Proof_type.goal Evd.sigma -> + Goal.goal Evd.sigma -> EConstr.constr -> ?ty:EConstr.types -> int -> EConstr.constr * EConstr.types * (int * EConstr.constr) list * - Proof_type.goal Evd.sigma + Goal.goal Evd.sigma val saturate : ?beta:bool -> ?bi_types:bool -> @@ -338,32 +338,32 @@ type name_hint = (int * EConstr.types array) option ref val gentac : (Geninterp.interp_sign -> (Ssrast.ssrdocc) * - Ssrmatching_plugin.Ssrmatching.cpattern -> Proof_type.tactic) + Ssrmatching_plugin.Ssrmatching.cpattern -> Tacmach.tactic) val genstac : ((Ssrast.ssrhyp list option * Ssrmatching_plugin.Ssrmatching.occ) * Ssrmatching_plugin.Ssrmatching.cpattern) list * Ssrast.ssrhyp list -> - Tacinterp.interp_sign -> Proof_type.tactic + Tacinterp.interp_sign -> Tacmach.tactic val pf_interp_gen : Tacinterp.interp_sign -> - Proof_type.goal Evd.sigma -> + Goal.goal Evd.sigma -> bool -> (Ssrast.ssrhyp list option * Ssrmatching_plugin.Ssrmatching.occ) * Ssrmatching_plugin.Ssrmatching.cpattern -> EConstr.t * EConstr.t * Ssrast.ssrhyp list * - Proof_type.goal Evd.sigma + Goal.goal Evd.sigma val pf_interp_gen_aux : Tacinterp.interp_sign -> - Proof_type.goal Evd.sigma -> + Goal.goal Evd.sigma -> bool -> (Ssrast.ssrhyp list option * Ssrmatching_plugin.Ssrmatching.occ) * Ssrmatching_plugin.Ssrmatching.cpattern -> bool * Ssrmatching_plugin.Ssrmatching.pattern * EConstr.t * EConstr.t * Ssrast.ssrhyp list * UState.t * - Proof_type.goal Evd.sigma + Goal.goal Evd.sigma val is_name_in_ipats : Id.t -> ssripats -> bool @@ -386,12 +386,12 @@ val interp_clr : val genclrtac : EConstr.constr -> - EConstr.constr list -> Ssrast.ssrhyp list -> Proof_type.tactic + EConstr.constr list -> Ssrast.ssrhyp list -> Tacmach.tactic val cleartac : ssrhyps -> v82tac -val tclMULT : int * ssrmmod -> Proof_type.tactic -> Proof_type.tactic +val tclMULT : int * ssrmmod -> Tacmach.tactic -> Tacmach.tactic -val unprotecttac : Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma +val unprotecttac : Goal.goal Evd.sigma -> Goal.goal list Evd.sigma val abs_wgen : bool -> @@ -401,8 +401,8 @@ val abs_wgen : ((Ssrast.ssrhyp_or_id * string) * Ssrmatching_plugin.Ssrmatching.cpattern option) option -> - Proof_type.goal Evd.sigma * EConstr.t list * EConstr.t -> - Proof_type.goal Evd.sigma * EConstr.t list * EConstr.t + Goal.goal Evd.sigma * EConstr.t list * EConstr.t -> + Goal.goal Evd.sigma * EConstr.t list * EConstr.t val clr_of_wgen : ssrhyps * ((ssrhyp_or_id * 'a) * 'b option) option -> diff --git a/plugins/ssr/ssrelim.mli b/plugins/ssr/ssrelim.mli index 8dc28d8b7..825b4758e 100644 --- a/plugins/ssr/ssrelim.mli +++ b/plugins/ssr/ssrelim.mli @@ -32,23 +32,23 @@ val ssrelim : (?ist:Ltac_plugin.Tacinterp.interp_sign -> 'a -> Ssrast.ssripat option -> - (Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma) -> - bool -> Ssrast.ssrhyp list -> Proof_type.tactic) -> - Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma + (Goal.goal Evd.sigma -> Goal.goal list Evd.sigma) -> + bool -> Ssrast.ssrhyp list -> Tacmach.tactic) -> + Goal.goal Evd.sigma -> Goal.goal list Evd.sigma val elimtac : EConstr.constr -> - Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma + Goal.goal Evd.sigma -> Goal.goal list Evd.sigma val casetac : EConstr.constr -> - Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma + Goal.goal Evd.sigma -> Goal.goal list Evd.sigma -val is_injection_case : EConstr.t -> Proof_type.goal Evd.sigma -> bool +val is_injection_case : EConstr.t -> Goal.goal Evd.sigma -> bool val perform_injection : EConstr.constr -> - Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma + Goal.goal Evd.sigma -> Goal.goal list Evd.sigma val ssrscasetac : bool -> EConstr.constr -> - Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma + Goal.goal Evd.sigma -> Goal.goal list Evd.sigma diff --git a/plugins/ssr/ssrequality.mli b/plugins/ssr/ssrequality.mli index f712002c1..f9ab5d74f 100644 --- a/plugins/ssr/ssrequality.mli +++ b/plugins/ssr/ssrequality.mli @@ -25,12 +25,12 @@ val mkclr : ssrclear -> ssrdocc val nodocc : ssrdocc val noclr : ssrdocc -val simpltac : Ssrast.ssrsimpl -> Proof_type.tactic +val simpltac : Ssrast.ssrsimpl -> Tacmach.tactic val newssrcongrtac : int * Ssrast.ssrterm -> Ltac_plugin.Tacinterp.interp_sign -> - Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma + Goal.goal Evd.sigma -> Goal.goal list Evd.sigma val mk_rwarg : @@ -45,7 +45,7 @@ val ssrinstancesofrule : Ltac_plugin.Tacinterp.interp_sign -> Ssrast.ssrdir -> Ssrast.ssrterm -> - Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma + Goal.goal Evd.sigma -> Goal.goal list Evd.sigma val ssrrewritetac : Ltac_plugin.Tacinterp.interp_sign -> @@ -53,11 +53,11 @@ val ssrrewritetac : (((Ssrast.ssrhyps option * Ssrmatching.occ) * Ssrmatching.rpattern option) * (ssrwkind * Ssrast.ssrterm))) - list -> Proof_type.tactic + list -> Tacmach.tactic -val ipat_rewrite : ssrocc -> ssrdir -> EConstr.t -> Proof_type.tactic +val ipat_rewrite : ssrocc -> ssrdir -> EConstr.t -> Tacmach.tactic val unlocktac : Ltac_plugin.Tacinterp.interp_sign -> (Ssrmatching.occ * Ssrast.ssrterm) list -> - Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma + Goal.goal Evd.sigma -> Goal.goal list Evd.sigma diff --git a/plugins/ssr/ssrfwd.mli b/plugins/ssr/ssrfwd.mli index ead361745..7f254074c 100644 --- a/plugins/ssr/ssrfwd.mli +++ b/plugins/ssr/ssrfwd.mli @@ -36,7 +36,7 @@ val ssrabstract : val basecuttac : string -> - EConstr.t -> Proof_type.goal Evd.sigma -> Evar.t list Evd.sigma + EConstr.t -> Goal.goal Evd.sigma -> Evar.t list Evd.sigma val wlogtac : Ltac_plugin.Tacinterp.interp_sign -> @@ -52,7 +52,7 @@ val wlogtac : Ltac_plugin.Tacinterp.Value.t Ssrast.ssrhint -> bool -> [< `Gen of Names.Id.t option option | `NoGen > `NoGen ] -> - Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma + Goal.goal Evd.sigma -> Goal.goal list Evd.sigma val sufftac : Ssrast.ist -> @@ -62,5 +62,5 @@ val sufftac : (Ssrast.ssrtermkind * (Glob_term.glob_constr * Constrexpr.constr_expr option))) * (bool * Tacinterp.Value.t option list)) -> - Proof_type.tactic + Tacmach.tactic diff --git a/plugins/ssr/ssripats.ml b/plugins/ssr/ssripats.ml index 7ae9e3824..06bbd749e 100644 --- a/plugins/ssr/ssripats.ml +++ b/plugins/ssr/ssripats.ml @@ -175,10 +175,10 @@ let move_top_with_view ~next c r v = type block_names = (int * EConstr.types array) option -let (introstac : ?ist:Tacinterp.interp_sign -> ssripats -> Proof_type.tactic), +let (introstac : ?ist:Tacinterp.interp_sign -> ssripats -> Tacmach.tactic), (tclEQINTROS : ?ind:block_names ref -> ?ist:Tacinterp.interp_sign -> - Proof_type.tactic -> Proof_type.tactic -> ssripats -> - Proof_type.tactic) + Tacmach.tactic -> Tacmach.tactic -> ssripats -> + Tacmach.tactic) = let rec ipattac ?ist ~next p : tac_ctx tac_a = fun gl -> diff --git a/plugins/ssr/ssripats.mli b/plugins/ssr/ssripats.mli index 5f5c7f34a..aefdc8e11 100644 --- a/plugins/ssr/ssripats.mli +++ b/plugins/ssr/ssripats.mli @@ -36,10 +36,10 @@ val elim_intro_tac : ?ist:Tacinterp.interp_sign -> [> `EConstr of 'a * 'b * EConstr.t ] -> Ssrast.ssripat option -> - Proof_type.tactic -> + Tacmach.tactic -> bool -> Ssrast.ssrhyp list -> - Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma + Goal.goal Evd.sigma -> Goal.goal list Evd.sigma (* "move=> top; tac top; clear top" respecting the speed *) val with_top : (EConstr.t -> v82tac) -> tac_ctx tac_a @@ -51,17 +51,17 @@ val ssrmovetac : (((Ssrast.ssrdocc * Ssrmatching.cpattern) list list * Ssrast.ssrclear) * Ssrast.ssripats)) -> - Proof_type.tactic + Tacmach.tactic -val movehnftac : Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma +val movehnftac : Goal.goal Evd.sigma -> Goal.goal list Evd.sigma val with_dgens : (Ssrast.ssrdocc * Ssrmatching.cpattern) list list * Ssrast.ssrclear -> ((Ssrast.ssrdocc * Ssrmatching.cpattern) list -> Ssrast.ssrdocc * Ssrmatching.cpattern -> - Ltac_plugin.Tacinterp.interp_sign -> Proof_type.tactic) -> - Ltac_plugin.Tacinterp.interp_sign -> Proof_type.tactic + Ltac_plugin.Tacinterp.interp_sign -> Tacmach.tactic) -> + Ltac_plugin.Tacinterp.interp_sign -> Tacmach.tactic val ssrcasetac : Ltac_plugin.Tacinterp.interp_sign -> @@ -69,7 +69,7 @@ val ssrcasetac : (Ssrast.ssripat option * (((Ssrast.ssrdocc * Ssrmatching.cpattern) list list * Ssrast.ssrclear) * Ssrast.ssripats)) -> - Proof_type.tactic + Tacmach.tactic val ssrapplytac : Tacinterp.interp_sign -> @@ -79,5 +79,5 @@ val ssrapplytac : (Ssrast.ssrtermkind * Tacexpr.glob_constr_and_expr)) list list * Ssrast.ssrhyps) * Ssrast.ssripats)) -> - Proof_type.tactic + Tacmach.tactic diff --git a/plugins/ssr/ssrprinters.mli b/plugins/ssr/ssrprinters.mli index 9207b9e43..8da9bc72b 100644 --- a/plugins/ssr/ssrprinters.mli +++ b/plugins/ssr/ssrprinters.mli @@ -12,7 +12,7 @@ open API open Ssrast val pp_term : - Proof_type.goal Evd.sigma -> EConstr.constr -> Pp.std_ppcmds + Goal.goal Evd.sigma -> EConstr.constr -> Pp.std_ppcmds val pr_spc : unit -> Pp.std_ppcmds val pr_bar : unit -> Pp.std_ppcmds diff --git a/plugins/ssr/ssrtacticals.mli b/plugins/ssr/ssrtacticals.mli index 1d1887138..297cfdfdc 100644 --- a/plugins/ssr/ssrtacticals.mli +++ b/plugins/ssr/ssrtacticals.mli @@ -17,7 +17,7 @@ val tclSEQAT : int Misctypes.or_var * (('a * Ltac_plugin.Tacinterp.Value.t option list) * Ltac_plugin.Tacinterp.Value.t option) -> - Proof_type.tactic + Tacmach.tactic val tclCLAUSES : Ltac_plugin.Tacinterp.interp_sign -> @@ -27,7 +27,7 @@ val tclCLAUSES : Ssrmatching_plugin.Ssrmatching.cpattern option) option) list * Ssrast.ssrclseq -> - Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma + Goal.goal Evd.sigma -> Goal.goal list Evd.sigma val hinttac : Tacinterp.interp_sign -> @@ -42,5 +42,5 @@ val ssrdotac : Ssrmatching_plugin.Ssrmatching.cpattern option) option) list * Ssrast.ssrclseq) -> - Proof_type.goal Evd.sigma -> Proof_type.goal list Evd.sigma + Goal.goal Evd.sigma -> Goal.goal list Evd.sigma diff --git a/plugins/ssr/ssrview.ml b/plugins/ssr/ssrview.ml index 91e40f368..cc142e091 100644 --- a/plugins/ssr/ssrview.ml +++ b/plugins/ssr/ssrview.ml @@ -80,7 +80,7 @@ let interp_view ist si env sigma gv v rid = snd (view_with (if view_nbimps < 0 then [] else viewtab.(0))) -let with_view ist ~next si env (gl0 : (Proof_type.goal * tac_ctx) Evd.sigma) c name cl prune (conclude : EConstr.t -> EConstr.t -> tac_ctx tac_a) clr = +let with_view ist ~next si env (gl0 : (Goal.goal * tac_ctx) Evd.sigma) c name cl prune (conclude : EConstr.t -> EConstr.t -> tac_ctx tac_a) clr = let c2r ist x = { ist with lfun = Id.Map.add top_id (Value.of_constr x) ist.lfun } in let terminate (sigma, c') = diff --git a/plugins/ssrmatching/ssrmatching.mli b/plugins/ssrmatching/ssrmatching.mli index c2bf12cb6..1853bc35d 100644 --- a/plugins/ssrmatching/ssrmatching.mli +++ b/plugins/ssrmatching/ssrmatching.mli @@ -3,11 +3,11 @@ open API open Grammar_API +open Goal open Genarg open Tacexpr open Environ open Evd -open Proof_type open Term (** ******** Small Scale Reflection pattern matching facilities ************* *) -- cgit v1.2.3