| author | hoelzl | 
| Thu, 26 May 2011 14:12:03 +0200 | |
| changeset 42986 | 11fd8c04ea24 | 
| parent 35267 | 8dfd816713c6 | 
| child 44064 | 5bce8ff0d9ae | 
| 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/GenHOL4Real.thy | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14516diff
changeset | 2 | Author: Sebastian Skalberg (TU Muenchen) | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14516diff
changeset | 3 | *) | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14516diff
changeset | 4 | |
| 16417 | 5 | theory GenHOL4Real imports GenHOL4Base begin | 
| 14516 | 6 | |
| 7 | import_segment "hol4"; | |
| 8 | ||
| 9 | setup_dump "../HOL" "HOL4Real"; | |
| 10 | ||
| 17566 | 11 | append_dump "theory HOL4Real imports HOL4Base begin"; | 
| 14516 | 12 | |
| 13 | import_theory realax; | |
| 14 | ||
| 15 | type_maps | |
| 16 | real > RealDef.real; | |
| 17 | ||
| 18 | const_maps | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 19 | real_0 > Groups.zero :: real | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 20 | real_1 > Groups.one :: real | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 21 | real_neg > Groups.uminus :: "real => real" | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 22 | inv > Groups.inverse :: "real => real" | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 23 | real_add > Groups.plus :: "[real,real] => real" | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 24 | real_mul > Groups.times :: "[real,real] => real" | 
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
34974diff
changeset | 25 | real_lt > Orderings.less :: "[real,real] => bool"; | 
| 14516 | 26 | |
| 27 | ignore_thms | |
| 28 | real_TY_DEF | |
| 29 | real_tybij | |
| 30 | real_0 | |
| 31 | real_1 | |
| 32 | real_neg | |
| 33 | real_inv | |
| 34 | real_add | |
| 35 | real_mul | |
| 36 | real_lt | |
| 37 | real_of_hreal | |
| 38 | hreal_of_real | |
| 39 | REAL_ISO_EQ | |
| 40 | REAL_POS | |
| 41 | SUP_ALLPOS_LEMMA1 | |
| 42 | SUP_ALLPOS_LEMMA2 | |
| 43 | SUP_ALLPOS_LEMMA3 | |
| 44 | SUP_ALLPOS_LEMMA4; | |
| 45 | ||
| 46 | end_import; | |
| 47 | ||
| 48 | import_theory real; | |
| 49 | ||
| 50 | const_maps | |
| 51 | real_gt > HOL4Compat.real_gt | |
| 52 | real_ge > HOL4Compat.real_ge | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
34974diff
changeset | 53 | real_lte > Orderings.less_eq :: "[real,real] => bool" | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 54 | real_sub > Groups.minus :: "[real,real] => real" | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 55 | "/" > Rings.divide :: "[real,real] => real" | 
| 34974 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: 
24996diff
changeset | 56 | pow > Power.power :: "[real,nat] => real" | 
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
34974diff
changeset | 57 | abs > Groups.abs :: "real => real" | 
| 15647 | 58 | real_of_num > RealDef.real :: "nat => real"; | 
| 14516 | 59 | |
| 60 | end_import; | |
| 61 | ||
| 62 | import_theory topology; | |
| 63 | end_import; | |
| 64 | ||
| 65 | import_theory nets; | |
| 66 | end_import; | |
| 67 | ||
| 68 | import_theory seq; | |
| 17694 | 69 | const_renames | 
| 70 | "-->" > "hol4-->"; | |
| 71 | ||
| 14516 | 72 | end_import; | 
| 73 | ||
| 74 | import_theory lim; | |
| 75 | end_import; | |
| 76 | ||
| 77 | import_theory powser; | |
| 78 | end_import; | |
| 79 | ||
| 80 | import_theory transc; | |
| 81 | end_import; | |
| 82 | ||
| 83 | import_theory poly; | |
| 84 | end_import; | |
| 85 | ||
| 86 | append_dump "end"; | |
| 87 | ||
| 88 | flush_dump; | |
| 89 | ||
| 90 | import_segment ""; | |
| 91 | ||
| 92 | end |