1 (* Title: Pure/markup_kind.ML |
1 (* Title: Pure/markup_kind.ML |
2 Author: Makarius |
2 Author: Makarius |
3 |
3 |
4 Formally defined kind for Markup.expression. |
4 Formally defined kind for Markup.notation and Markup.expression. |
5 *) |
5 *) |
6 |
6 |
7 signature MARKUP_KIND = |
7 signature MARKUP_KIND = |
8 sig |
8 sig |
|
9 val get_notation_kinds: Proof.context -> string list |
|
10 val check_notation_kind: Proof.context -> xstring * Position.T -> string |
|
11 val setup_notation_kind: binding -> theory -> string * theory |
|
12 val check_notation: Proof.context -> xstring * Position.T -> Markup.T |
|
13 val setup_notation: binding -> Markup.T |
|
14 val get_expression_kinds: Proof.context -> string list |
9 val check_expression_kind: Proof.context -> xstring * Position.T -> string |
15 val check_expression_kind: Proof.context -> xstring * Position.T -> string |
10 val setup_expression_kind: binding -> theory -> string * theory |
16 val setup_expression_kind: binding -> theory -> string * theory |
11 val check_expression: Proof.context -> xstring * Position.T -> Markup.T |
17 val check_expression: Proof.context -> xstring * Position.T -> Markup.T |
12 val setup_expression: binding -> Markup.T |
18 val setup_expression: binding -> Markup.T |
|
19 val markup_mixfix: Markup.T |
|
20 val markup_prefix: Markup.T |
|
21 val markup_postfix: Markup.T |
|
22 val markup_infix: Markup.T |
|
23 val markup_binder: Markup.T |
13 val markup_type: Markup.T |
24 val markup_type: Markup.T |
14 val markup_type_application: Markup.T |
25 val markup_type_application: Markup.T |
15 val markup_term: Markup.T |
26 val markup_term: Markup.T |
16 val markup_term_application: Markup.T |
27 val markup_term_application: Markup.T |
17 val markup_term_abstraction: Markup.T |
28 val markup_term_abstraction: Markup.T |
25 |
36 |
26 (* theory data *) |
37 (* theory data *) |
27 |
38 |
28 structure Data = Theory_Data |
39 structure Data = Theory_Data |
29 ( |
40 ( |
30 type T = unit Name_Space.table; |
41 type T = unit Name_Space.table * unit Name_Space.table; |
31 val empty : T = Name_Space.empty_table "markup_expression_kind"; |
42 val empty : T = |
32 fun merge data : T = Name_Space.merge_tables data; |
43 (Name_Space.empty_table "markup_notation_kind", |
|
44 Name_Space.empty_table "markup_expression_kind"); |
|
45 fun merge ((tab1, tab2), (tab1', tab2')) : T = |
|
46 (Name_Space.merge_tables (tab1, tab1'), |
|
47 Name_Space.merge_tables (tab2, tab2')); |
33 ); |
48 ); |
34 |
49 |
35 |
50 |
36 (* kind *) |
51 (* kind *) |
37 |
52 |
38 fun check_expression_kind ctxt = |
53 local |
39 Name_Space.check (Context.Proof ctxt) (Data.get (Proof_Context.theory_of ctxt)) #> #1; |
|
40 |
54 |
41 fun setup_expression_kind binding thy = |
55 fun get_kinds which ctxt = |
|
56 which (Data.get (Proof_Context.theory_of ctxt)) |
|
57 |> Name_Space.dest_table |
|
58 |> map #1 |
|
59 |> sort_strings; |
|
60 |
|
61 fun check_kind which ctxt = |
|
62 Name_Space.check (Context.Proof ctxt) (which (Data.get (Proof_Context.theory_of ctxt))) #> #1; |
|
63 |
|
64 fun setup_kind which ap binding thy = |
42 let |
65 let |
43 val context = Name_Space.map_naming (K Name_Space.global_naming) (Context.Theory thy); |
66 val context = Name_Space.map_naming (K Name_Space.global_naming) (Context.Theory thy); |
44 val (name, data') = Name_Space.define context true (binding, ()) (Data.get thy); |
67 val (name, tab') = Name_Space.define context true (binding, ()) (which (Data.get thy)); |
45 in (name, Data.put data' thy) end; |
68 in (name, (Data.map o ap) (K tab') thy) end; |
|
69 |
|
70 in |
|
71 |
|
72 val get_notation_kinds = get_kinds #1; |
|
73 val get_expression_kinds = get_kinds #2; |
|
74 |
|
75 val check_notation_kind = check_kind #1; |
|
76 val check_expression_kind = check_kind #2; |
|
77 |
|
78 val setup_notation_kind = setup_kind #1 apfst; |
|
79 val setup_expression_kind = setup_kind #2 apsnd; |
|
80 |
|
81 end; |
46 |
82 |
47 |
83 |
48 (* markup *) |
84 (* markup *) |
49 |
85 |
|
86 fun check_notation ctxt = check_notation_kind ctxt #> Markup.notation; |
|
87 |
50 fun check_expression ctxt = check_expression_kind ctxt #> Markup.expression; |
88 fun check_expression ctxt = check_expression_kind ctxt #> Markup.expression; |
|
89 |
|
90 fun setup_notation binding = |
|
91 Context.>>> (Context.map_theory_result (setup_notation_kind binding #>> Markup.notation)); |
51 |
92 |
52 fun setup_expression binding = |
93 fun setup_expression binding = |
53 Context.>>> (Context.map_theory_result (setup_expression_kind binding #>> Markup.expression)); |
94 Context.>>> (Context.map_theory_result (setup_expression_kind binding #>> Markup.expression)); |
54 |
95 |
55 |
96 |
56 (* concrete markup *) |
97 (* concrete markup *) |
|
98 |
|
99 val markup_mixfix = setup_notation (Binding.make ("mixfix", \<^here>)); |
|
100 val markup_prefix = setup_notation (Binding.make ("prefix", \<^here>)); |
|
101 val markup_postfix = setup_notation (Binding.make ("postfix", \<^here>)); |
|
102 val markup_infix = setup_notation (Binding.make ("infix", \<^here>)); |
|
103 val markup_binder = setup_notation (Binding.make ("binder", \<^here>)); |
57 |
104 |
58 val markup_type = setup_expression (Binding.make ("type", \<^here>)); |
105 val markup_type = setup_expression (Binding.make ("type", \<^here>)); |
59 val markup_type_application = setup_expression (Binding.make ("type_application", \<^here>)); |
106 val markup_type_application = setup_expression (Binding.make ("type_application", \<^here>)); |
60 |
107 |
61 val markup_term = setup_expression (Binding.make ("term", \<^here>)); |
108 val markup_term = setup_expression (Binding.make ("term", \<^here>)); |