author | nipkow |
Fri, 21 Jun 2013 09:00:26 +0200 | |
changeset 52402 | c2f30ba4bb98 |
parent 52143 | 36ffe23b25f8 |
child 56059 | 2390391584c2 |
permissions | -rw-r--r-- |
28565 | 1 |
theory Setup |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
48985
diff
changeset
|
2 |
imports Main |
28565 | 3 |
begin |
4 |
||
48950
9965099f51ad
more standard document preparation within session context;
wenzelm
parents:
48891
diff
changeset
|
5 |
ML_file "../antiquote_setup.ML" |
9965099f51ad
more standard document preparation within session context;
wenzelm
parents:
48891
diff
changeset
|
6 |
ML_file "../more_antiquote.ML" |
48891 | 7 |
|
43564
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
42224
diff
changeset
|
8 |
setup {* |
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
42224
diff
changeset
|
9 |
Antiquote_Setup.setup #> |
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
42224
diff
changeset
|
10 |
More_Antiquote.setup #> |
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
42224
diff
changeset
|
11 |
Code_Target.set_default_code_width 74 |
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
42224
diff
changeset
|
12 |
*} |
28565 | 13 |
|
14 |
syntax |
|
15 |
"_alpha" :: "type" ("\<alpha>") |
|
16 |
"_alpha_ofsort" :: "sort \<Rightarrow> type" ("\<alpha>()\<Colon>_" [0] 1000) |
|
17 |
"_beta" :: "type" ("\<beta>") |
|
18 |
"_beta_ofsort" :: "sort \<Rightarrow> type" ("\<beta>()\<Colon>_" [0] 1000) |
|
19 |
||
20 |
parse_ast_translation {* |
|
21 |
let |
|
42224
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents:
41413
diff
changeset
|
22 |
fun alpha_ast_tr [] = Ast.Variable "'a" |
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents:
41413
diff
changeset
|
23 |
| alpha_ast_tr asts = raise Ast.AST ("alpha_ast_tr", asts); |
28565 | 24 |
fun alpha_ofsort_ast_tr [ast] = |
42224
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents:
41413
diff
changeset
|
25 |
Ast.Appl [Ast.Constant @{syntax_const "_ofsort"}, Ast.Variable "'a", ast] |
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents:
41413
diff
changeset
|
26 |
| alpha_ofsort_ast_tr asts = raise Ast.AST ("alpha_ast_tr", asts); |
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents:
41413
diff
changeset
|
27 |
fun beta_ast_tr [] = Ast.Variable "'b" |
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents:
41413
diff
changeset
|
28 |
| beta_ast_tr asts = raise Ast.AST ("beta_ast_tr", asts); |
28565 | 29 |
fun beta_ofsort_ast_tr [ast] = |
42224
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents:
41413
diff
changeset
|
30 |
Ast.Appl [Ast.Constant @{syntax_const "_ofsort"}, Ast.Variable "'b", ast] |
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents:
41413
diff
changeset
|
31 |
| beta_ofsort_ast_tr asts = raise Ast.AST ("beta_ast_tr", asts); |
35113 | 32 |
in |
52143 | 33 |
[(@{syntax_const "_alpha"}, K alpha_ast_tr), |
34 |
(@{syntax_const "_alpha_ofsort"}, K alpha_ofsort_ast_tr), |
|
35 |
(@{syntax_const "_beta"}, K beta_ast_tr), |
|
36 |
(@{syntax_const "_beta_ofsort"}, K beta_ofsort_ast_tr)] |
|
35113 | 37 |
end |
28565 | 38 |
*} |
39 |
||
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
48985
diff
changeset
|
40 |
end |