author | wenzelm |
Fri, 13 Mar 2020 16:12:50 +0100 | |
changeset 71546 | 4dd5dadfc87d |
parent 69605 | a96320074298 |
child 80914 | d97fdabd9e2b |
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 |
||
69605 | 5 |
ML_file \<open>../antiquote_setup.ML\<close> |
6 |
ML_file \<open>../more_antiquote.ML\<close> |
|
48891 | 7 |
|
59324 | 8 |
declare [[default_code_width = 74]] |
28565 | 9 |
|
10 |
syntax |
|
11 |
"_alpha" :: "type" ("\<alpha>") |
|
71546 | 12 |
"_alpha_ofsort" :: "sort \<Rightarrow> type" ("\<alpha>' ::_" [0] 1000) |
28565 | 13 |
"_beta" :: "type" ("\<beta>") |
71546 | 14 |
"_beta_ofsort" :: "sort \<Rightarrow> type" ("\<beta>' ::_" [0] 1000) |
28565 | 15 |
|
61419 | 16 |
parse_ast_translation \<open> |
28565 | 17 |
let |
42224
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents:
41413
diff
changeset
|
18 |
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
|
19 |
| alpha_ast_tr asts = raise Ast.AST ("alpha_ast_tr", asts); |
28565 | 20 |
fun alpha_ofsort_ast_tr [ast] = |
69597 | 21 |
Ast.Appl [Ast.Constant \<^syntax_const>\<open>_ofsort\<close>, Ast.Variable "'a", ast] |
42224
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents:
41413
diff
changeset
|
22 |
| 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
|
23 |
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
|
24 |
| beta_ast_tr asts = raise Ast.AST ("beta_ast_tr", asts); |
28565 | 25 |
fun beta_ofsort_ast_tr [ast] = |
69597 | 26 |
Ast.Appl [Ast.Constant \<^syntax_const>\<open>_ofsort\<close>, Ast.Variable "'b", ast] |
42224
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents:
41413
diff
changeset
|
27 |
| beta_ofsort_ast_tr asts = raise Ast.AST ("beta_ast_tr", asts); |
35113 | 28 |
in |
69597 | 29 |
[(\<^syntax_const>\<open>_alpha\<close>, K alpha_ast_tr), |
30 |
(\<^syntax_const>\<open>_alpha_ofsort\<close>, K alpha_ofsort_ast_tr), |
|
31 |
(\<^syntax_const>\<open>_beta\<close>, K beta_ast_tr), |
|
32 |
(\<^syntax_const>\<open>_beta_ofsort\<close>, K beta_ofsort_ast_tr)] |
|
35113 | 33 |
end |
61419 | 34 |
\<close> |
28565 | 35 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
48985
diff
changeset
|
36 |
end |