| author | wenzelm | 
| Tue, 26 Aug 2008 11:42:46 +0200 | |
| changeset 28000 | ca56bbb99607 | 
| parent 20326 | cbf31171c147 | 
| child 34974 | 18b41bba42b5 | 
| permissions | -rw-r--r-- | 
| 17322 | 1 | (* Title: HOL/Import/Generate-HOLLight/GenHOLLight.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Steven Obua (TU Muenchen) | |
| 4 | *) | |
| 5 | ||
| 6 | theory GenHOLLight imports "../HOLLightCompat" "../HOL4Syntax" begin; | |
| 7 | ||
| 20275 | 8 | (*;skip_import_dir "/home/obua/tmp/cache" | 
| 19064 | 9 | |
| 20275 | 10 | ;skip_import on*) | 
| 19064 | 11 | |
| 17322 | 12 | ;import_segment "hollight"; | 
| 13 | ||
| 14 | setup_dump "../HOLLight" "HOLLight"; | |
| 15 | ||
| 16 | append_dump {*theory HOLLight = "../HOLLightCompat" + "../HOL4Syntax":*};
 | |
| 17 | ||
| 18 | ;import_theory hollight; | |
| 19 | ||
| 20 | ignore_thms | |
| 21 | TYDEF_1 | |
| 22 | DEF_one | |
| 23 | TYDEF_prod | |
| 24 | TYDEF_num | |
| 25 | IND_SUC_0_EXISTS | |
| 26 | DEF__0 | |
| 27 | DEF_SUC | |
| 28 | DEF_IND_SUC | |
| 29 | DEF_IND_0 | |
| 30 | DEF_NUM_REP | |
| 19064 | 31 | TYDEF_sum | 
| 17322 | 32 | DEF_INL | 
| 19203 | 33 | DEF_INR | 
| 20326 | 34 | (* TYDEF_option*); | 
| 17322 | 35 | |
| 36 | type_maps | |
| 37 | ind > Nat.ind | |
| 38 | bool > bool | |
| 39 | fun > fun | |
| 40 | N_1 > Product_Type.unit | |
| 41 | prod > "*" | |
| 19064 | 42 | num > nat | 
| 19203 | 43 | sum > "+" | 
| 20326 | 44 | (* option > Datatype.option*); | 
| 17322 | 45 | |
| 17444 | 46 | const_renames | 
| 47 | "==" > "eqeq" | |
| 17490 | 48 | ".." > "dotdot" | 
| 49 | "ALL" > ALL_list; | |
| 17444 | 50 | |
| 17322 | 51 | const_maps | 
| 52 | T > True | |
| 53 | F > False | |
| 54 | ONE_ONE > HOL4Setup.ONE_ONE | |
| 55 | ONTO > Fun.surj | |
| 56 | "=" > "op =" | |
| 57 | "==>" > "op -->" | |
| 58 | "/\\" > "op &" | |
| 59 | "\\/" > "op |" | |
| 60 | "!" > All | |
| 61 | "?" > Ex | |
| 62 | "?!" > Ex1 | |
| 63 | "~" > Not | |
| 64 | o > Fun.comp | |
| 65 | "@" > "Hilbert_Choice.Eps" | |
| 66 | I > Fun.id | |
| 67 | one > Product_Type.Unity | |
| 68 | LET > HOL4Compat.LET | |
| 69 | mk_pair > Product_Type.Pair_Rep | |
| 70 | "," > Pair | |
| 71 | REP_prod > Rep_Prod | |
| 72 | ABS_prod > Abs_Prod | |
| 73 | FST > fst | |
| 74 | SND > snd | |
| 75 | "_0" > 0 :: nat | |
| 76 | SUC > Suc | |
| 77 | PRE > HOLLightCompat.Pred | |
| 78 | NUMERAL > HOL4Compat.NUMERAL | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
19203diff
changeset | 79 | "+" > HOL.plus :: "nat \<Rightarrow> nat \<Rightarrow> nat" | 
| 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
19203diff
changeset | 80 | "*" > HOL.times :: "nat \<Rightarrow> nat \<Rightarrow> nat" | 
| 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
19203diff
changeset | 81 | "-" > HOL.minus :: "nat \<Rightarrow> nat \<Rightarrow> nat" | 
| 17322 | 82 | BIT0 > HOLLightCompat.NUMERAL_BIT0 | 
| 19064 | 83 | BIT1 > HOL4Compat.NUMERAL_BIT1 | 
| 84 | INL > Sum_Type.Inl | |
| 20326 | 85 | INR > Sum_Type.Inr | 
| 86 | (* NONE > Datatype.None | |
| 87 | SOME > Datatype.Some; | |
| 88 | HAS_SIZE > HOLLightCompat.HAS_SIZE; *) | |
| 17440 
df77edc4f5d0
fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
 obua parents: 
17322diff
changeset | 89 | |
| 
df77edc4f5d0
fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
 obua parents: 
17322diff
changeset | 90 | (*import_until "TYDEF_sum" | 
| 
df77edc4f5d0
fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
 obua parents: 
17322diff
changeset | 91 | |
| 
df77edc4f5d0
fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
 obua parents: 
17322diff
changeset | 92 | consts | 
| 
df77edc4f5d0
fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
 obua parents: 
17322diff
changeset | 93 | print_theorems | 
| 
df77edc4f5d0
fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
 obua parents: 
17322diff
changeset | 94 | |
| 
df77edc4f5d0
fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
 obua parents: 
17322diff
changeset | 95 | import_until *) | 
| 17322 | 96 | |
| 97 | end_import; | |
| 98 | ||
| 99 | append_dump "end"; | |
| 100 | ||
| 101 | flush_dump; | |
| 102 | ||
| 103 | import_segment ""; | |
| 104 | ||
| 105 | end |