| 1478 |      1 | (*  Title:      ZF/ex/ListN
 | 
| 515 |      2 |     ID:         $Id$
 | 
| 1478 |      3 |     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 | 
| 515 |      4 |     Copyright   1994  University of Cambridge
 | 
|  |      5 | 
 | 
|  |      6 | Inductive definition of lists of n elements
 | 
|  |      7 | 
 | 
|  |      8 | See Ch. Paulin-Mohring, Inductive Definitions in the System Coq.
 | 
|  |      9 | Research Report 92-49, LIP, ENS Lyon.  Dec 1992.
 | 
|  |     10 | *)
 | 
|  |     11 | 
 | 
|  |     12 | ListN = List +
 | 
| 1478 |     13 | consts  listn ::i=>i
 | 
| 515 |     14 | inductive
 | 
|  |     15 |   domains   "listn(A)" <= "nat*list(A)"
 | 
|  |     16 |   intrs
 | 
|  |     17 |     NilI  "<0,Nil> : listn(A)"
 | 
|  |     18 |     ConsI "[| a: A;  <n,l> : listn(A) |] ==> <succ(n), Cons(a,l)> : listn(A)"
 | 
| 3841 |     19 |   type_intrs "nat_typechecks @ list.intrs"
 | 
| 515 |     20 | end
 |