Module Proof_diffs
val string_to_diffs : string -> diffOpt
val make_goal : Environ.env -> Evd.evar_map -> Evar.t -> goal
val diff_goal : ?short:bool -> ?og_s:goal -> goal -> Pp.t list * Pp.t
Computes the diff between the goals of two Proofs and returns the highlighted lists of hypotheses and conclusions.
If the strings used to display the goal are not lexable (this is believed unlikely), this routine will generate a Diff_Failure. This routine may also raise Diff_Failure under some "impossible" conditions.
If you want to make your call especially bulletproof, catch these exceptions, print a user-visible message, then recall this routine with the second argument ('og_s') set to None, which will skip the diff.
The 'short' flag skips computing the diffs of the hypotheses, which allows the Subgoals XML command to fetch just the conclusion of a goal. This is useful, for example, when an IDE only needs to display the number of admitted goals or preview the next unsolved goal.
val tokenize_string : string -> string list
Convert a string to a list of token strings using the lexer
val make_goal_map : Proof.t -> Proof.t -> goal_map
Generates a map from
np
toop
that maps changed goals to their prior forms. The map doesn't include entries for unchanged goals; unchanged goals will have the same goal id in both versions.op
andnp
must be from the same proof document andop
must be for a state beforenp
.
type hyp_info
=
{
idents : string list;
rhs_pp : Pp.t;
}
val diff_hyps : string list list -> hyp_info CString.Map.t -> string list list -> hyp_info CString.Map.t -> Pp.t list
val diff_proofs : diff_opt:diffOpt -> ?old:Proof.t -> Proof.t -> Pp.t
val notify_proof_diff_failure : string -> unit