| author | haftmann | 
| Tue, 20 Mar 2018 09:27:39 +0000 | |
| changeset 67905 | fe0f4eeceeb7 | 
| parent 61419 | 3c3f8b182e4b | 
| child 69597 | ff784d5a5bfb | 
| 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  | 
|
| 59324 | 8  | 
declare [[default_code_width = 74]]  | 
| 28565 | 9  | 
|
10  | 
syntax  | 
|
11  | 
  "_alpha" :: "type"  ("\<alpha>")
 | 
|
| 61076 | 12  | 
  "_alpha_ofsort" :: "sort \<Rightarrow> type"  ("\<alpha>()::_" [0] 1000)
 | 
| 28565 | 13  | 
  "_beta" :: "type"  ("\<beta>")
 | 
| 61076 | 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] =  | 
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
41413 
diff
changeset
 | 
21  | 
          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
 | 
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] =  | 
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
41413 
diff
changeset
 | 
26  | 
          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
 | 
27  | 
      | beta_ofsort_ast_tr asts = raise Ast.AST ("beta_ast_tr", asts);
 | 
| 35113 | 28  | 
in  | 
| 52143 | 29  | 
   [(@{syntax_const "_alpha"}, K alpha_ast_tr),
 | 
30  | 
    (@{syntax_const "_alpha_ofsort"}, K alpha_ofsort_ast_tr),
 | 
|
31  | 
    (@{syntax_const "_beta"}, K beta_ast_tr),
 | 
|
32  | 
    (@{syntax_const "_beta_ofsort"}, 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  |