author | wenzelm |
Fri, 08 Apr 2011 15:02:11 +0200 | |
changeset 42288 | 2074b31650e6 |
parent 39066 | 4517a4049588 |
child 42293 | 6cca0343ea48 |
permissions | -rw-r--r-- |
28213 | 1 |
theory Setup |
39066 | 2 |
imports |
3 |
Complex_Main |
|
4 |
More_List RBT Dlist Mapping |
|
30227 | 5 |
uses |
6 |
"../../antiquote_setup.ML" |
|
7 |
"../../more_antiquote.ML" |
|
28213 | 8 |
begin |
9 |
||
38503 | 10 |
setup {* |
11 |
let |
|
12 |
val typ = Simple_Syntax.read_typ; |
|
42288
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
wenzelm
parents:
39066
diff
changeset
|
13 |
val typeT = Syntax_Ext.typeT; |
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
wenzelm
parents:
39066
diff
changeset
|
14 |
val spropT = Syntax_Ext.spropT; |
38503 | 15 |
in |
16 |
Sign.del_modesyntax_i (Symbol.xsymbolsN, false) [ |
|
17 |
("_constrain", typ "logic => type => logic", Mixfix ("_\<Colon>_", [4, 0], 3)), |
|
18 |
("_constrain", [spropT, typeT] ---> spropT, Mixfix ("_\<Colon>_", [4, 0], 3))] |
|
19 |
#> Sign.add_modesyntax_i (Symbol.xsymbolsN, false) [ |
|
20 |
("_constrain", typ "logic => type => logic", Mixfix ("_ \<Colon> _", [4, 0], 3)), |
|
21 |
("_constrain", [spropT, typeT] ---> spropT, Mixfix ("_ \<Colon> _", [4, 0], 3))] |
|
22 |
end |
|
23 |
*} |
|
24 |
||
34071 | 25 |
setup {* Code_Target.set_default_code_width 74 *} |
26 |
||
38798
89f273ab1d42
expanded some aliases from structure Unsynchronized;
wenzelm
parents:
38503
diff
changeset
|
27 |
ML_command {* unique_names := false *} |
28213 | 28 |
|
29 |
end |