diff options
| author | Gaëtan Gilbert | 2019-04-24 13:56:25 +0200 |
|---|---|---|
| committer | Gaëtan Gilbert | 2019-04-24 14:01:11 +0200 |
| commit | 7fd23f30ef6c52aabc8d1084171d906e1080f3f2 (patch) | |
| tree | ec73c181399217e12f56d77c868d7b8d4be82382 /engine | |
| parent | ed1b663afdd02355877c51dc9c759ddfd91b0854 (diff) | |
Fix proof bullet error helper (nosuchgoal)
The [int] is incorrect for list focusing, we could work a bit harder
to fix that. It's only used for pluralisation in the error message "no
such goal(s)" so we could also ignore the issue.
Diffstat (limited to 'engine')
| -rw-r--r-- | engine/proofview.ml | 16 | ||||
| -rw-r--r-- | engine/proofview.mli | 19 |
2 files changed, 17 insertions, 18 deletions
diff --git a/engine/proofview.ml b/engine/proofview.ml index 86dcb8ab3b..f278c83912 100644 --- a/engine/proofview.ml +++ b/engine/proofview.ml @@ -352,11 +352,12 @@ let _ = CErrors.register_handler begin function | _ -> raise CErrors.Unhandled end -(** [tclFOCUS_gen nosuchgoal i j t] applies [t] in a context where +(** [tclFOCUS ?nosuchgoal i j t] applies [t] in a context where only the goals numbered [i] to [j] are focused (the rest of the goals is restored at the end of the tactic). If the range [i]-[j] is not valid, then it [tclFOCUS_gen nosuchgoal i j t] is [nosuchgoal]. *) -let tclFOCUS_gen nosuchgoal i j t = +let tclFOCUS ?nosuchgoal i j t = + let nosuchgoal = Option.default (tclZERO (NoSuchGoals (j+1-i))) nosuchgoal in let open Proof in Pv.get >>= fun initial -> try @@ -367,10 +368,9 @@ let tclFOCUS_gen nosuchgoal i j t = return result with CList.IndexOutOfRange -> nosuchgoal -let tclFOCUS i j t = tclFOCUS_gen (tclZERO (NoSuchGoals (j+1-i))) i j t -let tclTRYFOCUS i j t = tclFOCUS_gen (tclUNIT ()) i j t +let tclTRYFOCUS i j t = tclFOCUS ~nosuchgoal:(tclUNIT ()) i j t -let tclFOCUSLIST l t = +let tclFOCUSLIST ?(nosuchgoal=tclZERO (NoSuchGoals 0)) l t = let open Proof in Comb.get >>= fun comb -> let n = CList.length comb in @@ -384,7 +384,7 @@ let tclFOCUSLIST l t = in let l = CList.map_filter sanitize l in match l with - | [] -> tclZERO (NoSuchGoals 0) + | [] -> nosuchgoal | (mi, _) :: _ -> (* Get the left-most goal to focus. This goal won't move, and we will then place all the other goals to focus to the right. *) @@ -401,7 +401,7 @@ let tclFOCUSLIST l t = (** Like {!tclFOCUS} but selects a single goal by name. *) -let tclFOCUSID id t = +let tclFOCUSID ?(nosuchgoal=tclZERO (NoSuchGoals 1)) id t = let open Proof in Pv.get >>= fun initial -> try @@ -421,7 +421,7 @@ let tclFOCUSID id t = t >>= fun result -> Comb.set initial.comb >> return result - with Not_found -> tclZERO (NoSuchGoals 1) + with Not_found -> nosuchgoal (** {7 Dispatching on goals} *) diff --git a/engine/proofview.mli b/engine/proofview.mli index 8685430c62..9455dae643 100644 --- a/engine/proofview.mli +++ b/engine/proofview.mli @@ -244,14 +244,12 @@ val tclBREAK : (iexn -> iexn option) -> 'a tactic -> 'a tactic (** [tclFOCUS i j t] applies [t] after focusing on the goals number [i] to [j] (see {!focus}). The rest of the goals is restored after the tactic action. If the specified range doesn't correspond to - existing goals, fails with [NoSuchGoals] (a user error). this - exception is caught at toplevel with a default message + a hook - message that can be customized by [set_nosuchgoals_hook] below. - This hook is used to add a suggestion about bullets when - applicable. *) + existing goals, fails with the [nosuchgoal] argument, by default + raising [NoSuchGoals] (a user error). This exception is caught at + toplevel with a default message. *) exception NoSuchGoals of int -val tclFOCUS : int -> int -> 'a tactic -> 'a tactic +val tclFOCUS : ?nosuchgoal:'a tactic -> int -> int -> 'a tactic -> 'a tactic (** [tclFOCUSLIST li t] applies [t] on the list of focused goals described by [li]. Each element of [li] is a pair [(i, j)] denoting @@ -260,13 +258,14 @@ val tclFOCUS : int -> int -> 'a tactic -> 'a tactic intervals. If the set of such goals is not a single range, then it will move goals such that it is a single range. (So, for instance, [[1, 3-5]; idtac.] is not the identity.) - If the set of such goals is empty, it will fail. *) -val tclFOCUSLIST : (int * int) list -> 'a tactic -> 'a tactic + If the set of such goals is empty, it will fail with [nosuchgoal], + by default raising [NoSuchGoals 0]. *) +val tclFOCUSLIST : ?nosuchgoal:'a tactic -> (int * int) list -> 'a tactic -> 'a tactic (** [tclFOCUSID x t] applies [t] on a (single) focused goal like {!tclFOCUS}. The goal is found by its name rather than its - number.*) -val tclFOCUSID : Names.Id.t -> 'a tactic -> 'a tactic + number. Fails with [nosuchgoal], by default raising [NoSuchGoals 1]. *) +val tclFOCUSID : ?nosuchgoal:'a tactic -> Names.Id.t -> 'a tactic -> 'a tactic (** [tclTRYFOCUS i j t] behaves like {!tclFOCUS}, except that if the specified range doesn't correspond to existing goals, behaves like |
