File: bug_1963.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 (20 lines) | stat: -rw-r--r-- 557 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
(* Check that "dependent inversion" behaves correctly w.r.t to universes *)

Require Import Eqdep.

Set Implicit Arguments.

Inductive illist(A:Type) : nat -> Type :=
  illistn : illist A 0
| illistc : forall n:nat, A -> illist A n -> illist A (S n).

Inductive isig (A:Type)(P:A -> Type) : Type :=
  iexists : forall x : A, P x -> isig P.

Lemma inv : forall (A:Type)(n n':nat)(ts':illist A n'), n' = S n ->
  isig (fun t => isig (fun ts =>
    eq_dep nat (fun n => illist A n) n' ts' (S n) (illistc t ts))).
Proof.
intros.
dependent inversion ts'.
Abort.