File: bug_13946.v

package info (click to toggle)
coq 8.20.1%2Bdfsg-1
  • links: PTS, VCS
  • area: main
  • in suites: sid, trixie
  • size: 44,116 kB
  • sloc: ml: 234,160; sh: 4,301; python: 3,270; ansic: 2,644; makefile: 882; lisp: 172; javascript: 63; xml: 24; sed: 2
file content (21 lines) | stat: -rw-r--r-- 974 bytes parent folder | download | duplicates (4)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
Succeed Inductive Foo (A:Prop) : Prop := foo : _ tt.

Class ElemOf A B := elem_of: A -> B -> Prop.
Infix "∈" := elem_of (at level 70) .
Inductive Types := tnil: Types | tcons: Type -> Types -> Types.
Implicit Type As: Types.
Infix "^::" := tcons (at level 60, right associativity).
Notation "^[ ]" := tnil (at level 5, format "^[ ]").

Inductive hlist F : Types -> Type :=
| hnil: hlist F ^[]
| hcons {A As} : F A -> hlist F As -> hlist F (A ^:: As).
Infix "+::" := (hcons _) (at level 60, right associativity).

Fail Inductive elem_of_hlist  {F A As} : ElemOf (F A) (hlist F As) :=
  | elem_of_hlist_here F A As (x : F A) (l : hlist F As) : x ∈ (x +:: l)
  | elem_of_list_further F A As (x y : F A) (l : hlist F As) : x ∈ l -> x ∈ (y +:: l).

Fail Inductive elem_of_hlist  {F A As} : ElemOf (F A) (hlist F As) :=
  | elem_of_hlist_here (x : F A) (l : hlist F As) : x ∈ (x +:: l)
  | elem_of_list_further (x y : F A) (l : hlist F As) : x ∈ l -> x ∈ (y +:: l).