aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Anton Trunov <anton.a.trunov@gmail.com>2018-03-04 18:14:21 +0100
committerGravatar Anton Trunov <anton.a.trunov@gmail.com>2018-03-04 19:24:19 +0100
commitd147b10a7fede25214416165b6a6ac10b6b827de (patch)
tree0f987d97c061dbecf0d476f7ca2116885b3964a1
parent5233abd2f876af6b9e5b705b2ab4df8cdb50b12e (diff)
ssr: add Prenex Implicits for Some_inj to use it as a view
-rw-r--r--plugins/ssr/ssrfun.v3
-rw-r--r--test-suite/bugs/closed/6910.v5
2 files changed, 8 insertions, 0 deletions
diff --git a/plugins/ssr/ssrfun.v b/plugins/ssr/ssrfun.v
index 3aabff824..f44596cb7 100644
--- a/plugins/ssr/ssrfun.v
+++ b/plugins/ssr/ssrfun.v
@@ -637,6 +637,9 @@ End Injections.
Lemma Some_inj {T} : injective (@Some T). Proof. by move=> x y []. Qed.
+(* Force implicits to use as a view. *)
+Prenex Implicits Some_inj.
+
(* cancellation lemmas for dependent type casts. *)
Lemma esymK T x y : cancel (@esym T x y) (@esym T y x).
Proof. by case: y /. Qed.
diff --git a/test-suite/bugs/closed/6910.v b/test-suite/bugs/closed/6910.v
new file mode 100644
index 000000000..5167a5364
--- /dev/null
+++ b/test-suite/bugs/closed/6910.v
@@ -0,0 +1,5 @@
+From Coq Require Import ssreflect ssrfun.
+
+(* We should be able to use Some_inj as a view: *)
+Lemma foo (x y : nat) : Some x = Some y -> x = y.
+Proof. by move/Some_inj. Qed.