| author | wenzelm | 
| Sun, 30 Aug 2015 21:26:42 +0200 | |
| changeset 61057 | 5f6a1e31f3ad | 
| parent 59378 | 065f349852e6 | 
| child 61143 | 5f898411ce87 | 
| permissions | -rw-r--r-- | 
| 28213 | 1 | theory Setup | 
| 39066 | 2 | imports | 
| 3 | Complex_Main | |
| 59378 
065f349852e6
separate image for prerequisites of codegen tutorial
 haftmann parents: 
59377diff
changeset | 4 | "~~/src/Tools/Permanent_Interpretation" | 
| 46517 | 5 | "~~/src/HOL/Library/Dlist" | 
| 6 | "~~/src/HOL/Library/RBT" | |
| 7 | "~~/src/HOL/Library/Mapping" | |
| 51162 | 8 | "~~/src/HOL/Library/IArray" | 
| 28213 | 9 | begin | 
| 10 | ||
| 48951 
b9238cbcdd41
more standard document preparation within session context;
 wenzelm parents: 
48891diff
changeset | 11 | ML_file "../antiquote_setup.ML" | 
| 
b9238cbcdd41
more standard document preparation within session context;
 wenzelm parents: 
48891diff
changeset | 12 | ML_file "../more_antiquote.ML" | 
| 48891 | 13 | |
| 59377 | 14 | setup \<open> | 
| 38503 | 15 | let | 
| 16 | val typ = Simple_Syntax.read_typ; | |
| 17 | in | |
| 56240 | 18 | Sign.del_syntax (Symbol.xsymbolsN, false) | 
| 42293 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 19 |    [("_constrain", typ "logic => type => logic", Mixfix ("_\<Colon>_", [4, 0], 3)),
 | 
| 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 20 |     ("_constrain", typ "prop' => type => prop'", Mixfix ("_\<Colon>_", [4, 0], 3))] #>
 | 
| 56240 | 21 | Sign.add_syntax (Symbol.xsymbolsN, false) | 
| 42293 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 22 |    [("_constrain", typ "logic => type => logic", Mixfix ("_ \<Colon>  _", [4, 0], 3)),
 | 
| 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 23 |     ("_constrain", typ "prop' => type => prop'", Mixfix ("_ \<Colon> _", [4, 0], 3))]
 | 
| 38503 | 24 | end | 
| 59377 | 25 | \<close> | 
| 38503 | 26 | |
| 59324 | 27 | declare [[default_code_width = 74]] | 
| 34071 | 28 | |
| 42669 
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
 wenzelm parents: 
42358diff
changeset | 29 | declare [[names_unique = false]] | 
| 28213 | 30 | |
| 31 | end |