File: bug_2255.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 (22 lines) | stat: -rw-r--r-- 769 bytes parent folder | download | duplicates (5)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
(* Check injection in presence of dependencies hidden in applicative terms *)

Inductive TupleT : nat -> Type :=
  nilT : TupleT 0
| consT {n} A : (A -> TupleT n) -> TupleT (S n).

Inductive Tuple : forall n, TupleT n -> Type :=
  nil : Tuple _ nilT
| cons {n} A (x : A) (F : A -> TupleT n) : Tuple _ (F x) -> Tuple _ (consT A F).

Goal forall n A F x X n0 A0 x0 F0 H0 (H : existT (fun n0 : nat => {H0 : TupleT
n0 & Tuple n0 H0})
         (S n0)
         (existT (fun H0 : TupleT (S n0) => Tuple (S n0) H0)
            (consT A0 F0) (cons A0 x0 F0 H0)) =
       existT (fun n0 : nat => {H0 : TupleT n0 & Tuple n0 H0})
         (S n)
         (existT (fun H0 : TupleT (S n) => Tuple (S n) H0)
            (consT A F) (cons A x F X))), False.
intros.
injection H.
Abort.