Module Tacticals
Tacticals i.e. functions from tactics to tactics.
type tactic= Proofview.V82.tac
val tclIDTAC : tacticval tclIDTAC_MESSAGE : Pp.t -> tacticval tclORELSE0 : tactic -> tactic -> tacticval tclORELSE : tactic -> tactic -> tacticval tclTHEN : tactic -> tactic -> tacticval tclTHENLIST : tactic list -> tacticval tclTHEN_i : tactic -> (int -> tactic) -> tacticval tclTHENFIRST : tactic -> tactic -> tacticval tclTHENLAST : tactic -> tactic -> tacticval tclTHENS : tactic -> tactic list -> tacticval tclTHENSV : tactic -> tactic array -> tacticval tclTHENSLASTn : tactic -> tactic -> tactic array -> tacticval tclTHENSFIRSTn : tactic -> tactic array -> tactic -> tacticval tclREPEAT : tactic -> tacticval tclFIRST : tactic list -> tacticval tclTRY : tactic -> tacticval tclCOMPLETE : tactic -> tacticval tclAT_LEAST_ONCE : tactic -> tacticval tclFAIL : int -> Pp.t -> tacticval tclFAIL_lazy : int -> Pp.t Stdlib.Lazy.t -> tacticval tclDO : int -> tactic -> tacticval tclPROGRESS : tactic -> tacticval tclSHOWHYPS : tactic -> tacticval tclTHENTRY : tactic -> tactic -> tacticval tclMAP : ('a -> tactic) -> 'a list -> tactic
Tacticals applying to hypotheses
val onNthHypId : int -> (Names.Id.t -> tactic) -> tacticval onNthHyp : int -> (EConstr.constr -> tactic) -> tacticval onNthDecl : int -> (EConstr.named_declaration -> tactic) -> tacticval onLastHypId : (Names.Id.t -> tactic) -> tacticval onLastHyp : (EConstr.constr -> tactic) -> tacticval onLastDecl : (EConstr.named_declaration -> tactic) -> tacticval onNLastHypsId : int -> (Names.Id.t list -> tactic) -> tacticval onNLastHyps : int -> (EConstr.constr list -> tactic) -> tacticval onNLastDecls : int -> (EConstr.named_context -> tactic) -> tacticval lastHypId : Goal.goal Evd.sigma -> Names.Id.tval lastHyp : Goal.goal Evd.sigma -> EConstr.constrval lastDecl : Goal.goal Evd.sigma -> EConstr.named_declarationval nLastHypsId : int -> Goal.goal Evd.sigma -> Names.Id.t listval nLastHyps : int -> Goal.goal Evd.sigma -> EConstr.constr listval nLastDecls : int -> Goal.goal Evd.sigma -> EConstr.named_contextval afterHyp : Names.Id.t -> Goal.goal Evd.sigma -> EConstr.named_contextval ifOnHyp : ((Names.Id.t * EConstr.types) -> bool) -> (Names.Id.t -> tactic) -> (Names.Id.t -> tactic) -> Names.Id.t -> tacticval onHyps : (Goal.goal Evd.sigma -> EConstr.named_context) -> (EConstr.named_context -> tactic) -> tactic
Tacticals applying to goal components
val onAllHyps : (Names.Id.t -> tactic) -> tacticval onAllHypsAndConcl : (Names.Id.t option -> tactic) -> tacticval onClause : (Names.Id.t option -> tactic) -> Locus.clause -> tacticval onClauseLR : (Names.Id.t option -> tactic) -> Locus.clause -> tactic
Elimination tacticals.
type branch_args= private{ity : Constr.pinductive;the type we were eliminating on
largs : EConstr.constr list;its arguments
branchnum : int;the branch number
pred : EConstr.constr;the predicate we used
nassums : int;number of assumptions/letin to be introduced
branchsign : bool list;the signature of the branch. true=assumption, false=let-in
branchnames : Tactypes.intro_patterns;}type branch_assumptions= private{ba : branch_args;the branch args
assums : EConstr.named_context;}the list of assumptions introduced
val get_and_check_or_and_pattern : ?loc:Loc.t -> Tactypes.delayed_open_constr Tactypes.or_and_intro_pattern_expr -> bool list array -> Tactypes.intro_patterns arrayget_and_check_or_and_pattern loc pats branchsignreturns an appropriate error message if |pats| <> |branchsign|; extends them if no pattern is given for let-ins in the case of a conjunctive pattern
val fix_empty_or_and_pattern : int -> Tactypes.delayed_open_constr Tactypes.or_and_intro_pattern_expr -> Tactypes.delayed_open_constr Tactypes.or_and_intro_pattern_exprTolerate "
" to mean a disjunctive pattern of any length
val compute_constructor_signatures : rec_flag:bool -> (Names.inductive * 'a) -> bool list arrayval compute_induction_names : bool list array -> Tactypes.or_and_intro_pattern option -> Tactypes.intro_patterns arrayUseful for
as intro_patternmodifier
val elimination_sort_of_goal : Goal.goal Evd.sigma -> Sorts.familyval elimination_sort_of_hyp : Names.Id.t -> Goal.goal Evd.sigma -> Sorts.familyval elimination_sort_of_clause : Names.Id.t option -> Goal.goal Evd.sigma -> Sorts.familyval pf_with_evars : (Goal.goal Evd.sigma -> Evd.evar_map * 'a) -> ('a -> tactic) -> tacticval pf_constr_of_global : Names.GlobRef.t -> (EConstr.constr -> tactic) -> tactic
module New : sig ... endThe tacticals in the module
Neware the tactical of Ltac. Their semantics is an extension of the tacticals in this file for the multi-goal backtracking tactics. They do not have the same semantics as the similarly named tacticals inProofview. The tactical ofProofvieware used in the definition of the tacticals ofTacticals.New, but they are more atomic. In particularTacticals.New.tclORELSEsees lack of progress as a failure, whereasProofview.tclORELSEdoesn't. Additionally every tactic which can catch failure (tclOR,tclORELSE,tclTRY,tclREPEAt, etc…) are run into each goal independently (failures and backtracks are localised to a given goal).