src/Doc/Classes/Setup.thy
author immler
Sun, 03 Nov 2019 21:46:46 -0500
changeset 71034 e0755162093f
parent 69605 a96320074298
child 71546 4dd5dadfc87d
permissions -rw-r--r--
replace approximation oracle by less ad-hoc @{computation}s
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28565
haftmann
parents:
diff changeset
     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
haftmann
parents:
diff changeset
     3
begin
haftmann
parents:
diff changeset
     4
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69597
diff changeset
     5
ML_file \<open>../antiquote_setup.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69597
diff changeset
     6
ML_file \<open>../more_antiquote.ML\<close>
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 43564
diff changeset
     7
59324
f5f9993a168d prefer option for default code printing width
haftmann
parents: 56061
diff changeset
     8
declare [[default_code_width = 74]]
28565
haftmann
parents:
diff changeset
     9
haftmann
parents:
diff changeset
    10
syntax
haftmann
parents:
diff changeset
    11
  "_alpha" :: "type"  ("\<alpha>")
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 59324
diff changeset
    12
  "_alpha_ofsort" :: "sort \<Rightarrow> type"  ("\<alpha>()::_" [0] 1000)
28565
haftmann
parents:
diff changeset
    13
  "_beta" :: "type"  ("\<beta>")
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 59324
diff changeset
    14
  "_beta_ofsort" :: "sort \<Rightarrow> type"  ("\<beta>()::_" [0] 1000)
28565
haftmann
parents:
diff changeset
    15
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 61076
diff changeset
    16
parse_ast_translation \<open>
28565
haftmann
parents:
diff changeset
    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
haftmann
parents:
diff changeset
    20
    fun alpha_ofsort_ast_tr [ast] =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61419
diff changeset
    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
haftmann
parents:
diff changeset
    25
    fun beta_ofsort_ast_tr [ast] =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61419
diff changeset
    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
1a0c129bb2e0 modernized translations;
wenzelm
parents: 34071
diff changeset
    28
  in
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61419
diff changeset
    29
   [(\<^syntax_const>\<open>_alpha\<close>, K alpha_ast_tr),
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61419
diff changeset
    30
    (\<^syntax_const>\<open>_alpha_ofsort\<close>, K alpha_ofsort_ast_tr),
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61419
diff changeset
    31
    (\<^syntax_const>\<open>_beta\<close>, K beta_ast_tr),
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61419
diff changeset
    32
    (\<^syntax_const>\<open>_beta_ofsort\<close>, K beta_ofsort_ast_tr)]
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 34071
diff changeset
    33
  end
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 61076
diff changeset
    34
\<close>
28565
haftmann
parents:
diff changeset
    35
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48985
diff changeset
    36
end