| author | wenzelm | 
| Thu, 08 Nov 2012 20:25:48 +0100 | |
| changeset 50078 | 02aa7f6e530d | 
| parent 48992 | 0518bf89c777 | 
| child 51686 | 532e0ac5a66d | 
| permissions | -rw-r--r-- | 
| 27340 | 1 | (* Title: Pure/ML/ml_antiquote.ML | 
| 2 | Author: Makarius | |
| 3 | ||
| 4 | Common ML antiquotations. | |
| 5 | *) | |
| 6 | ||
| 7 | signature ML_ANTIQUOTE = | |
| 8 | sig | |
| 9 | val variant: string -> Proof.context -> string * Proof.context | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 10 | val macro: binding -> Proof.context context_parser -> theory -> theory | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 11 | val inline: binding -> string context_parser -> theory -> theory | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 12 | val declaration: string -> binding -> string context_parser -> theory -> theory | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 13 | val value: binding -> string context_parser -> theory -> theory | 
| 27340 | 14 | end; | 
| 15 | ||
| 16 | structure ML_Antiquote: ML_ANTIQUOTE = | |
| 17 | struct | |
| 18 | ||
| 19 | (** generic tools **) | |
| 20 | ||
| 21 | (* ML names *) | |
| 22 | ||
| 48776 
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
 wenzelm parents: 
48646diff
changeset | 23 | val init_context = ML_Syntax.reserved |> Name.declare "ML_context"; | 
| 
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
 wenzelm parents: 
48646diff
changeset | 24 | |
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
36950diff
changeset | 25 | structure Names = Proof_Data | 
| 27340 | 26 | ( | 
| 27 | type T = Name.context; | |
| 48776 
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
 wenzelm parents: 
48646diff
changeset | 28 | fun init _ = init_context; | 
| 27340 | 29 | ); | 
| 30 | ||
| 31 | fun variant a ctxt = | |
| 32 | let | |
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
36950diff
changeset | 33 | val names = Names.get ctxt; | 
| 43326 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 wenzelm parents: 
42468diff
changeset | 34 | val (b, names') = Name.variant a names; | 
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
36950diff
changeset | 35 | val ctxt' = Names.put names' ctxt; | 
| 27340 | 36 | in (b, ctxt') end; | 
| 37 | ||
| 38 | ||
| 39 | (* specific antiquotations *) | |
| 40 | ||
| 27379 | 41 | fun macro name scan = ML_Context.add_antiq name | 
| 27868 | 42 | (fn _ => scan :|-- (fn ctxt => Scan.depend (fn _ => Scan.succeed | 
| 35019 
1ec0a3ff229e
simplified interface for ML antiquotations, struct_name is always "Isabelle";
 wenzelm parents: 
33519diff
changeset | 43 |     (Context.Proof ctxt, fn background => (K ("", ""), background)))));
 | 
| 27379 | 44 | |
| 27340 | 45 | fun inline name scan = ML_Context.add_antiq name | 
| 35019 
1ec0a3ff229e
simplified interface for ML antiquotations, struct_name is always "Isabelle";
 wenzelm parents: 
33519diff
changeset | 46 |   (fn _ => scan >> (fn s => fn background => (K ("", s), background)));
 | 
| 27340 | 47 | |
| 48 | fun declaration kind name scan = ML_Context.add_antiq name | |
| 35019 
1ec0a3ff229e
simplified interface for ML antiquotations, struct_name is always "Isabelle";
 wenzelm parents: 
33519diff
changeset | 49 | (fn _ => scan >> (fn s => fn background => | 
| 27340 | 50 | let | 
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 51 | val (a, background') = | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 52 | variant (translate_string (fn "." => "_" | c => c) (Binding.name_of name)) background; | 
| 27340 | 53 | val env = kind ^ " " ^ a ^ " = " ^ s ^ ";\n"; | 
| 41486 
82c1e348bc18
reverted 08240feb69c7 -- breaks positions of reports;
 wenzelm parents: 
41376diff
changeset | 54 | val body = "Isabelle." ^ a; | 
| 27340 | 55 | in (K (env, body), background') end)); | 
| 56 | ||
| 57 | val value = declaration "val"; | |
| 58 | ||
| 59 | ||
| 60 | ||
| 30231 | 61 | (** misc antiquotations **) | 
| 27340 | 62 | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 63 | val _ = Context.>> (Context.map_theory | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 64 | |
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 65 | (inline (Binding.name "assert") | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 66 | (Scan.succeed "(fn b => if b then () else raise General.Fail \"Assertion failed\")") #> | 
| 40110 | 67 | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 68 | inline (Binding.name "make_string") (Scan.succeed ml_make_string) #> | 
| 36162 
0bd034a80a9a
added ML antiquotation @{make_string}, which produces proper pretty printed version in Poly/ML 5.3.0 or later;
 wenzelm parents: 
35670diff
changeset | 69 | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 70 | value (Binding.name "binding") | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46948diff
changeset | 71 | (Scan.lift (Parse.position Args.name) >> ML_Syntax.make_binding) #> | 
| 27340 | 72 | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 73 | value (Binding.name "theory") | 
| 48927 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48864diff
changeset | 74 | (Args.theory -- Scan.lift (Parse.position Args.name) >> (fn (thy, (name, pos)) => | 
| 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48864diff
changeset | 75 | (Position.report pos (Theory.get_markup (Context.get_theory thy name)); | 
| 
ef462b5558eb
theory def/ref position reports, which enable hyperlinks etc.;
 wenzelm parents: 
48864diff
changeset | 76 | "Context.get_theory (Proof_Context.theory_of ML_context) " ^ ML_Syntax.print_string name)) | 
| 48781 | 77 | || Scan.succeed "Proof_Context.theory_of ML_context") #> | 
| 27340 | 78 | |
| 48781 | 79 | inline (Binding.name "context") (Scan.succeed "Isabelle.ML_context") #> | 
| 27340 | 80 | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 81 | inline (Binding.name "typ") (Args.typ >> (ML_Syntax.atomic o ML_Syntax.print_typ)) #> | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 82 | inline (Binding.name "term") (Args.term >> (ML_Syntax.atomic o ML_Syntax.print_term)) #> | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 83 | inline (Binding.name "prop") (Args.prop >> (ML_Syntax.atomic o ML_Syntax.print_term)) #> | 
| 27340 | 84 | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 85 | macro (Binding.name "let") | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 86 | (Args.context -- | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 87 | Scan.lift | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 88 | (Parse.and_list1 (Parse.and_list1 Args.name_source -- (Args.$$$ "=" |-- Args.name_source))) | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 89 | >> (fn (ctxt, args) => #2 (Proof_Context.match_bind true args ctxt))) #> | 
| 27379 | 90 | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 91 | macro (Binding.name "note") | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 92 | (Args.context :|-- (fn ctxt => | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 93 | Parse.and_list1' (Scan.lift (Args.opt_thm_name I "=") -- Attrib.thms | 
| 47815 | 94 | >> (fn ((a, srcs), ths) => ((a, map (Attrib.attribute_cmd ctxt) srcs), [(ths, [])]))) | 
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 95 | >> (fn args => #2 (Proof_Context.note_thmss "" args ctxt)))) #> | 
| 27379 | 96 | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 97 | value (Binding.name "ctyp") (Args.typ >> (fn T => | 
| 48781 | 98 | "Thm.ctyp_of (Proof_Context.theory_of ML_context) " ^ | 
| 99 | ML_Syntax.atomic (ML_Syntax.print_typ T))) #> | |
| 27340 | 100 | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 101 | value (Binding.name "cterm") (Args.term >> (fn t => | 
| 48781 | 102 | "Thm.cterm_of (Proof_Context.theory_of ML_context) " ^ | 
| 103 | ML_Syntax.atomic (ML_Syntax.print_term t))) #> | |
| 27340 | 104 | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 105 | value (Binding.name "cprop") (Args.prop >> (fn t => | 
| 48781 | 106 | "Thm.cterm_of (Proof_Context.theory_of ML_context) " ^ | 
| 107 | ML_Syntax.atomic (ML_Syntax.print_term t))) #> | |
| 27340 | 108 | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 109 | value (Binding.name "cpat") | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 110 | (Args.context -- | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 111 | Scan.lift Args.name_source >> uncurry Proof_Context.read_term_pattern >> (fn t => | 
| 48781 | 112 | "Thm.cterm_of (Proof_Context.theory_of ML_context) " ^ | 
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 113 | ML_Syntax.atomic (ML_Syntax.print_term t))))); | 
| 27340 | 114 | |
| 115 | ||
| 35396 | 116 | (* type classes *) | 
| 117 | ||
| 35670 | 118 | fun class syn = Args.context -- Scan.lift Args.name_source >> (fn (ctxt, s) => | 
| 42360 | 119 | Proof_Context.read_class ctxt s | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
41486diff
changeset | 120 | |> syn ? Lexicon.mark_class | 
| 35396 | 121 | |> ML_Syntax.print_string); | 
| 122 | ||
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 123 | val _ = Context.>> (Context.map_theory | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 124 | (inline (Binding.name "class") (class false) #> | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 125 | inline (Binding.name "class_syntax") (class true) #> | 
| 35396 | 126 | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 127 | inline (Binding.name "sort") | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 128 | (Args.context -- Scan.lift Args.name_source >> (fn (ctxt, s) => | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 129 | ML_Syntax.atomic (ML_Syntax.print_sort (Syntax.read_sort ctxt s)))))); | 
| 35396 | 130 | |
| 131 | ||
| 35361 
4c7c849b70aa
more orthogonal antiquotations for type constructors;
 wenzelm parents: 
35262diff
changeset | 132 | (* type constructors *) | 
| 27340 | 133 | |
| 36950 | 134 | fun type_name kind check = Args.context -- Scan.lift (Parse.position Args.name_source) | 
| 35401 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 135 | >> (fn (ctxt, (s, pos)) => | 
| 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 136 | let | 
| 42360 | 137 | val Type (c, _) = Proof_Context.read_type_name_proper ctxt false s; | 
| 42468 | 138 | val decl = Type.the_decl (Proof_Context.tsig_of ctxt) (c, pos); | 
| 35401 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 139 | val res = | 
| 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 140 | (case try check (c, decl) of | 
| 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 141 | SOME res => res | 
| 48992 | 142 |         | NONE => error ("Not a " ^ kind ^ ": " ^ quote c ^ Position.here pos));
 | 
| 35401 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 143 | in ML_Syntax.print_string res end); | 
| 27340 | 144 | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 145 | val _ = Context.>> (Context.map_theory | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 146 | (inline (Binding.name "type_name") | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 147 | (type_name "logical type" (fn (c, Type.LogicalType _) => c)) #> | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 148 | inline (Binding.name "type_abbrev") | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 149 | (type_name "type abbreviation" (fn (c, Type.Abbreviation _) => c)) #> | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 150 | inline (Binding.name "nonterminal") | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 151 | (type_name "nonterminal" (fn (c, Type.Nonterminal) => c)) #> | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 152 | inline (Binding.name "type_syntax") | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 153 | (type_name "type" (fn (c, _) => Lexicon.mark_type c)))); | 
| 35361 
4c7c849b70aa
more orthogonal antiquotations for type constructors;
 wenzelm parents: 
35262diff
changeset | 154 | |
| 
4c7c849b70aa
more orthogonal antiquotations for type constructors;
 wenzelm parents: 
35262diff
changeset | 155 | |
| 
4c7c849b70aa
more orthogonal antiquotations for type constructors;
 wenzelm parents: 
35262diff
changeset | 156 | (* constants *) | 
| 27340 | 157 | |
| 36950 | 158 | fun const_name check = Args.context -- Scan.lift (Parse.position Args.name_source) | 
| 35401 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 159 | >> (fn (ctxt, (s, pos)) => | 
| 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 160 | let | 
| 42360 | 161 | val Const (c, _) = Proof_Context.read_const_proper ctxt false s; | 
| 162 | val res = check (Proof_Context.consts_of ctxt, c) | |
| 48992 | 163 | handle TYPE (msg, _, _) => error (msg ^ Position.here pos); | 
| 35401 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 164 | in ML_Syntax.print_string res end); | 
| 27340 | 165 | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 166 | val _ = Context.>> (Context.map_theory | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 167 | (inline (Binding.name "const_name") | 
| 43794 
49cbbe2768a8
sub-structural sharing after Syntax.check phase, with global interning of logical entities (the latter is relevant when bypassing default parsing via YXML);
 wenzelm parents: 
43560diff
changeset | 168 | (const_name (fn (consts, c) => (Consts.the_const consts c; c))) #> | 
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 169 | inline (Binding.name "const_abbrev") | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 170 | (const_name (fn (consts, c) => (Consts.the_abbreviation consts c; c))) #> | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 171 | inline (Binding.name "const_syntax") | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 172 | (const_name (fn (_, c) => Lexicon.mark_const c)) #> | 
| 35401 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 173 | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 174 | inline (Binding.name "syntax_const") | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 175 | (Args.context -- Scan.lift (Parse.position Args.name) >> (fn (ctxt, (c, pos)) => | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 176 | if is_some (Syntax.lookup_const (Proof_Context.syn_of ctxt) c) | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 177 | then ML_Syntax.print_string c | 
| 48992 | 178 |       else error ("Unknown syntax const: " ^ quote c ^ Position.here pos))) #>
 | 
| 27340 | 179 | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 180 | inline (Binding.name "const") | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 181 | (Args.context -- Scan.lift Args.name_source -- Scan.optional | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 182 |         (Scan.lift (Args.$$$ "(") |-- Parse.enum1' "," Args.typ --| Scan.lift (Args.$$$ ")")) []
 | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 183 | >> (fn ((ctxt, raw_c), Ts) => | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 184 | let | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 185 | val Const (c, _) = Proof_Context.read_const_proper ctxt true raw_c; | 
| 46987 | 186 | val consts = Proof_Context.consts_of ctxt; | 
| 187 | val n = length (Consts.typargs consts (c, Consts.type_scheme consts c)); | |
| 188 | val _ = length Ts <> n andalso | |
| 189 |             error ("Constant requires " ^ string_of_int n ^ " type argument(s): " ^
 | |
| 190 |               quote c ^ enclose "(" ")" (commas (replicate n "_")));
 | |
| 191 | val const = Const (c, Consts.instance consts (c, Ts)); | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43326diff
changeset | 192 | in ML_Syntax.atomic (ML_Syntax.print_term const) end)))); | 
| 27340 | 193 | |
| 46948 | 194 | |
| 195 | (* outer syntax *) | |
| 196 | ||
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46948diff
changeset | 197 | fun with_keyword f = | 
| 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46948diff
changeset | 198 | Args.theory -- Scan.lift (Parse.position Parse.string) >> (fn (thy, (name, pos)) => | 
| 48646 
91281e9472d8
more official command specifications, including source position;
 wenzelm parents: 
47815diff
changeset | 199 | (f ((name, Thy_Header.the_keyword thy name), pos) | 
| 48992 | 200 | handle ERROR msg => error (msg ^ Position.here pos))); | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46948diff
changeset | 201 | |
| 46948 | 202 | val _ = Context.>> (Context.map_theory | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46948diff
changeset | 203 | (value (Binding.name "keyword") | 
| 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46948diff
changeset | 204 | (with_keyword | 
| 48646 
91281e9472d8
more official command specifications, including source position;
 wenzelm parents: 
47815diff
changeset | 205 | (fn ((name, NONE), _) => "Parse.$$$ " ^ ML_Syntax.print_string name | 
| 
91281e9472d8
more official command specifications, including source position;
 wenzelm parents: 
47815diff
changeset | 206 | | ((name, SOME _), pos) => | 
| 48992 | 207 |             error ("Expected minor keyword " ^ quote name ^ Position.here pos))) #>
 | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46948diff
changeset | 208 | value (Binding.name "command_spec") | 
| 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46948diff
changeset | 209 | (with_keyword | 
| 48646 
91281e9472d8
more official command specifications, including source position;
 wenzelm parents: 
47815diff
changeset | 210 | (fn ((name, SOME kind), pos) => | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46948diff
changeset | 211 | "Keyword.command_spec " ^ ML_Syntax.atomic | 
| 48646 
91281e9472d8
more official command specifications, including source position;
 wenzelm parents: 
47815diff
changeset | 212 | ((ML_Syntax.print_pair | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46948diff
changeset | 213 | (ML_Syntax.print_pair ML_Syntax.print_string | 
| 48864 
3ee314ae1e0a
added keyword kind "thy_load" (with optional list of file extensions);
 wenzelm parents: 
48781diff
changeset | 214 | (ML_Syntax.print_pair | 
| 
3ee314ae1e0a
added keyword kind "thy_load" (with optional list of file extensions);
 wenzelm parents: 
48781diff
changeset | 215 | (ML_Syntax.print_pair ML_Syntax.print_string | 
| 
3ee314ae1e0a
added keyword kind "thy_load" (with optional list of file extensions);
 wenzelm parents: 
48781diff
changeset | 216 | (ML_Syntax.print_list ML_Syntax.print_string)) | 
| 48646 
91281e9472d8
more official command specifications, including source position;
 wenzelm parents: 
47815diff
changeset | 217 | (ML_Syntax.print_list ML_Syntax.print_string))) | 
| 
91281e9472d8
more official command specifications, including source position;
 wenzelm parents: 
47815diff
changeset | 218 | ML_Syntax.print_position) ((name, kind), pos)) | 
| 
91281e9472d8
more official command specifications, including source position;
 wenzelm parents: 
47815diff
changeset | 219 | | ((name, NONE), pos) => | 
| 48992 | 220 |             error ("Expected command keyword " ^ quote name ^ Position.here pos)))));
 | 
| 46948 | 221 | |
| 27340 | 222 | end; | 
| 223 |