File: bug_4782.v

package info (click to toggle)
coq 8.16.1%2Bdfsg-1
  • links: PTS, VCS
  • area: main
  • in suites: bookworm
  • size: 40,596 kB
  • sloc: ml: 219,376; sh: 3,545; python: 3,231; ansic: 2,529; makefile: 767; lisp: 279; javascript: 63; xml: 24; sed: 2
file content (25 lines) | stat: -rw-r--r-- 757 bytes parent folder | download | duplicates (3)
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
(* About typing of with bindings *)

Record r : Type := mk_r { type : Type; cond : type -> Prop }.

Inductive p : Prop := consp : forall (e : r) (x : type e), cond e x -> p.

Goal p.
Fail apply consp with (fun _ : bool => mk_r unit (fun x => True)) nil.
Abort.

(* A simplification of an example from coquelicot, which was failing
   at some time after a fix #4782 was committed. *)

Record T := { dom : Type }.
Definition pairT A B := {| dom := (dom A * dom B)%type |}.
Class C (A:Type).
Parameter B:T.
Instance c (A:T) : C (dom A) := {}.
Instance cn : C (dom B) := {}.
Parameter F : forall A:T, C (dom A) -> forall x:dom A, x=x -> A = A.
Set Typeclasses Debug.
Goal forall (A:T) (x:dom A), pairT A A = pairT A A.
intros.
apply (F _ _)  with (x,x).
Abort.