| author | wenzelm | 
| Wed, 31 Aug 2005 15:46:36 +0200 | |
| changeset 17199 | 59c1bfc81d91 | 
| parent 16417 | 9bc16273c2d4 | 
| child 17566 | 484ff733f29c | 
| permissions | -rw-r--r-- | 
| 14620 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14516diff
changeset | 1 | (* Title: HOL/Import/Generate-HOL/GenHOL4Vec.thy | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14516diff
changeset | 2 | ID: $Id$ | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14516diff
changeset | 3 | Author: Sebastian Skalberg (TU Muenchen) | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14516diff
changeset | 4 | *) | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14516diff
changeset | 5 | |
| 16417 | 6 | theory GenHOL4Vec imports GenHOL4Base begin | 
| 14516 | 7 | |
| 8 | import_segment "hol4"; | |
| 9 | ||
| 10 | setup_dump "../HOL" "HOL4Vec"; | |
| 11 | ||
| 12 | append_dump "theory HOL4Vec = HOL4Base:"; | |
| 13 | ||
| 14 | import_theory res_quan; | |
| 15 | end_import; | |
| 16 | ||
| 17 | import_theory word_base; | |
| 18 | ||
| 19 | const_renames | |
| 20 | BIT > bit; | |
| 21 | ||
| 22 | end_import; | |
| 23 | ||
| 24 | import_theory word_num; | |
| 25 | end_import; | |
| 26 | ||
| 27 | import_theory word_bitop; | |
| 28 | end_import; | |
| 29 | ||
| 30 | import_theory bword_num; | |
| 31 | end_import; | |
| 32 | ||
| 33 | import_theory bword_arith; | |
| 34 | end_import; | |
| 35 | ||
| 36 | import_theory bword_bitop; | |
| 37 | end_import; | |
| 38 | ||
| 39 | append_dump "end"; | |
| 40 | ||
| 41 | flush_dump; | |
| 42 | ||
| 43 | import_segment ""; | |
| 44 | ||
| 45 | end |