aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--plugins/ssrmatching/ssrmatching.ml414
-rw-r--r--test-suite/bugs/closed/5692.v38
2 files changed, 45 insertions, 7 deletions
diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4
index 2e5522b83..e3e34616b 100644
--- a/plugins/ssrmatching/ssrmatching.ml4
+++ b/plugins/ssrmatching/ssrmatching.ml4
@@ -502,16 +502,16 @@ let ungen_upat lhs (sigma, uc, t) u =
let nb_cs_proj_args pc f u =
let na k =
List.length (snd (lookup_canonical_conversion (ConstRef pc, k))).o_TCOMPS in
- try match kind_of_term f with
- | Prod _ -> na Prod_cs
- | Sort s -> na (Sort_cs (family_of_sort s))
- | Const (c',_) when Constant.equal c' pc ->
- begin match kind_of_term u.up_f with
+ let nargs_of_proj t = match kind_of_term t with
| App(_,args) -> Array.length args
| Proj _ -> 0 (* if splay_app calls expand_projection, this has to be
the number of arguments including the projected *)
- | _ -> assert false
- end
+ | _ -> assert false in
+ try match kind_of_term f with
+ | Prod _ -> na Prod_cs
+ | Sort s -> na (Sort_cs (family_of_sort s))
+ | Const (c',_) when Constant.equal c' pc -> nargs_of_proj u.up_f
+ | Proj (c',_) when Constant.equal (Projection.constant c') pc -> nargs_of_proj u.up_f
| Var _ | Ind _ | Construct _ | Const _ -> na (Const_cs (global_of_constr f))
| _ -> -1
with Not_found -> -1
diff --git a/test-suite/bugs/closed/5692.v b/test-suite/bugs/closed/5692.v
new file mode 100644
index 000000000..55ef7abe4
--- /dev/null
+++ b/test-suite/bugs/closed/5692.v
@@ -0,0 +1,38 @@
+Set Primitive Projections.
+Require Import ZArith ssreflect.
+
+Module Test3.
+
+Set Primitive Projections.
+
+Structure semigroup := SemiGroup {
+ sg_car :> Type;
+ sg_op : sg_car -> sg_car -> sg_car;
+}.
+
+Structure group := Something {
+ group_car :> Type;
+ group_op : group_car -> group_car -> group_car;
+ group_neg : group_car -> group_car;
+ group_neg_op' x y : group_neg (group_op x y) = group_op (group_neg x) (group_neg y)
+}.
+
+Coercion group_sg (X : group) : semigroup :=
+ SemiGroup (group_car X) (group_op X).
+Canonical Structure group_sg.
+
+Axiom group_neg_op : forall (X : group) (x y : X),
+ group_neg X (sg_op (group_sg X) x y) = sg_op (group_sg X) (group_neg X x) (group_neg X y).
+
+Canonical Structure Z_sg := SemiGroup Z Z.add .
+Canonical Structure Z_group := Something Z Z.add Z.opp Z.opp_add_distr.
+
+Lemma foo (x y : Z) :
+ sg_op Z_sg (group_neg Z_group x) (group_neg Z_group y) =
+ group_neg Z_group (sg_op Z_sg x y).
+Proof.
+ rewrite -group_neg_op.
+ reflexivity.
+Qed.
+
+End Test3.