diff options
| author | Enrico Tassi | 2020-06-13 21:17:17 +0200 |
|---|---|---|
| committer | Enrico Tassi | 2020-06-15 10:59:53 +0200 |
| commit | 8cbb01a852c950f044c6f2c9a9ada2626d2bbfc5 (patch) | |
| tree | ba5e13815381882c2c27cc3eb24f3edc53486ba4 | |
| parent | 73c3dd110402f182229d8a21458758424256d64c (diff) | |
[ssr] fix env handling in error message (fix #12507)
| -rw-r--r-- | plugins/ssr/ssrequality.ml | 10 | ||||
| -rw-r--r-- | test-suite/ssr/rewrtite_err_msg.v | 30 |
2 files changed, 37 insertions, 3 deletions
diff --git a/plugins/ssr/ssrequality.ml b/plugins/ssr/ssrequality.ml index e25fc1100d..da623703a2 100644 --- a/plugins/ssr/ssrequality.ml +++ b/plugins/ssr/ssrequality.ml @@ -394,12 +394,16 @@ let pirrel_rewrite ?(under=false) ?(map_redex=id_map_redex) pred rdx rdx_ty new_ let hd_ty, miss = match EConstr.kind sigma c with | App (hd, args) -> let hd_ty = Retyping.get_type_of env sigma hd in - let names = let rec aux t = function 0 -> [] | n -> + let names = let rec aux env t = function 0 -> [] | n -> let t = Reductionops.whd_all env sigma t in let open EConstr in match kind_of_type sigma t with - | ProdType (name, _, t) -> name.binder_name :: aux t (n-1) - | _ -> assert false in aux hd_ty (Array.length args) in + | ProdType (name, ty, t) -> + name.binder_name :: + aux (EConstr.push_rel (Context.Rel.Declaration.LocalAssum (name,ty)) env) t (n-1) + | _ -> + (* In the case the head is an HO constant it may accept more arguments *) + CList.init n (fun _ -> Names.Name.Anonymous) in aux env hd_ty (Array.length args) in hd_ty, Util.List.map_filter (fun (t, name) -> let evs = Evar.Set.elements (Evarutil.undefined_evars_of_term sigma t) in let open_evs = List.filter (fun k -> diff --git a/test-suite/ssr/rewrtite_err_msg.v b/test-suite/ssr/rewrtite_err_msg.v new file mode 100644 index 0000000000..d749fd264b --- /dev/null +++ b/test-suite/ssr/rewrtite_err_msg.v @@ -0,0 +1,30 @@ +Require Import ssreflect ssrbool. + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Axiom finGroupType : Type. +Axiom group : finGroupType -> Type. +Axiom abelian : forall gT : finGroupType, group gT -> Prop. +Arguments abelian {_} _. +Axiom carrier : finGroupType -> Type. +Coercion carrier : finGroupType >-> Sortclass. +Axiom mem : forall gT : finGroupType, gT -> group gT -> Prop. +Arguments mem {_} _ _. +Axiom mul : forall gT : finGroupType, gT -> gT -> gT. +Arguments mul {_} _ _. +Definition centralised gT (G : group gT) (x : gT) := forall y, mul x y = mul y x. +Arguments centralised {gT} _. +Axiom b : bool. + +Axiom centsP : forall (gT : finGroupType) (A B : group gT), + reflect (forall a, mem a A -> centralised B a) b. +Arguments centsP {_ _ _}. + +Lemma commute_abelian (gT : finGroupType) (G : group gT) + (G_abelian : abelian G) (g g' : gT) (gG : mem g G) (g'G : mem g' G) : + mul g g' = mul g' g. +Proof. +Fail rewrite (centsP _). (* fails but without an anomaly *) +Abort.
\ No newline at end of file |
