aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Enrico Tassi <Enrico.Tassi@inria.fr>2016-11-07 16:03:53 +0100
committerGravatar Enrico Tassi <Enrico.Tassi@inria.fr>2017-02-10 13:05:23 +0100
commitf30822ba46d08d65597e0e3230ea6ad86c98453f (patch)
tree0b938d18a8f38f3ca08969e022ffa55f1fd676fa
parentebbc4a8cf3dde3b18388f3723b1641ba5511db9b (diff)
Proofview: tclINDEPENDENTL
-rw-r--r--engine/proofview.ml29
-rw-r--r--engine/proofview.mli1
-rw-r--r--tactics/tacticals.ml13
-rw-r--r--tactics/tacticals.mli1
4 files changed, 43 insertions, 1 deletions
diff --git a/engine/proofview.ml b/engine/proofview.ml
index 2fbabb749..721389af4 100644
--- a/engine/proofview.ml
+++ b/engine/proofview.ml
@@ -453,6 +453,25 @@ let iter_goal i =
Solution.get >>= fun evd ->
Comb.set CList.(undefined evd (flatten (rev subgoals)))
+(** List iter but allocates a list of results *)
+let map_goal i =
+ let rev = List.rev in (* hem... Proof masks List... *)
+ let open Proof in
+ Comb.get >>= fun initial ->
+ Proof.List.fold_left begin fun (acc, subgoals as cur) goal ->
+ Solution.get >>= fun step ->
+ match Evarutil.advance step goal with
+ | None -> return cur
+ | Some goal ->
+ Comb.set [goal] >>
+ i goal >>= fun res ->
+ Proof.map (fun comb -> comb :: subgoals) Comb.get >>= fun x ->
+ return (res :: acc, x)
+ end ([],[]) initial >>= fun (results_rev, subgoals) ->
+ Solution.get >>= fun evd ->
+ Comb.set CList.(undefined evd (flatten (rev subgoals))) >>
+ return (rev results_rev)
+
(** A variant of [Monad.List.fold_left2] where the first list is the
list of focused goals. The argument tactic is executed in a focus
comprising only of the current goal, a goal which has been solved
@@ -585,7 +604,15 @@ let tclINDEPENDENT tac =
let tac = InfoL.tag (Info.DBranch) tac in
InfoL.tag (Info.Dispatch) (iter_goal (fun _ -> tac))
-
+let tclINDEPENDENTL tac =
+ let open Proof in
+ Pv.get >>= fun initial ->
+ match initial.comb with
+ | [] -> tclUNIT []
+ | [_] -> tac >>= fun x -> return [x]
+ | _ ->
+ let tac = InfoL.tag (Info.DBranch) tac in
+ InfoL.tag (Info.Dispatch) (map_goal (fun _ -> tac))
(** {7 Goal manipulation} *)
diff --git a/engine/proofview.mli b/engine/proofview.mli
index 90be2f90a..294b03dca 100644
--- a/engine/proofview.mli
+++ b/engine/proofview.mli
@@ -292,6 +292,7 @@ val tclEXTEND : unit tactic list -> unit tactic -> unit tactic list -> unit tact
independent of backtracking in another. It is equivalent to
[tclEXTEND [] tac []]. *)
val tclINDEPENDENT : unit tactic -> unit tactic
+val tclINDEPENDENTL: 'a tactic -> 'a list tactic
(** {7 Goal manipulation} *)
diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml
index 93c04e373..c5562b326 100644
--- a/tactics/tacticals.ml
+++ b/tactics/tacticals.ml
@@ -368,6 +368,16 @@ module New = struct
catch_failerror e <*> t2
end
end
+
+ let tclORELSE0L t1 t2 =
+ tclINDEPENDENTL begin
+ tclORELSE
+ t1
+ begin fun e ->
+ catch_failerror e <*> t2
+ end
+ end
+
let tclORELSE t1 t2 =
tclORELSE0 (tclPROGRESS t1) t2
@@ -419,6 +429,9 @@ module New = struct
let tclTRY t =
tclORELSE0 t (tclUNIT ())
+
+ let tclTRYb t =
+ tclORELSE0L (t <*> tclUNIT true) (tclUNIT false)
let tclIFTHENELSE t1 t2 t3 =
tclINDEPENDENT begin
diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli
index 18cf03c51..7aacc52f3 100644
--- a/tactics/tacticals.mli
+++ b/tactics/tacticals.mli
@@ -209,6 +209,7 @@ module New : sig
val tclMAP : ('a -> unit tactic) -> 'a list -> unit tactic
val tclTRY : unit tactic -> unit tactic
+ val tclTRYb : unit tactic -> bool list tactic
val tclFIRST : unit tactic list -> unit tactic
val tclIFTHENELSE : unit tactic -> unit tactic -> unit tactic -> unit tactic
val tclIFTHENSVELSE : unit tactic -> unit tactic array -> unit tactic -> unit tactic