| author | blanchet | 
| Sun, 01 May 2011 18:37:24 +0200 | |
| changeset 42563 | e70ffe3846d0 | 
| parent 42468 | aea61c5f88c3 | 
| child 43326 | 47cf4bc789aa | 
| 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 | |
| 30513 | 9 | val macro: string -> Proof.context context_parser -> unit | 
| 27340 | 10 | val variant: string -> Proof.context -> string * Proof.context | 
| 30513 | 11 | val inline: string -> string context_parser -> unit | 
| 12 | val declaration: string -> string -> string context_parser -> unit | |
| 13 | val value: string -> string context_parser -> unit | |
| 27340 | 14 | end; | 
| 15 | ||
| 16 | structure ML_Antiquote: ML_ANTIQUOTE = | |
| 17 | struct | |
| 18 | ||
| 19 | (** generic tools **) | |
| 20 | ||
| 21 | (* ML names *) | |
| 22 | ||
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
36950diff
changeset | 23 | structure Names = Proof_Data | 
| 27340 | 24 | ( | 
| 25 | type T = Name.context; | |
| 26 | fun init _ = ML_Syntax.reserved; | |
| 27 | ); | |
| 28 | ||
| 29 | fun variant a ctxt = | |
| 30 | let | |
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
36950diff
changeset | 31 | val names = Names.get ctxt; | 
| 27340 | 32 | val ([b], names') = Name.variants [a] names; | 
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
36950diff
changeset | 33 | val ctxt' = Names.put names' ctxt; | 
| 27340 | 34 | in (b, ctxt') end; | 
| 35 | ||
| 36 | ||
| 37 | (* specific antiquotations *) | |
| 38 | ||
| 27379 | 39 | fun macro name scan = ML_Context.add_antiq name | 
| 27868 | 40 | (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 | 41 |     (Context.Proof ctxt, fn background => (K ("", ""), background)))));
 | 
| 27379 | 42 | |
| 27340 | 43 | 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 | 44 |   (fn _ => scan >> (fn s => fn background => (K ("", s), background)));
 | 
| 27340 | 45 | |
| 46 | 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 | 47 | (fn _ => scan >> (fn s => fn background => | 
| 27340 | 48 | let | 
| 37304 | 49 | val (a, background') = variant (translate_string (fn "." => "_" | c => c) name) background; | 
| 27340 | 50 | val env = kind ^ " " ^ a ^ " = " ^ s ^ ";\n"; | 
| 41486 
82c1e348bc18
reverted 08240feb69c7 -- breaks positions of reports;
 wenzelm parents: 
41376diff
changeset | 51 | val body = "Isabelle." ^ a; | 
| 27340 | 52 | in (K (env, body), background') end)); | 
| 53 | ||
| 54 | val value = declaration "val"; | |
| 55 | ||
| 56 | ||
| 57 | ||
| 30231 | 58 | (** misc antiquotations **) | 
| 27340 | 59 | |
| 40110 | 60 | val _ = inline "assert" | 
| 61 | (Scan.succeed "(fn b => if b then () else raise General.Fail \"Assertion failed\")"); | |
| 62 | ||
| 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 | 63 | val _ = inline "make_string" (Scan.succeed ml_make_string); | 
| 
0bd034a80a9a
added ML antiquotation @{make_string}, which produces proper pretty printed version in Poly/ML 5.3.0 or later;
 wenzelm parents: 
35670diff
changeset | 64 | |
| 30721 
0579dec9f8ba
@{binding} is not a constructor term and should not be inlined, otherwise we loose value polymorphism;
 wenzelm parents: 
30524diff
changeset | 65 | val _ = value "binding" | 
| 36950 | 66 | (Scan.lift (Parse.position Args.name) | 
| 67 | >> (fn name => ML_Syntax.atomic (ML_Syntax.make_binding name))); | |
| 27340 | 68 | |
| 69 | val _ = value "theory" | |
| 37868 
59eed00bfd8e
ML antiquotations @{theory} and @{theory_ref} refer to the theory ancestry, not any accidental theory loader state;
 wenzelm parents: 
37304diff
changeset | 70 | (Scan.lift Args.name >> (fn name => | 
| 
59eed00bfd8e
ML antiquotations @{theory} and @{theory_ref} refer to the theory ancestry, not any accidental theory loader state;
 wenzelm parents: 
37304diff
changeset | 71 | "Context.get_theory (ML_Context.the_global_context ()) " ^ ML_Syntax.print_string name) | 
| 27340 | 72 | || Scan.succeed "ML_Context.the_global_context ()"); | 
| 73 | ||
| 74 | val _ = value "context" (Scan.succeed "ML_Context.the_local_context ()"); | |
| 75 | ||
| 76 | val _ = inline "typ" (Args.typ >> (ML_Syntax.atomic o ML_Syntax.print_typ)); | |
| 77 | val _ = inline "term" (Args.term >> (ML_Syntax.atomic o ML_Syntax.print_term)); | |
| 78 | val _ = inline "prop" (Args.prop >> (ML_Syntax.atomic o ML_Syntax.print_term)); | |
| 79 | ||
| 27379 | 80 | val _ = macro "let" (Args.context -- | 
| 36950 | 81 | Scan.lift | 
| 82 | (Parse.and_list1 (Parse.and_list1 Args.name_source -- (Args.$$$ "=" |-- Args.name_source))) | |
| 42360 | 83 | >> (fn (ctxt, args) => #2 (Proof_Context.match_bind true args ctxt))); | 
| 27379 | 84 | |
| 85 | val _ = macro "note" (Args.context :|-- (fn ctxt => | |
| 36950 | 86 | Parse.and_list1' (Scan.lift (Args.opt_thm_name I "=") -- Attrib.thms >> (fn ((a, srcs), ths) => | 
| 42360 | 87 | ((a, map (Attrib.attribute (Proof_Context.theory_of ctxt)) srcs), [(ths, [])]))) | 
| 88 | >> (fn args => #2 (Proof_Context.note_thmss "" args ctxt)))); | |
| 27379 | 89 | |
| 27340 | 90 | val _ = value "ctyp" (Args.typ >> (fn T => | 
| 91 | "Thm.ctyp_of (ML_Context.the_global_context ()) " ^ ML_Syntax.atomic (ML_Syntax.print_typ T))); | |
| 92 | ||
| 93 | val _ = value "cterm" (Args.term >> (fn t => | |
| 94 | "Thm.cterm_of (ML_Context.the_global_context ()) " ^ ML_Syntax.atomic (ML_Syntax.print_term t))); | |
| 95 | ||
| 96 | val _ = value "cprop" (Args.prop >> (fn t => | |
| 97 | "Thm.cterm_of (ML_Context.the_global_context ()) " ^ ML_Syntax.atomic (ML_Syntax.print_term t))); | |
| 98 | ||
| 99 | val _ = value "cpat" | |
| 42360 | 100 | (Args.context -- Scan.lift Args.name_source >> uncurry Proof_Context.read_term_pattern >> (fn t => | 
| 27340 | 101 | "Thm.cterm_of (ML_Context.the_global_context ()) " ^ ML_Syntax.atomic (ML_Syntax.print_term t))); | 
| 102 | ||
| 103 | ||
| 35396 | 104 | (* type classes *) | 
| 105 | ||
| 35670 | 106 | fun class syn = Args.context -- Scan.lift Args.name_source >> (fn (ctxt, s) => | 
| 42360 | 107 | Proof_Context.read_class ctxt s | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
41486diff
changeset | 108 | |> syn ? Lexicon.mark_class | 
| 35396 | 109 | |> ML_Syntax.print_string); | 
| 110 | ||
| 111 | val _ = inline "class" (class false); | |
| 112 | val _ = inline "class_syntax" (class true); | |
| 113 | ||
| 114 | val _ = inline "sort" (Args.context -- Scan.lift Args.name_source >> (fn (ctxt, s) => | |
| 115 | ML_Syntax.atomic (ML_Syntax.print_sort (Syntax.read_sort ctxt s)))); | |
| 116 | ||
| 117 | ||
| 35361 
4c7c849b70aa
more orthogonal antiquotations for type constructors;
 wenzelm parents: 
35262diff
changeset | 118 | (* type constructors *) | 
| 27340 | 119 | |
| 36950 | 120 | 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 | 121 | >> (fn (ctxt, (s, pos)) => | 
| 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 122 | let | 
| 42360 | 123 | val Type (c, _) = Proof_Context.read_type_name_proper ctxt false s; | 
| 42468 | 124 | 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 | 125 | val res = | 
| 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 126 | (case try check (c, decl) of | 
| 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 127 | SOME res => res | 
| 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 128 |         | NONE => error ("Not a " ^ kind ^ ": " ^ quote c ^ Position.str_of pos));
 | 
| 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 129 | in ML_Syntax.print_string res end); | 
| 27340 | 130 | |
| 35401 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 131 | val _ = inline "type_name" (type_name "logical type" (fn (c, Type.LogicalType _) => c)); | 
| 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 132 | val _ = inline "type_abbrev" (type_name "type abbreviation" (fn (c, Type.Abbreviation _) => c)); | 
| 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 133 | val _ = inline "nonterminal" (type_name "nonterminal" (fn (c, Type.Nonterminal) => c)); | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
41486diff
changeset | 134 | val _ = inline "type_syntax" (type_name "type" (fn (c, _) => Lexicon.mark_type c)); | 
| 35361 
4c7c849b70aa
more orthogonal antiquotations for type constructors;
 wenzelm parents: 
35262diff
changeset | 135 | |
| 
4c7c849b70aa
more orthogonal antiquotations for type constructors;
 wenzelm parents: 
35262diff
changeset | 136 | |
| 
4c7c849b70aa
more orthogonal antiquotations for type constructors;
 wenzelm parents: 
35262diff
changeset | 137 | (* constants *) | 
| 27340 | 138 | |
| 36950 | 139 | 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 | 140 | >> (fn (ctxt, (s, pos)) => | 
| 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 141 | let | 
| 42360 | 142 | val Const (c, _) = Proof_Context.read_const_proper ctxt false s; | 
| 143 | val res = check (Proof_Context.consts_of ctxt, c) | |
| 35401 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 144 | handle TYPE (msg, _, _) => error (msg ^ Position.str_of pos); | 
| 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 145 | in ML_Syntax.print_string res end); | 
| 27340 | 146 | |
| 35401 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 147 | val _ = inline "const_name" (const_name (fn (consts, c) => (Consts.the_type consts c; c))); | 
| 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 148 | val _ = inline "const_abbrev" (const_name (fn (consts, c) => (Consts.the_abbreviation consts c; c))); | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
41486diff
changeset | 149 | val _ = inline "const_syntax" (const_name (fn (_, c) => Lexicon.mark_const c)); | 
| 35401 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 150 | |
| 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 151 | |
| 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 152 | val _ = inline "syntax_const" | 
| 36950 | 153 | (Args.context -- Scan.lift (Parse.position Args.name) >> (fn (ctxt, (c, pos)) => | 
| 42360 | 154 | if is_some (Syntax.lookup_const (Proof_Context.syn_of ctxt) c) | 
| 42298 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42290diff
changeset | 155 | then ML_Syntax.print_string c | 
| 35401 
bfcbab8592ba
clarified @{const_name} (only logical consts) vs. @{const_abbrev};
 wenzelm parents: 
35396diff
changeset | 156 |     else error ("Unknown syntax const: " ^ quote c ^ Position.str_of pos)));
 | 
| 27340 | 157 | |
| 158 | val _ = inline "const" | |
| 27882 
eaa9fef9f4c1
Args.name_source(_position) for proper position information;
 wenzelm parents: 
27868diff
changeset | 159 | (Args.context -- Scan.lift Args.name_source -- Scan.optional | 
| 36950 | 160 |       (Scan.lift (Args.$$$ "(") |-- Parse.enum1' "," Args.typ --| Scan.lift (Args.$$$ ")")) []
 | 
| 35139 | 161 | >> (fn ((ctxt, raw_c), Ts) => | 
| 162 | let | |
| 42360 | 163 | val Const (c, _) = Proof_Context.read_const_proper ctxt true raw_c; | 
| 164 | val const = Const (c, Consts.instance (Proof_Context.consts_of ctxt) (c, Ts)); | |
| 35139 | 165 | in ML_Syntax.atomic (ML_Syntax.print_term const) end)); | 
| 27340 | 166 | |
| 167 | end; | |
| 168 |