| author | wenzelm | 
| Sun, 23 Feb 2025 15:40:44 +0100 | |
| changeset 82223 | 706562be40fc | 
| parent 81933 | cb05f8d3fd05 | 
| permissions | -rw-r--r-- | 
| 
47258
 
880e587eee9f
Modernized HOL-Import for HOL Light
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/Import/Import_Setup.thy  | 
| 
 
880e587eee9f
Modernized HOL-Import for HOL Light
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents:  
diff
changeset
 | 
2  | 
Author: Cezary Kaliszyk, University of Innsbruck  | 
| 
 
880e587eee9f
Modernized HOL-Import for HOL Light
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents:  
diff
changeset
 | 
3  | 
Author: Alexander Krauss, QAware GmbH  | 
| 81933 | 4  | 
Author: Makarius  | 
| 
47258
 
880e587eee9f
Modernized HOL-Import for HOL Light
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
880e587eee9f
Modernized HOL-Import for HOL Light
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents:  
diff
changeset
 | 
6  | 
|
| 81829 | 7  | 
section \<open>Importer machinery\<close>  | 
| 
47258
 
880e587eee9f
Modernized HOL-Import for HOL Light
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents:  
diff
changeset
 | 
8  | 
|
| 
 
880e587eee9f
Modernized HOL-Import for HOL Light
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents:  
diff
changeset
 | 
9  | 
theory Import_Setup  | 
| 
65552
 
f533820e7248
theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
 
wenzelm 
parents: 
63167 
diff
changeset
 | 
10  | 
imports Main  | 
| 48891 | 11  | 
keywords "import_type_map" "import_const_map" "import_file" :: thy_decl  | 
| 
47258
 
880e587eee9f
Modernized HOL-Import for HOL Light
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents:  
diff
changeset
 | 
12  | 
begin  | 
| 
 
880e587eee9f
Modernized HOL-Import for HOL Light
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents:  
diff
changeset
 | 
13  | 
|
| 69605 | 14  | 
ML_file \<open>import_data.ML\<close>  | 
| 81904 | 15  | 
|
16  | 
text \<open>  | 
|
17  | 
Initial type and constant maps, for types and constants that are not  | 
|
18  | 
defined, which means their definitions do not appear in the proof dump.  | 
|
19  | 
\<close>  | 
|
20  | 
import_type_map bool : bool  | 
|
21  | 
import_type_map "fun" : "fun"  | 
|
22  | 
import_type_map ind : ind  | 
|
23  | 
import_const_map "=" : HOL.eq  | 
|
24  | 
import_const_map "@" : Eps  | 
|
25  | 
||
| 69605 | 26  | 
ML_file \<open>import_rule.ML\<close>  | 
| 
47258
 
880e587eee9f
Modernized HOL-Import for HOL Light
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents:  
diff
changeset
 | 
27  | 
|
| 
 
880e587eee9f
Modernized HOL-Import for HOL Light
 
Cezary Kaliszyk <cezarykaliszyk@gmail.com> 
parents:  
diff
changeset
 | 
28  | 
end  |