author | wenzelm |
Sun, 12 Jan 2025 13:27:47 +0100 | |
changeset 81776 | c6d8db03dfdc |
parent 81631 | 2d4838ffb67e |
permissions | -rw-r--r-- |
81776 | 1 |
(* Title: Pure/PIDE/markup_kind.ML |
80889 | 2 |
Author: Makarius |
3 |
||
80920
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
4 |
Formally defined kind for Markup.notation and Markup.expression. |
80889 | 5 |
*) |
6 |
||
80919 | 7 |
signature MARKUP_KIND = |
80889 | 8 |
sig |
80920
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
9 |
val get_notation_kinds: Proof.context -> string list |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
10 |
val check_notation_kind: Proof.context -> xstring * Position.T -> string |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
11 |
val setup_notation_kind: binding -> theory -> string * theory |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
12 |
val check_notation: Proof.context -> xstring * Position.T -> Markup.T |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
13 |
val setup_notation: binding -> Markup.T |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
14 |
val get_expression_kinds: Proof.context -> string list |
80919 | 15 |
val check_expression_kind: Proof.context -> xstring * Position.T -> string |
16 |
val setup_expression_kind: binding -> theory -> string * theory |
|
17 |
val check_expression: Proof.context -> xstring * Position.T -> Markup.T |
|
18 |
val setup_expression: binding -> Markup.T |
|
80941
fd7a70babec1
more markup <expression kind="item"> in Isabelle/Scala, with pro-forma Markup_Kind.setup in Isabelle/ML;
wenzelm
parents:
80923
diff
changeset
|
19 |
val markup_item: Markup.T |
81631
2d4838ffb67e
more markup: for diagnostic purposes of ambig_msgs;
wenzelm
parents:
81571
diff
changeset
|
20 |
val markup_syntax: Markup.T |
80920
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
21 |
val markup_mixfix: Markup.T |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
22 |
val markup_prefix: Markup.T |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
23 |
val markup_postfix: Markup.T |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
24 |
val markup_infix: Markup.T |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
25 |
val markup_binder: Markup.T |
81125 | 26 |
val markup_pattern: Markup.T |
81124 | 27 |
val markup_type_literal: Markup.T |
81118 | 28 |
val markup_literal: Markup.T |
81571 | 29 |
val markup_index: Markup.T |
80912 | 30 |
val markup_type_application: Markup.T |
80921
a37ed1aeb163
clarified inner syntax markup: use "notation" uniformly;
wenzelm
parents:
80920
diff
changeset
|
31 |
val markup_application: Markup.T |
a37ed1aeb163
clarified inner syntax markup: use "notation" uniformly;
wenzelm
parents:
80920
diff
changeset
|
32 |
val markup_abstraction: Markup.T |
80923 | 33 |
val markup_judgment: Markup.T |
80889 | 34 |
end; |
35 |
||
80919 | 36 |
structure Markup_Kind: MARKUP_KIND = |
80889 | 37 |
struct |
38 |
||
39 |
(* theory data *) |
|
40 |
||
41 |
structure Data = Theory_Data |
|
42 |
( |
|
80920
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
43 |
type T = unit Name_Space.table * unit Name_Space.table; |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
44 |
val empty : T = |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
45 |
(Name_Space.empty_table "markup_notation_kind", |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
46 |
Name_Space.empty_table "markup_expression_kind"); |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
47 |
fun merge ((tab1, tab2), (tab1', tab2')) : T = |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
48 |
(Name_Space.merge_tables (tab1, tab1'), |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
49 |
Name_Space.merge_tables (tab2, tab2')); |
80889 | 50 |
); |
51 |
||
80902 | 52 |
|
53 |
(* kind *) |
|
54 |
||
80920
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
55 |
local |
80889 | 56 |
|
80920
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
57 |
fun get_kinds which ctxt = |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
58 |
which (Data.get (Proof_Context.theory_of ctxt)) |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
59 |
|> Name_Space.dest_table |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
60 |
|> map #1 |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
61 |
|> sort_strings; |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
62 |
|
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
63 |
fun check_kind which ctxt = |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
64 |
Name_Space.check (Context.Proof ctxt) (which (Data.get (Proof_Context.theory_of ctxt))) #> #1; |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
65 |
|
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
66 |
fun setup_kind which ap binding thy = |
80889 | 67 |
let |
68 |
val context = Name_Space.map_naming (K Name_Space.global_naming) (Context.Theory thy); |
|
80920
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
69 |
val (name, tab') = Name_Space.define context true (binding, ()) (which (Data.get thy)); |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
70 |
in (name, (Data.map o ap) (K tab') thy) end; |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
71 |
|
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
72 |
in |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
73 |
|
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
74 |
val get_notation_kinds = get_kinds #1; |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
75 |
val get_expression_kinds = get_kinds #2; |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
76 |
|
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
77 |
val check_notation_kind = check_kind #1; |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
78 |
val check_expression_kind = check_kind #2; |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
79 |
|
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
80 |
val setup_notation_kind = setup_kind #1 apfst; |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
81 |
val setup_expression_kind = setup_kind #2 apsnd; |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
82 |
|
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
83 |
end; |
80889 | 84 |
|
80902 | 85 |
|
86 |
(* markup *) |
|
87 |
||
80920
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
88 |
fun check_notation ctxt = check_notation_kind ctxt #> Markup.notation; |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
89 |
|
80919 | 90 |
fun check_expression ctxt = check_expression_kind ctxt #> Markup.expression; |
80902 | 91 |
|
80920
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
92 |
fun setup_notation binding = |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
93 |
Context.>>> (Context.map_theory_result (setup_notation_kind binding #>> Markup.notation)); |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
94 |
|
80919 | 95 |
fun setup_expression binding = |
96 |
Context.>>> (Context.map_theory_result (setup_expression_kind binding #>> Markup.expression)); |
|
80889 | 97 |
|
80912 | 98 |
|
99 |
(* concrete markup *) |
|
100 |
||
80941
fd7a70babec1
more markup <expression kind="item"> in Isabelle/Scala, with pro-forma Markup_Kind.setup in Isabelle/ML;
wenzelm
parents:
80923
diff
changeset
|
101 |
val markup_item = setup_expression (Binding.make ("item", \<^here>)); |
fd7a70babec1
more markup <expression kind="item"> in Isabelle/Scala, with pro-forma Markup_Kind.setup in Isabelle/ML;
wenzelm
parents:
80923
diff
changeset
|
102 |
|
81631
2d4838ffb67e
more markup: for diagnostic purposes of ambig_msgs;
wenzelm
parents:
81571
diff
changeset
|
103 |
val markup_syntax = setup_expression (Binding.make ("syntax", \<^here>)); |
2d4838ffb67e
more markup: for diagnostic purposes of ambig_msgs;
wenzelm
parents:
81571
diff
changeset
|
104 |
|
80920
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
105 |
val markup_mixfix = setup_notation (Binding.make ("mixfix", \<^here>)); |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
106 |
val markup_prefix = setup_notation (Binding.make ("prefix", \<^here>)); |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
107 |
val markup_postfix = setup_notation (Binding.make ("postfix", \<^here>)); |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
108 |
val markup_infix = setup_notation (Binding.make ("infix", \<^here>)); |
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
109 |
val markup_binder = setup_notation (Binding.make ("binder", \<^here>)); |
81125 | 110 |
val markup_pattern = setup_notation (Binding.make ("pattern", \<^here>)); |
81124 | 111 |
val markup_type_literal = setup_notation (Binding.make ("type_literal", \<^here>)); |
81118 | 112 |
val markup_literal = setup_notation (Binding.make ("literal", \<^here>)); |
81571 | 113 |
val markup_index = setup_notation (Binding.make ("index", \<^here>)); |
80920
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
wenzelm
parents:
80919
diff
changeset
|
114 |
|
80921
a37ed1aeb163
clarified inner syntax markup: use "notation" uniformly;
wenzelm
parents:
80920
diff
changeset
|
115 |
val markup_type_application = setup_notation (Binding.make ("type_application", \<^here>)); |
a37ed1aeb163
clarified inner syntax markup: use "notation" uniformly;
wenzelm
parents:
80920
diff
changeset
|
116 |
val markup_application = setup_notation (Binding.make ("application", \<^here>)); |
a37ed1aeb163
clarified inner syntax markup: use "notation" uniformly;
wenzelm
parents:
80920
diff
changeset
|
117 |
val markup_abstraction = setup_notation (Binding.make ("abstraction", \<^here>)); |
80912 | 118 |
|
80923 | 119 |
val markup_judgment = setup_notation (Binding.make ("judgment", \<^here>)); |
120 |
||
80889 | 121 |
end; |