author | wenzelm |
Sun, 30 Jan 2011 13:02:18 +0100 | |
changeset 41648 | 6d736d983d5c |
parent 41589 | bbd861837ebc |
child 44367 | 74c08021ab2e |
permissions | -rw-r--r-- |
14620
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
skalberg
parents:
14516
diff
changeset
|
1 |
(* Title: HOL/Import/Generate-HOL/GenHOL4Base.thy |
41589 | 2 |
Author: Sebastian Skalberg, TU Muenchen |
14620
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
skalberg
parents:
14516
diff
changeset
|
3 |
*) |
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
skalberg
parents:
14516
diff
changeset
|
4 |
|
16417 | 5 |
theory GenHOL4Base imports "../HOL4Compat" "../HOL4Syntax" begin; |
14516 | 6 |
|
7 |
import_segment "hol4"; |
|
8 |
||
9 |
setup_dump "../HOL" "HOL4Base"; |
|
10 |
||
17566 | 11 |
append_dump {*theory HOL4Base imports "../HOL4Compat" "../HOL4Syntax" begin*}; |
14516 | 12 |
|
13 |
import_theory bool; |
|
14 |
||
15 |
const_maps |
|
16 |
T > True |
|
17 |
F > False |
|
18 |
"!" > All |
|
38795
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
37678
diff
changeset
|
19 |
"/\\" > HOL.conj |
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
37678
diff
changeset
|
20 |
"\\/" > HOL.disj |
14516 | 21 |
"?" > Ex |
22 |
"?!" > Ex1 |
|
23 |
"~" > Not |
|
17188 | 24 |
COND > HOL.If |
14516 | 25 |
bool_case > Datatype.bool.bool_case |
26 |
ONE_ONE > HOL4Setup.ONE_ONE |
|
17628 | 27 |
ONTO > Fun.surj |
14516 | 28 |
TYPE_DEFINITION > HOL4Setup.TYPE_DEFINITION |
29 |
LET > HOL4Compat.LET; |
|
30 |
||
31 |
ignore_thms |
|
32 |
BOUNDED_DEF |
|
33 |
BOUNDED_THM |
|
34 |
UNBOUNDED_DEF |
|
35 |
UNBOUNDED_THM; |
|
36 |
||
37 |
end_import; |
|
38 |
||
39 |
import_theory combin; |
|
40 |
||
41 |
const_maps |
|
42 |
o > Fun.comp; |
|
43 |
||
44 |
end_import; |
|
45 |
||
46 |
import_theory sum; |
|
47 |
||
48 |
type_maps |
|
49 |
sum > "+"; |
|
50 |
||
51 |
const_maps |
|
15647 | 52 |
INL > Sum_Type.Inl |
53 |
INR > Sum_Type.Inr |
|
14516 | 54 |
ISL > HOL4Compat.ISL |
55 |
ISR > HOL4Compat.ISR |
|
56 |
OUTL > HOL4Compat.OUTL |
|
57 |
OUTR > HOL4Compat.OUTR |
|
58 |
sum_case > Datatype.sum.sum_case; |
|
59 |
||
60 |
ignore_thms |
|
61 |
sum_TY_DEF |
|
62 |
sum_ISO_DEF |
|
63 |
IS_SUM_REP |
|
64 |
INL_DEF |
|
65 |
INR_DEF |
|
66 |
sum_axiom |
|
67 |
sum_Axiom; |
|
68 |
||
69 |
end_import; |
|
70 |
||
71 |
import_theory one; |
|
72 |
||
73 |
type_maps |
|
74 |
one > Product_Type.unit; |
|
75 |
||
76 |
const_maps |
|
77 |
one > Product_Type.Unity; |
|
78 |
||
79 |
ignore_thms |
|
80 |
one_TY_DEF |
|
81 |
one_axiom |
|
82 |
one_Axiom |
|
83 |
one_DEF; |
|
84 |
||
85 |
end_import; |
|
86 |
||
87 |
import_theory option; |
|
88 |
||
89 |
type_maps |
|
30235
58d147683393
Made Option a separate theory and renamed option_map to Option.map
nipkow
parents:
24996
diff
changeset
|
90 |
option > Option.option; |
14516 | 91 |
|
92 |
const_maps |
|
30235
58d147683393
Made Option a separate theory and renamed option_map to Option.map
nipkow
parents:
24996
diff
changeset
|
93 |
NONE > Option.option.None |
58d147683393
Made Option a separate theory and renamed option_map to Option.map
nipkow
parents:
24996
diff
changeset
|
94 |
SOME > Option.option.Some |
58d147683393
Made Option a separate theory and renamed option_map to Option.map
nipkow
parents:
24996
diff
changeset
|
95 |
option_case > Option.option.option_case |
58d147683393
Made Option a separate theory and renamed option_map to Option.map
nipkow
parents:
24996
diff
changeset
|
96 |
OPTION_MAP > Option.map |
58d147683393
Made Option a separate theory and renamed option_map to Option.map
nipkow
parents:
24996
diff
changeset
|
97 |
THE > Option.the |
14516 | 98 |
IS_SOME > HOL4Compat.IS_SOME |
99 |
IS_NONE > HOL4Compat.IS_NONE |
|
100 |
OPTION_JOIN > HOL4Compat.OPTION_JOIN; |
|
101 |
||
102 |
ignore_thms |
|
103 |
option_axiom |
|
104 |
option_Axiom |
|
105 |
option_TY_DEF |
|
106 |
option_REP_ABS_DEF |
|
107 |
SOME_DEF |
|
108 |
NONE_DEF; |
|
109 |
||
110 |
end_import; |
|
111 |
||
112 |
import_theory marker; |
|
113 |
end_import; |
|
114 |
||
115 |
import_theory relation; |
|
116 |
||
117 |
const_renames |
|
118 |
reflexive > pred_reflexive; |
|
119 |
||
120 |
end_import; |
|
121 |
||
122 |
import_theory pair; |
|
123 |
||
124 |
type_maps |
|
37678
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
haftmann
parents:
37391
diff
changeset
|
125 |
prod > Product_Type.prod; |
14516 | 126 |
|
127 |
const_maps |
|
128 |
"," > Pair |
|
129 |
FST > fst |
|
130 |
SND > snd |
|
131 |
CURRY > curry |
|
132 |
UNCURRY > split |
|
40607 | 133 |
"##" > map_pair |
14516 | 134 |
pair_case > split; |
135 |
||
136 |
ignore_thms |
|
137 |
prod_TY_DEF |
|
138 |
MK_PAIR_DEF |
|
139 |
IS_PAIR_DEF |
|
140 |
ABS_REP_prod |
|
141 |
COMMA_DEF; |
|
142 |
||
143 |
end_import; |
|
144 |
||
145 |
import_theory num; |
|
146 |
||
147 |
type_maps |
|
148 |
num > nat; |
|
149 |
||
150 |
const_maps |
|
151 |
SUC > Suc |
|
152 |
0 > 0 :: nat; |
|
153 |
||
154 |
ignore_thms |
|
155 |
num_TY_DEF |
|
156 |
num_ISO_DEF |
|
157 |
IS_NUM_REP |
|
158 |
ZERO_REP_DEF |
|
159 |
SUC_REP_DEF |
|
160 |
ZERO_DEF |
|
161 |
SUC_DEF; |
|
162 |
||
163 |
end_import; |
|
164 |
||
165 |
import_theory prim_rec; |
|
166 |
||
167 |
const_maps |
|
35092
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents:
34974
diff
changeset
|
168 |
"<" > Orderings.less :: "[nat,nat]=>bool"; |
14516 | 169 |
|
170 |
end_import; |
|
171 |
||
172 |
import_theory arithmetic; |
|
173 |
||
174 |
const_maps |
|
175 |
ALT_ZERO > HOL4Compat.ALT_ZERO |
|
176 |
NUMERAL_BIT1 > HOL4Compat.NUMERAL_BIT1 |
|
177 |
NUMERAL_BIT2 > HOL4Compat.NUMERAL_BIT2 |
|
178 |
NUMERAL > HOL4Compat.NUMERAL |
|
179 |
num_case > Nat.nat.nat_case |
|
180 |
">" > HOL4Compat.nat_gt |
|
181 |
">=" > HOL4Compat.nat_ge |
|
182 |
FUNPOW > HOL4Compat.FUNPOW |
|
35092
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents:
34974
diff
changeset
|
183 |
"<=" > Orderings.less_eq :: "[nat,nat]=>bool" |
35267
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
35092
diff
changeset
|
184 |
"+" > Groups.plus :: "[nat,nat]=>nat" |
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
35092
diff
changeset
|
185 |
"*" > Groups.times :: "[nat,nat]=>nat" |
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
35092
diff
changeset
|
186 |
"-" > Groups.minus :: "[nat,nat]=>nat" |
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
30235
diff
changeset
|
187 |
MIN > Orderings.min :: "[nat,nat]=>nat" |
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
30235
diff
changeset
|
188 |
MAX > Orderings.max :: "[nat,nat]=>nat" |
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
30235
diff
changeset
|
189 |
DIV > Divides.div :: "[nat,nat]=>nat" |
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
30235
diff
changeset
|
190 |
MOD > Divides.mod :: "[nat,nat]=>nat" |
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
30235
diff
changeset
|
191 |
EXP > Power.power :: "[nat,nat]=>nat"; |
14516 | 192 |
|
193 |
end_import; |
|
194 |
||
195 |
import_theory hrat; |
|
196 |
end_import; |
|
197 |
||
198 |
import_theory hreal; |
|
199 |
end_import; |
|
200 |
||
201 |
import_theory numeral; |
|
202 |
end_import; |
|
203 |
||
204 |
import_theory ind_type; |
|
205 |
end_import; |
|
206 |
||
207 |
import_theory divides; |
|
208 |
||
209 |
const_maps |
|
23686 | 210 |
divides > Divides.times_class.dvd :: "[nat,nat]=>bool"; |
14516 | 211 |
|
212 |
end_import; |
|
213 |
||
214 |
import_theory prime; |
|
215 |
end_import; |
|
216 |
||
217 |
import_theory list; |
|
218 |
||
219 |
type_maps |
|
220 |
list > List.list; |
|
221 |
||
222 |
const_maps |
|
223 |
CONS > List.list.Cons |
|
224 |
NIL > List.list.Nil |
|
225 |
list_case > List.list.list_case |
|
226 |
NULL > List.null |
|
227 |
HD > List.hd |
|
228 |
TL > List.tl |
|
229 |
MAP > List.map |
|
230 |
MEM > "List.op mem" |
|
231 |
FILTER > List.filter |
|
232 |
FOLDL > List.foldl |
|
233 |
EVERY > List.list_all |
|
234 |
REVERSE > List.rev |
|
235 |
LAST > List.last |
|
236 |
FRONT > List.butlast |
|
23029 | 237 |
APPEND > List.append |
14516 | 238 |
FLAT > List.concat |
239 |
LENGTH > Nat.size |
|
240 |
REPLICATE > List.replicate |
|
241 |
list_size > HOL4Compat.list_size |
|
242 |
SUM > HOL4Compat.sum |
|
243 |
FOLDR > HOL4Compat.FOLDR |
|
244 |
EXISTS > HOL4Compat.list_exists |
|
245 |
MAP2 > HOL4Compat.map2 |
|
246 |
ZIP > HOL4Compat.ZIP |
|
247 |
UNZIP > HOL4Compat.unzip; |
|
248 |
||
249 |
ignore_thms |
|
250 |
list_TY_DEF |
|
251 |
list_repfns |
|
252 |
list0_def |
|
253 |
list1_def |
|
254 |
NIL |
|
255 |
CONS_def; |
|
256 |
||
257 |
end_import; |
|
258 |
||
259 |
import_theory pred_set; |
|
260 |
end_import; |
|
261 |
||
262 |
import_theory operator; |
|
263 |
end_import; |
|
264 |
||
265 |
import_theory rich_list; |
|
266 |
end_import; |
|
267 |
||
268 |
import_theory state_transformer; |
|
269 |
end_import; |
|
270 |
||
271 |
append_dump "end"; |
|
272 |
||
273 |
flush_dump; |
|
274 |
||
275 |
import_segment ""; |
|
276 |
||
277 |
end |