aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr>2016-10-12 17:43:51 +0200
committerGravatar Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr>2016-10-12 17:44:34 +0200
commit112e974ec90b2afc51a7cffeba49e5777f3ea80f (patch)
tree1bf579fca94adea5ee4a423c0a96c8832e68861e
parent8b2a08ecd123515778584596918666e5f49076f7 (diff)
parent6d55121c90ec50319a3de6a6907726fbcdc2f835 (diff)
Merge branch 'v8.5' into v8.6
-rw-r--r--engine/ftactic.ml23
-rw-r--r--engine/proofview.ml6
-rw-r--r--engine/proofview.mli7
-rw-r--r--kernel/entries.mli7
-rw-r--r--lib/monad.ml11
-rw-r--r--lib/monad.mli3
-rw-r--r--proofs/refine.ml22
-rw-r--r--test-suite/bugs/closed/4416.v3
-rw-r--r--test-suite/bugs/closed/4863.v32
9 files changed, 109 insertions, 5 deletions
diff --git a/engine/ftactic.ml b/engine/ftactic.ml
index 588709873..aeaaea7e4 100644
--- a/engine/ftactic.ml
+++ b/engine/ftactic.ml
@@ -29,13 +29,28 @@ let bind (type a) (type b) (m : a t) (f : a -> b t) : b t = m >>= function
| Uniform x ->
(** We dispatch the uniform result on each goal under focus, as we know
that the [m] argument was actually dependent. *)
- Proofview.Goal.goals >>= fun l ->
- let ans = List.map (fun _ -> x) l in
+ Proofview.Goal.goals >>= fun goals ->
+ let ans = List.map (fun g -> (g,x)) goals in
Proofview.tclUNIT ans
- | Depends l -> Proofview.tclUNIT l
+ | Depends l ->
+ Proofview.Goal.goals >>= fun goals ->
+ Proofview.tclUNIT (List.combine goals l)
+ in
+ (* After the tactic has run, some goals which were previously
+ produced may have been solved by side effects. The values
+ attached to such goals must be discarded, otherwise the list of
+ result would not have the same length as the list of focused
+ goals, which is an invariant of the [Ftactic] module. It is the
+ reason why a goal is attached to each result above. *)
+ let filter (g,x) =
+ g >>= fun g ->
+ Proofview.Goal.unsolved g >>= function
+ | true -> Proofview.tclUNIT (Some x)
+ | false -> Proofview.tclUNIT None
in
Proofview.tclDISPATCHL (List.map f l) >>= fun l ->
- Proofview.tclUNIT (Depends (List.concat l))
+ Proofview.Monad.List.map_filter filter (List.concat l) >>= fun filtered ->
+ Proofview.tclUNIT (Depends filtered)
let goals = Proofview.Goal.goals >>= fun l -> Proofview.tclUNIT (Depends l)
let set_sigma r =
diff --git a/engine/proofview.ml b/engine/proofview.ml
index a2838a2de..f2f400515 100644
--- a/engine/proofview.ml
+++ b/engine/proofview.ml
@@ -929,6 +929,8 @@ module Unsafe = struct
{ step with comb = step.comb @ gls }
end
+ let tclSETENV = Env.set
+
let tclGETGOALS = Comb.get
let tclSETGOALS = Comb.set
@@ -1129,6 +1131,10 @@ module Goal = struct
in
tclUNIT (CList.map_filter map step.comb)
+ let unsolved { self=self } =
+ tclEVARMAP >>= fun sigma ->
+ tclUNIT (not (Option.is_empty (advance sigma self)))
+
(* compatibility *)
let goal { self=self } = self
diff --git a/engine/proofview.mli b/engine/proofview.mli
index bc68f11ff..fae75f825 100644
--- a/engine/proofview.mli
+++ b/engine/proofview.mli
@@ -409,6 +409,9 @@ module Unsafe : sig
(** Like {!tclEVARS} but also checks whether goals have been solved. *)
val tclEVARSADVANCE : Evd.evar_map -> unit tactic
+ (** Set the global environment of the tactic *)
+ val tclSETENV : Environ.env -> unit tactic
+
(** [tclNEWGOALS gls] adds the goals [gls] to the ones currently
being proved, appending them to the list of focused goals. If a
goal is already solved, it is not added. *)
@@ -518,6 +521,10 @@ module Goal : sig
FIXME: encapsulate the level in an existential type. *)
val goals : ([ `LZ ], 'r) t tactic list tactic
+ (** [unsolved g] is [true] if [g] is still unsolved in the current
+ proof state. *)
+ val unsolved : ('a, 'r) t -> bool tactic
+
(** Compatibility: avoid if possible *)
val goal : ([ `NF ], 'r) t -> Evar.t
diff --git a/kernel/entries.mli b/kernel/entries.mli
index df2c4653f..ea7c266bc 100644
--- a/kernel/entries.mli
+++ b/kernel/entries.mli
@@ -98,7 +98,12 @@ type module_entry =
| MExpr of
module_params_entry * module_struct_entry * module_struct_entry option
-type seff_env = [ `Nothing | `Opaque of Constr.t * Univ.universe_context_set ]
+
+type seff_env =
+ [ `Nothing
+ (* The proof term and its universes.
+ Same as the constant_body's but not in an ephemeron *)
+ | `Opaque of Constr.t * Univ.universe_context_set ]
type side_eff =
| SEsubproof of constant * Declarations.constant_body * seff_env
diff --git a/lib/monad.ml b/lib/monad.ml
index a1714a41b..2e55e9698 100644
--- a/lib/monad.ml
+++ b/lib/monad.ml
@@ -64,6 +64,9 @@ module type ListS = sig
its second argument in a tail position. *)
val iter : ('a -> unit t) -> 'a list -> unit t
+ (** Like the regular {!CList.map_filter}. The monadic effects are threaded left*)
+ val map_filter : ('a -> 'b option t) -> 'a list -> 'b list t
+
(** {6 Two-list iterators} *)
@@ -138,6 +141,14 @@ module Make (M:Def) : S with type +'a t = 'a M.t = struct
| a::b::l -> f a >> f b >> iter f l
+ let rec map_filter f = function
+ | [] -> return []
+ | a::l ->
+ f a >>= function
+ | None -> map_filter f l
+ | Some b ->
+ map_filter f l >>= fun filtered ->
+ return (b::filtered)
let rec fold_left2 r f x l1 l2 =
match l1,l2 with
diff --git a/lib/monad.mli b/lib/monad.mli
index c8655efa0..f7de71f53 100644
--- a/lib/monad.mli
+++ b/lib/monad.mli
@@ -66,6 +66,9 @@ module type ListS = sig
its second argument in a tail position. *)
val iter : ('a -> unit t) -> 'a list -> unit t
+ (** Like the regular {!CList.map_filter}. The monadic effects are threaded left*)
+ val map_filter : ('a -> 'b option t) -> 'a list -> 'b list t
+
(** {6 Two-list iterators} *)
diff --git a/proofs/refine.ml b/proofs/refine.ml
index dc6f4cea1..e5114a2ec 100644
--- a/proofs/refine.ml
+++ b/proofs/refine.ml
@@ -51,6 +51,23 @@ let typecheck_proof c concl env sigma =
let (pr_constrv,pr_constr) =
Hook.make ~default:(fun _env _sigma _c -> Pp.str"<constr>") ()
+(* Get the side-effect's constant declarations to update the monad's
+ * environmnent *)
+let add_if_undefined kn cb env =
+ try ignore(Environ.lookup_constant kn env); env
+ with Not_found -> Environ.add_constant kn cb env
+
+(* Add the side effects to the monad's environment, if not already done. *)
+let add_side_effect env = function
+ | { Entries.eff = Entries.SEsubproof (kn, cb, eff_env) } ->
+ add_if_undefined kn cb env
+ | { Entries.eff = Entries.SEscheme (l,_) } ->
+ List.fold_left (fun env (_,kn,cb,eff_env) ->
+ add_if_undefined kn cb env) env l
+
+let add_side_effects env effects =
+ List.fold_left (fun env eff -> add_side_effect env eff) env effects
+
let make_refine_enter ?(unsafe = true) f =
{ enter = fun gl ->
let gl = Proofview.Goal.assume gl in
@@ -66,6 +83,10 @@ let make_refine_enter ?(unsafe = true) f =
let ((v,c), sigma) = Sigma.run (Evd.reset_future_goals sigma) f in
let evs = Evd.future_goals sigma in
let evkmain = Evd.principal_future_goal sigma in
+ (** Redo the effects in sigma in the monad's env *)
+ let privates_csts = Evd.eval_side_effects sigma in
+ let sideff = Safe_typing.side_effects_of_private_constants privates_csts in
+ let env = add_side_effects env sideff in
(** Check that the introduced evars are well-typed *)
let fold accu ev = typecheck_evar ev env accu in
let sigma = if unsafe then sigma else CList.fold_left fold sigma evs in
@@ -94,6 +115,7 @@ let make_refine_enter ?(unsafe = true) f =
let sigma = CList.fold_left Proofview.Unsafe.mark_as_goal sigma comb in
let trace () = Pp.(hov 2 (str"simple refine"++spc()++ Hook.get pr_constrv env sigma c)) in
Proofview.Trace.name_tactic trace (Proofview.tclUNIT v) >>= fun v ->
+ Proofview.Unsafe.tclSETENV (Environ.reset_context env) <*>
Proofview.Unsafe.tclEVARS sigma <*>
Proofview.Unsafe.tclSETGOALS comb <*>
Proofview.tclUNIT v
diff --git a/test-suite/bugs/closed/4416.v b/test-suite/bugs/closed/4416.v
new file mode 100644
index 000000000..b97a8ce64
--- /dev/null
+++ b/test-suite/bugs/closed/4416.v
@@ -0,0 +1,3 @@
+Goal exists x, x.
+unshelve refine (ex_intro _ _ _); match goal with _ => refine (_ _) end.
+(* Error: Incorrect number of goals (expected 2 tactics). *) \ No newline at end of file
diff --git a/test-suite/bugs/closed/4863.v b/test-suite/bugs/closed/4863.v
new file mode 100644
index 000000000..e884355fd
--- /dev/null
+++ b/test-suite/bugs/closed/4863.v
@@ -0,0 +1,32 @@
+Require Import Classes.DecidableClass.
+
+Inductive Foo : Set :=
+| foo1 | foo2.
+
+Instance Decidable_sumbool : forall P, {P}+{~P} -> Decidable P.
+Proof.
+ intros P H.
+ refine (Build_Decidable _ (if H then true else false) _).
+ intuition congruence.
+Qed.
+
+Hint Extern 100 ({?A = ?B}+{~ ?A = ?B}) => abstract (abstract (abstract (decide equality))) : typeclass_instances.
+
+Goal forall (a b : Foo), {a=b}+{a<>b}.
+intros.
+abstract (abstract (decide equality)). (*abstract works here*)
+Qed.
+
+Check ltac:(abstract (exact I)) : True.
+
+Goal forall (a b : Foo), Decidable (a=b) * Decidable (a=b).
+intros.
+split. typeclasses eauto. typeclasses eauto. Qed.
+
+Goal forall (a b : Foo), Decidable (a=b) * Decidable (a=b).
+intros.
+split.
+refine _.
+refine _.
+Defined.
+(*fails*)