File: bug_20008.v

package info (click to toggle)
coq 9.1.0%2Bdfsg-3
  • links: PTS, VCS
  • area: main
  • in suites: experimental
  • size: 35,964 kB
  • sloc: ml: 239,908; sh: 4,355; python: 2,985; ansic: 2,644; makefile: 874; lisp: 171; javascript: 63; xml: 24; sed: 2
file content (90 lines) | stat: -rw-r--r-- 1,428 bytes parent folder | download
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
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
Set Primitive Projections.

Record R A := mk { p : A -> A -> A }.

Arguments mk {_} _.
Arguments p {_} _ !_ !_.

(* using nat *)

Goal forall x, p (mk (fun _ _ => 2)) x x = 0.
Proof.
  Fail progress cbn.
Abort.

Goal forall x, p (mk (fun _ _ => 2)) 3 x = 0.
Proof.
  Fail progress cbn.
Abort.

Goal p (mk (fun _ _ => 2)) 3 4 = 0.
Proof.
  progress cbn.
Abort.

Arguments p {_} _ !_ _.

Goal forall x, p (mk (fun _ _ => 2)) x x = 0.
Proof.
  Fail progress cbn.
Abort.

Goal forall x, p (mk (fun _ _ => 2)) 3 x = 0.
Proof.
  progress cbn.
Abort.

Goal forall x, p (mk (fun _ _ => 2)) x 4 = 0.
Proof.
  Fail progress cbn.
Abort.

Goal p (mk (fun _ _ => 2)) 3 4 = 0.
Proof.
  progress cbn.
Abort.

(* using prim int *)
Require Import PrimInt63.

Open Scope int63_scope.

Arguments p {_} _ !_ !_.

Goal forall x, p (mk (fun _ _ => 2)) x x = 0.
Proof.
  intros x. Check x : int. (* sanity check that the scope is correct *)
  Fail progress cbn.
Abort.

Goal forall x, p (mk (fun _ _ => 2)) 3 x = 0.
Proof.
  Fail progress cbn.
Abort.

Goal p (mk (fun _ _ => 2)) 3 4 = 0.
Proof.
  progress cbn.
Abort.

Arguments p {_} _ !_ _.

Goal forall x, p (mk (fun _ _ => 2)) x x = 0.
Proof.
  Fail progress cbn.
Abort.

Goal forall x, p (mk (fun _ _ => 2)) 3 x = 0.
Proof.
  progress cbn.
Abort.

Goal forall x, p (mk (fun _ _ => 2)) x 4 = 0.
Proof.
  Fail progress cbn.
Abort.

Goal p (mk (fun _ _ => 2)) 3 4 = 0.
Proof.
  progress cbn.
Abort.