| author | wenzelm | 
| Mon, 13 Oct 2014 22:43:29 +0200 | |
| changeset 58668 | 1891f17c6124 | 
| parent 58634 | 9f10d82e8188 | 
| child 58928 | 23d0ffd48006 | 
| permissions | -rw-r--r-- | 
| 56205 | 1  | 
(* Title: Pure/ML/ml_antiquotations.ML  | 
2  | 
Author: Makarius  | 
|
3  | 
||
4  | 
Miscellaneous ML antiquotations.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
structure ML_Antiquotations: sig end =  | 
|
8  | 
struct  | 
|
9  | 
||
| 58632 | 10  | 
(* ML support *)  | 
11  | 
||
12  | 
val _ = Theory.setup  | 
|
13  | 
 (ML_Antiquotation.inline @{binding assert}
 | 
|
14  | 
(Scan.succeed "(fn b => if b then () else raise General.Fail \"Assertion failed\")") #>  | 
|
15  | 
||
16  | 
  ML_Antiquotation.inline @{binding make_string} (Scan.succeed ml_make_string) #>
 | 
|
17  | 
||
18  | 
  ML_Antiquotation.declaration @{binding print}
 | 
|
19  | 
(Scan.lift (Scan.optional Args.name "Output.writeln"))  | 
|
20  | 
(fn src => fn output => fn ctxt =>  | 
|
21  | 
let  | 
|
22  | 
val (_, pos) = Token.name_of_src src;  | 
|
23  | 
val (a, ctxt') = ML_Antiquotation.variant "output" ctxt;  | 
|
24  | 
val env =  | 
|
25  | 
"val " ^ a ^ ": string -> unit =\n\  | 
|
26  | 
            \  (" ^ output ^ ") o (fn s => s ^ Position.here (" ^
 | 
|
27  | 
ML_Syntax.print_position pos ^ "));\n";  | 
|
28  | 
val body =  | 
|
29  | 
            "(fn x => (Isabelle." ^ a ^ " (" ^ ml_make_string ^ " x); x))";
 | 
|
30  | 
in (K (env, body), ctxt') end));  | 
|
31  | 
||
32  | 
||
33  | 
(* embedded source *)  | 
|
34  | 
||
35  | 
val _ = Theory.setup  | 
|
36  | 
 (ML_Antiquotation.value @{binding source}
 | 
|
37  | 
    (Scan.lift Args.text_source_position >> (fn {delimited, text, pos} =>
 | 
|
38  | 
      "{delimited = " ^ Bool.toString delimited ^
 | 
|
39  | 
", text = " ^ ML_Syntax.print_string text ^  | 
|
40  | 
", pos = " ^ ML_Syntax.print_position pos ^ "}")));  | 
|
41  | 
||
42  | 
||
43  | 
(* formal entities *)  | 
|
44  | 
||
| 56205 | 45  | 
val _ = Theory.setup  | 
| 56467 | 46  | 
 (ML_Antiquotation.value @{binding system_option}
 | 
| 56465 | 47  | 
(Args.context -- Scan.lift (Parse.position Args.name) >> (fn (ctxt, (name, pos)) =>  | 
| 56467 | 48  | 
(Context_Position.report ctxt pos (Options.default_markup (name, pos));  | 
49  | 
ML_Syntax.print_string name))) #>  | 
|
| 56205 | 50  | 
|
51  | 
  ML_Antiquotation.value @{binding theory}
 | 
|
52  | 
(Args.context -- Scan.lift (Parse.position Args.name) >> (fn (ctxt, (name, pos)) =>  | 
|
53  | 
(Context_Position.report ctxt pos  | 
|
54  | 
(Theory.get_markup (Context.get_theory (Proof_Context.theory_of ctxt) name));  | 
|
55  | 
"Context.get_theory (Proof_Context.theory_of ML_context) " ^ ML_Syntax.print_string name))  | 
|
56  | 
|| Scan.succeed "Proof_Context.theory_of ML_context") #>  | 
|
57  | 
||
58  | 
  ML_Antiquotation.value @{binding theory_context}
 | 
|
59  | 
(Args.context -- Scan.lift (Parse.position Args.name) >> (fn (ctxt, (name, pos)) =>  | 
|
60  | 
(Context_Position.report ctxt pos  | 
|
61  | 
(Theory.get_markup (Context.get_theory (Proof_Context.theory_of ctxt) name));  | 
|
62  | 
"Proof_Context.get_global (Proof_Context.theory_of ML_context) " ^  | 
|
63  | 
ML_Syntax.print_string name))) #>  | 
|
64  | 
||
65  | 
  ML_Antiquotation.inline @{binding context} (Scan.succeed "Isabelle.ML_context") #>
 | 
|
66  | 
||
67  | 
  ML_Antiquotation.inline @{binding typ} (Args.typ >> (ML_Syntax.atomic o ML_Syntax.print_typ)) #>
 | 
|
68  | 
  ML_Antiquotation.inline @{binding term} (Args.term >> (ML_Syntax.atomic o ML_Syntax.print_term)) #>
 | 
|
69  | 
  ML_Antiquotation.inline @{binding prop} (Args.prop >> (ML_Syntax.atomic o ML_Syntax.print_term)) #>
 | 
|
70  | 
||
71  | 
  ML_Antiquotation.value @{binding ctyp} (Args.typ >> (fn T =>
 | 
|
72  | 
"Thm.ctyp_of (Proof_Context.theory_of ML_context) " ^  | 
|
73  | 
ML_Syntax.atomic (ML_Syntax.print_typ T))) #>  | 
|
74  | 
||
75  | 
  ML_Antiquotation.value @{binding cterm} (Args.term >> (fn t =>
 | 
|
76  | 
"Thm.cterm_of (Proof_Context.theory_of ML_context) " ^  | 
|
77  | 
ML_Syntax.atomic (ML_Syntax.print_term t))) #>  | 
|
78  | 
||
79  | 
  ML_Antiquotation.value @{binding cprop} (Args.prop >> (fn t =>
 | 
|
80  | 
"Thm.cterm_of (Proof_Context.theory_of ML_context) " ^  | 
|
81  | 
ML_Syntax.atomic (ML_Syntax.print_term t))) #>  | 
|
82  | 
||
83  | 
  ML_Antiquotation.value @{binding cpat}
 | 
|
84  | 
(Args.context --  | 
|
85  | 
Scan.lift Args.name_inner_syntax >> uncurry Proof_Context.read_term_pattern >> (fn t =>  | 
|
86  | 
"Thm.cterm_of (Proof_Context.theory_of ML_context) " ^  | 
|
87  | 
ML_Syntax.atomic (ML_Syntax.print_term t))));  | 
|
88  | 
||
89  | 
||
90  | 
(* type classes *)  | 
|
91  | 
||
92  | 
fun class syn = Args.context -- Scan.lift Args.name_inner_syntax >> (fn (ctxt, s) =>  | 
|
93  | 
Proof_Context.read_class ctxt s  | 
|
94  | 
|> syn ? Lexicon.mark_class  | 
|
95  | 
|> ML_Syntax.print_string);  | 
|
96  | 
||
97  | 
val _ = Theory.setup  | 
|
98  | 
 (ML_Antiquotation.inline @{binding class} (class false) #>
 | 
|
99  | 
  ML_Antiquotation.inline @{binding class_syntax} (class true) #>
 | 
|
100  | 
||
101  | 
  ML_Antiquotation.inline @{binding sort}
 | 
|
102  | 
(Args.context -- Scan.lift Args.name_inner_syntax >> (fn (ctxt, s) =>  | 
|
103  | 
ML_Syntax.atomic (ML_Syntax.print_sort (Syntax.read_sort ctxt s)))));  | 
|
104  | 
||
105  | 
||
106  | 
(* type constructors *)  | 
|
107  | 
||
108  | 
fun type_name kind check = Args.context -- Scan.lift (Parse.position Args.name_inner_syntax)  | 
|
109  | 
>> (fn (ctxt, (s, pos)) =>  | 
|
110  | 
let  | 
|
111  | 
      val Type (c, _) = Proof_Context.read_type_name {proper = true, strict = false} ctxt s;
 | 
|
112  | 
val decl = Type.the_decl (Proof_Context.tsig_of ctxt) (c, pos);  | 
|
113  | 
val res =  | 
|
114  | 
(case try check (c, decl) of  | 
|
115  | 
SOME res => res  | 
|
116  | 
        | NONE => error ("Not a " ^ kind ^ ": " ^ quote c ^ Position.here pos));
 | 
|
117  | 
in ML_Syntax.print_string res end);  | 
|
118  | 
||
119  | 
val _ = Theory.setup  | 
|
120  | 
 (ML_Antiquotation.inline @{binding type_name}
 | 
|
121  | 
(type_name "logical type" (fn (c, Type.LogicalType _) => c)) #>  | 
|
122  | 
  ML_Antiquotation.inline @{binding type_abbrev}
 | 
|
123  | 
(type_name "type abbreviation" (fn (c, Type.Abbreviation _) => c)) #>  | 
|
124  | 
  ML_Antiquotation.inline @{binding nonterminal}
 | 
|
125  | 
(type_name "nonterminal" (fn (c, Type.Nonterminal) => c)) #>  | 
|
126  | 
  ML_Antiquotation.inline @{binding type_syntax}
 | 
|
127  | 
(type_name "type" (fn (c, _) => Lexicon.mark_type c)));  | 
|
128  | 
||
129  | 
||
130  | 
(* constants *)  | 
|
131  | 
||
132  | 
fun const_name check = Args.context -- Scan.lift (Parse.position Args.name_inner_syntax)  | 
|
133  | 
>> (fn (ctxt, (s, pos)) =>  | 
|
134  | 
let  | 
|
135  | 
      val Const (c, _) = Proof_Context.read_const {proper = true, strict = false} ctxt s;
 | 
|
136  | 
val res = check (Proof_Context.consts_of ctxt, c)  | 
|
137  | 
handle TYPE (msg, _, _) => error (msg ^ Position.here pos);  | 
|
138  | 
in ML_Syntax.print_string res end);  | 
|
139  | 
||
140  | 
val _ = Theory.setup  | 
|
141  | 
 (ML_Antiquotation.inline @{binding const_name}
 | 
|
142  | 
(const_name (fn (consts, c) => (Consts.the_const consts c; c))) #>  | 
|
143  | 
  ML_Antiquotation.inline @{binding const_abbrev}
 | 
|
144  | 
(const_name (fn (consts, c) => (Consts.the_abbreviation consts c; c))) #>  | 
|
145  | 
  ML_Antiquotation.inline @{binding const_syntax}
 | 
|
146  | 
(const_name (fn (_, c) => Lexicon.mark_const c)) #>  | 
|
147  | 
||
148  | 
  ML_Antiquotation.inline @{binding syntax_const}
 | 
|
149  | 
(Args.context -- Scan.lift (Parse.position Args.name) >> (fn (ctxt, (c, pos)) =>  | 
|
150  | 
if is_some (Syntax.lookup_const (Proof_Context.syn_of ctxt) c)  | 
|
151  | 
then ML_Syntax.print_string c  | 
|
152  | 
      else error ("Unknown syntax const: " ^ quote c ^ Position.here pos))) #>
 | 
|
153  | 
||
154  | 
  ML_Antiquotation.inline @{binding const}
 | 
|
| 56251 | 155  | 
(Args.context -- Scan.lift (Parse.position Args.name_inner_syntax) -- Scan.optional  | 
| 56205 | 156  | 
        (Scan.lift (Args.$$$ "(") |-- Parse.enum1' "," Args.typ --| Scan.lift (Args.$$$ ")")) []
 | 
| 56251 | 157  | 
>> (fn ((ctxt, (raw_c, pos)), Ts) =>  | 
| 56205 | 158  | 
let  | 
159  | 
val Const (c, _) =  | 
|
160  | 
            Proof_Context.read_const {proper = true, strict = true} ctxt raw_c;
 | 
|
161  | 
val consts = Proof_Context.consts_of ctxt;  | 
|
162  | 
val n = length (Consts.typargs consts (c, Consts.type_scheme consts c));  | 
|
163  | 
val _ = length Ts <> n andalso  | 
|
164  | 
            error ("Constant requires " ^ string_of_int n ^ " type argument(s): " ^
 | 
|
| 56251 | 165  | 
              quote c ^ enclose "(" ")" (commas (replicate n "_")) ^ Position.here pos);
 | 
| 56205 | 166  | 
val const = Const (c, Consts.instance consts (c, Ts));  | 
167  | 
in ML_Syntax.atomic (ML_Syntax.print_term const) end)));  | 
|
168  | 
||
169  | 
||
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
170  | 
(* basic combinators *)  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
171  | 
|
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
172  | 
local  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
173  | 
|
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
174  | 
val parameter = Parse.position Parse.nat >> (fn (n, pos) =>  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
175  | 
  if n > 1 then n else error ("Bad parameter: " ^ string_of_int n ^ Position.here pos));
 | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
176  | 
|
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
177  | 
fun indices n = map string_of_int (1 upto n);  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
178  | 
|
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
179  | 
fun empty n = replicate_string n " []";  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
180  | 
fun dummy n = replicate_string n " _";  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
181  | 
fun vars x n = implode (map (fn a => " " ^ x ^ a) (indices n));  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
182  | 
fun cons n = implode (map (fn a => " (x" ^ a ^ " :: xs" ^ a ^ ")") (indices n));  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
183  | 
|
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
184  | 
val tuple = enclose "(" ")" o commas;
 | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
185  | 
fun tuple_empty n = tuple (replicate n "[]");  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
186  | 
fun tuple_vars x n = tuple (map (fn a => x ^ a) (indices n));  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
187  | 
fun tuple_cons n = "(" ^ tuple_vars "x" n ^ " :: xs)"
 | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
188  | 
fun cons_tuple n = tuple (map (fn a => "x" ^ a ^ " :: xs" ^ a) (indices n));  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
189  | 
|
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
190  | 
in  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
191  | 
|
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
192  | 
val _ = Theory.setup  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
193  | 
 (ML_Antiquotation.value @{binding map}
 | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
194  | 
(Scan.lift parameter >> (fn n =>  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
195  | 
"fn f =>\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
196  | 
\ let\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
197  | 
\ fun map _" ^ empty n ^ " = []\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
198  | 
\ | map f" ^ cons n ^ " = f" ^ vars "x" n ^ " :: map f" ^ vars "xs" n ^ "\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
199  | 
\ | map _" ^ dummy n ^ " = raise ListPair.UnequalLengths\n" ^  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
200  | 
" in map f end")) #>  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
201  | 
  ML_Antiquotation.value @{binding fold}
 | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
202  | 
(Scan.lift parameter >> (fn n =>  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
203  | 
"fn f =>\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
204  | 
\ let\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
205  | 
\ fun fold _" ^ empty n ^ " a = a\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
206  | 
\ | fold f" ^ cons n ^ " a = fold f" ^ vars "xs" n ^ " (f" ^ vars "x" n ^ " a)\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
207  | 
\ | fold _" ^ dummy n ^ " _ = raise ListPair.UnequalLengths\n" ^  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
208  | 
" in fold f end")) #>  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
209  | 
  ML_Antiquotation.value @{binding fold_map}
 | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
210  | 
(Scan.lift parameter >> (fn n =>  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
211  | 
"fn f =>\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
212  | 
\ let\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
213  | 
\ fun fold_map _" ^ empty n ^ " a = ([], a)\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
214  | 
\ | fold_map f" ^ cons n ^ " a =\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
215  | 
\ let\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
216  | 
\ val (x, a') = f" ^ vars "x" n ^ " a\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
217  | 
\ val (xs, a'') = fold_map f" ^ vars "xs" n ^ " a'\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
218  | 
\ in (x :: xs, a'') end\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
219  | 
\ | fold_map _" ^ dummy n ^ " _ = raise ListPair.UnequalLengths\n" ^  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
220  | 
" in fold_map f end")) #>  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
221  | 
  ML_Antiquotation.value @{binding split_list}
 | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
222  | 
(Scan.lift parameter >> (fn n =>  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
223  | 
"fn list =>\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
224  | 
\ let\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
225  | 
\ fun split_list [] =" ^ tuple_empty n ^ "\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
226  | 
\ | split_list" ^ tuple_cons n ^ " =\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
227  | 
\ let val" ^ tuple_vars "xs" n ^ " = split_list xs\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
228  | 
\ in " ^ cons_tuple n ^ "end\n\  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
229  | 
\ in split_list list end")))  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
230  | 
|
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
231  | 
end;  | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
232  | 
|
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58632 
diff
changeset
 | 
233  | 
|
| 56205 | 234  | 
(* outer syntax *)  | 
235  | 
||
236  | 
fun with_keyword f =  | 
|
237  | 
Args.theory -- Scan.lift (Parse.position Parse.string) >> (fn (thy, (name, pos)) =>  | 
|
238  | 
(f ((name, Thy_Header.the_keyword thy name), pos)  | 
|
239  | 
handle ERROR msg => error (msg ^ Position.here pos)));  | 
|
240  | 
||
241  | 
val _ = Theory.setup  | 
|
242  | 
 (ML_Antiquotation.value @{binding keyword}
 | 
|
243  | 
(with_keyword  | 
|
244  | 
(fn ((name, NONE), _) => "Parse.$$$ " ^ ML_Syntax.print_string name  | 
|
245  | 
| ((name, SOME _), pos) =>  | 
|
246  | 
            error ("Expected minor keyword " ^ quote name ^ Position.here pos))) #>
 | 
|
247  | 
  ML_Antiquotation.value @{binding command_spec}
 | 
|
248  | 
(with_keyword  | 
|
249  | 
(fn ((name, SOME kind), pos) =>  | 
|
250  | 
"Keyword.command_spec " ^ ML_Syntax.atomic  | 
|
251  | 
((ML_Syntax.print_pair  | 
|
252  | 
(ML_Syntax.print_pair ML_Syntax.print_string  | 
|
253  | 
(ML_Syntax.print_pair  | 
|
254  | 
(ML_Syntax.print_pair ML_Syntax.print_string  | 
|
255  | 
(ML_Syntax.print_list ML_Syntax.print_string))  | 
|
256  | 
(ML_Syntax.print_list ML_Syntax.print_string)))  | 
|
257  | 
ML_Syntax.print_position) ((name, kind), pos))  | 
|
258  | 
| ((name, NONE), pos) =>  | 
|
259  | 
            error ("Expected command keyword " ^ quote name ^ Position.here pos))));
 | 
|
260  | 
||
261  | 
end;  | 
|
262  |