aboutsummaryrefslogtreecommitdiff
path: root/test-suite/bugs/closed/bug_9512.v
blob: f42e32cf25f588147432cdc854b033a7208d3cd3 (plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
Require Import Coq.ZArith.BinInt Coq.micromega.Lia.

Set Primitive Projections.
Record params := { width : Z }.
Definition p : params := Build_params 64.

Definition width' := width.
Set Printing All.

Lemma foo : width p = 0%Z -> width p = 0%Z.
  intros.

  assert_succeeds (enough True; [lia|]).

(*   H : @eq Z (width p) Z0 *)
(*   ============================ *)
(*   @eq Z (width p) Z0 *)

  change (width' p = 0%Z) in H;cbv [width'] in H.
  (* check that we correctly got the compat constant in H *)
  Fail match goal with H : ?l = _ |- ?l' = _ => constr_eq l l' end.

(*   H : @eq Z (width p) Z0 *)
(*   ============================ *)
(*   @eq Z (width p) Z0 *)

  assert_succeeds (enough True; [lia|]).

  lia.
  (* Tactic failure:  Cannot find witness. *)
Qed.