| author | wenzelm | 
| Tue, 27 Jul 1999 22:04:30 +0200 | |
| changeset 7108 | 0229ce6735f6 | 
| parent 71 | 729fe026c5f3 | 
| permissions | -rw-r--r-- | 
| 0 | 1 | (* Title: ZF/ex/enum | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1993 University of Cambridge | |
| 5 | ||
| 6 | Example of a BIG enumeration type | |
| 7 | ||
| 56 | 8 | Can go up to at least 100 constructors, but it takes nearly 7 minutes... | 
| 0 | 9 | *) | 
| 10 | ||
| 11 | ||
| 56 | 12 | (*An enumeration type with 60 contructors! -- takes about 150 seconds!*) | 
| 0 | 13 | fun mk_ids a 0 = [] | 
| 14 | | mk_ids a n = a :: mk_ids (bump_string a) (n-1); | |
| 15 | ||
| 16 | val consts = mk_ids "con1" 60; | |
| 17 | ||
| 18 | structure Enum = Datatype_Fun | |
| 19 | (val thy = Univ.thy; | |
| 20 | val rec_specs = | |
| 21 |       [("enum", "univ(0)",
 | |
| 22 | [(consts, "i")])]; | |
| 23 | val rec_styp = "i"; | |
| 24 | val ext = None | |
| 25 | val sintrs = map (fn const => const ^ " : enum") consts; | |
| 26 | val monos = []; | |
| 71 
729fe026c5f3
sample datatype defs now use datatype_intrs, datatype_elims
 lcp parents: 
56diff
changeset | 27 | val type_intrs = datatype_intrs | 
| 0 | 28 | val type_elims = []); | 
| 29 | ||
| 38 | 30 | goal Enum.thy "con59 ~= con60"; | 
| 56 | 31 | by (simp_tac (ZF_ss addsimps Enum.free_iffs) 1); | 
| 0 | 32 | result(); | 
| 33 |