| author | Christian Urban <urbanc@in.tum.de> | 
| Sun, 08 Feb 2009 08:45:25 +0000 | |
| changeset 104 | 5dcad9348e4d | 
| parent 96 | 018bfa718982 | 
| child 117 | 796c6ea633b3 | 
| permissions | -rw-r--r-- | 
| 
90
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
1  | 
|
| 24 | 2  | 
structure Chunks =  | 
3  | 
struct  | 
|
4  | 
||
| 
90
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
5  | 
(* rebuilding the output_list function from ThyOutput in order to *)  | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
6  | 
(* enable the option [gray] *)  | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
7  | 
|
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
8  | 
val gray = ref false;  | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
9  | 
|
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
10  | 
fun boolean "" = true  | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
11  | 
| boolean "true" = true  | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
12  | 
| boolean "false" = false  | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
13  | 
  | boolean s = error ("Bad boolean value: " ^ quote s);
 | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
14  | 
|
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
15  | 
fun output_list pretty src ctxt xs =  | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
16  | 
map (pretty ctxt) xs (*always pretty in order to exhibit errors!*)  | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
17  | 
|> (if ! ThyOutput.source then K [ThyOutput.pretty_text (ThyOutput.str_of_source src)] else I)  | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
18  | 
|> (if ! ThyOutput.quotes then map Pretty.quote else I)  | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
19  | 
|> (if ! ThyOutput.display then  | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
20  | 
map (Output.output o Pretty.string_of o Pretty.indent (! ThyOutput.indent))  | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
21  | 
#> space_implode "\\isasep\\isanewline%\n"  | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
22  | 
    #> (if ! gray then (enclose "\\begin{graybox}%\n" "%\n\\end{graybox}") else I) 
 | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
23  | 
    #> enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
 | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
24  | 
else  | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
25  | 
map (Output.output o (if ! ThyOutput.break then Pretty.string_of else Pretty.str_of))  | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
26  | 
#> space_implode "\\isasep\\isanewline%\n"  | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
27  | 
    #> enclose "\\isa{" "}");
 | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
28  | 
|
| 
96
 
018bfa718982
(re)defined the document antiquotation @subgoal in order to be closer to what is actually printed
 
Christian Urban <urbanc@in.tum.de> 
parents: 
90 
diff
changeset
 | 
29  | 
fun output pretty src ctxt = output_list pretty src ctxt o single  | 
| 
 
018bfa718982
(re)defined the document antiquotation @subgoal in order to be closer to what is actually printed
 
Christian Urban <urbanc@in.tum.de> 
parents: 
90 
diff
changeset
 | 
30  | 
|
| 
90
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
31  | 
val _ = ThyOutput.add_options  | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
32  | 
 [("gray", Library.setmp gray o boolean)]
 | 
| 
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
33  | 
|
| 
96
 
018bfa718982
(re)defined the document antiquotation @subgoal in order to be closer to what is actually printed
 
Christian Urban <urbanc@in.tum.de> 
parents: 
90 
diff
changeset
 | 
34  | 
|
| 
 
018bfa718982
(re)defined the document antiquotation @subgoal in order to be closer to what is actually printed
 
Christian Urban <urbanc@in.tum.de> 
parents: 
90 
diff
changeset
 | 
35  | 
|
| 
 
018bfa718982
(re)defined the document antiquotation @subgoal in order to be closer to what is actually printed
 
Christian Urban <urbanc@in.tum.de> 
parents: 
90 
diff
changeset
 | 
36  | 
|
| 24 | 37  | 
(** theory data **)  | 
38  | 
||
39  | 
structure ChunkData = TheoryDataFun  | 
|
40  | 
(  | 
|
41  | 
type T = (ML_Lex.token list * stamp) NameSpace.table  | 
|
42  | 
val empty = NameSpace.empty_table;  | 
|
43  | 
val copy = I;  | 
|
44  | 
val extend = I;  | 
|
45  | 
fun merge _ tables : T = NameSpace.merge_tables (eq_snd (op =)) tables handle Symtab.DUP dup =>  | 
|
46  | 
    error ("Attempt to merge different versions of ML chunk " ^ quote dup);
 | 
|
47  | 
);  | 
|
48  | 
||
49  | 
fun declare_chunk name thy =  | 
|
| 55 | 50  | 
(Sign.full_bname thy name,  | 
51  | 
ChunkData.map (snd o NameSpace.bind (Sign.naming_of thy)  | 
|
52  | 
(Binding.name name, ([], stamp ()))) thy  | 
|
| 24 | 53  | 
   handle Symtab.DUP _ => error ("Duplicate definition of ML chunk " ^ quote name));
 | 
54  | 
||
55  | 
fun augment_chunk tok name =  | 
|
56  | 
ChunkData.map (apsnd (Symtab.map_entry name (apfst (cons tok))));  | 
|
57  | 
||
58  | 
fun the_chunk thy name =  | 
|
59  | 
let val (space, tab) = ChunkData.get thy  | 
|
60  | 
in case Symtab.lookup tab (NameSpace.intern space name) of  | 
|
61  | 
      NONE => error ("Unknown ML chunk " ^ quote name)
 | 
|
62  | 
| SOME (toks, _) => rev toks  | 
|
63  | 
end;  | 
|
64  | 
||
65  | 
||
66  | 
(** parsers **)  | 
|
67  | 
||
68  | 
val not_eof = ML_Lex.stopper |> Scan.is_stopper #> not;  | 
|
69  | 
||
70  | 
val scan_space = Scan.many Symbol.is_blank;  | 
|
71  | 
||
72  | 
fun scan_cmt scan =  | 
|
73  | 
Scan.one (ML_Lex.kind_of #> equal ML_Lex.Comment) >>  | 
|
74  | 
(ML_Lex.content_of #> Symbol.explode #> Scan.finite Symbol.stopper scan #> fst);  | 
|
75  | 
||
76  | 
val scan_chunks = Scan.depend (fn (open_chunks, thy) =>  | 
|
77  | 
scan_cmt ((Scan.this_string "(*" -- scan_space -- Scan.this_string "@chunk" --  | 
|
78  | 
scan_space) |-- Symbol.scan_id --|  | 
|
79  | 
(scan_space -- Scan.this_string "*)") >> (fn name =>  | 
|
80  | 
let val (name', thy') = declare_chunk name thy  | 
|
81  | 
in ((name' :: open_chunks, thy'), "") end))  | 
|
82  | 
||  | 
|
83  | 
scan_cmt (Scan.this_string "(*" -- scan_space -- Scan.this_string "@end" --  | 
|
84  | 
scan_space -- Scan.this_string "*)" >> (fn _ =>  | 
|
85  | 
(case open_chunks of  | 
|
86  | 
[] => error "end without matching chunk encountered"  | 
|
87  | 
| _ :: open_chunks' => ((open_chunks', thy), ""))))  | 
|
88  | 
||  | 
|
89  | 
Scan.one not_eof >> (fn tok =>  | 
|
90  | 
((open_chunks, fold (augment_chunk tok) open_chunks thy), ML_Lex.content_of tok)));  | 
|
91  | 
||
92  | 
fun chunks_of s thy =  | 
|
93  | 
let val toks = Source.exhaust (ML_Lex.source (Source.of_string s))  | 
|
94  | 
in case Scan.finite' ML_Lex.stopper (Scan.repeat (Scan.error scan_chunks))  | 
|
95  | 
(([], thy), toks) of  | 
|
96  | 
(toks', (([], thy'), _)) => (implode toks', thy')  | 
|
97  | 
| (_, ((open_chunks, _), _)) =>  | 
|
98  | 
        error ("Open chunks at end of text: " ^ commas_quote open_chunks)
 | 
|
99  | 
end;  | 
|
100  | 
||
101  | 
||
102  | 
(** use command **)  | 
|
103  | 
||
104  | 
fun load_ml dir raw_path thy =  | 
|
105  | 
(case ThyLoad.check_ml dir raw_path of  | 
|
106  | 
    NONE => error ("Could not find ML file " ^ quote (Path.implode raw_path))
 | 
|
107  | 
| SOME (path, _) =>  | 
|
108  | 
let val (s, thy') = chunks_of (File.read path) thy  | 
|
109  | 
in  | 
|
110  | 
Context.theory_map (ThyInfo.exec_file false raw_path) thy'  | 
|
111  | 
end);  | 
|
112  | 
||
113  | 
fun use_chunks path thy =  | 
|
114  | 
load_ml (ThyInfo.master_directory (Context.theory_name thy)) path thy;  | 
|
115  | 
||
116  | 
val _ =  | 
|
117  | 
OuterSyntax.command "use_chunks" "eval ML text from file"  | 
|
118  | 
(OuterKeyword.tag_ml OuterKeyword.thy_decl)  | 
|
119  | 
(OuterParse.path >> (Toplevel.theory o use_chunks));  | 
|
120  | 
||
121  | 
||
122  | 
(** antiquotation **)  | 
|
123  | 
||
124  | 
val str_of_source = space_implode " " o map OuterLex.unparse o #2 o #1 o Args.dest_src;  | 
|
125  | 
||
126  | 
val verbatim = space_implode "\\verb,|," o map (enclose "\\verb|" "|") o space_explode "|";  | 
|
127  | 
||
128  | 
val transform_cmts =  | 
|
129  | 
Scan.finite ML_Lex.stopper (Scan.repeat  | 
|
130  | 
( scan_cmt  | 
|
131  | 
          (Scan.this_string "(*{" |--
 | 
|
132  | 
Scan.repeat (Scan.unless (Scan.this_string "}*)")  | 
|
133  | 
(Scan.one (not o Symbol.is_eof))) --|  | 
|
134  | 
Scan.this_string "}*)" >>  | 
|
| 
57
 
065f472c09ab
Repaired output of marginal comments in ML antiquotation.
 
berghofe 
parents: 
55 
diff
changeset
 | 
135  | 
           (Symbol.encode_raw o enclose "\\hfill\\textrm{\\textit{" "}}" o implode))
 | 
| 24 | 136  | 
|| Scan.one not_eof >> ML_Lex.content_of)) #>  | 
137  | 
fst;  | 
|
138  | 
||
139  | 
fun output_chunk src ctxt name =  | 
|
140  | 
let  | 
|
141  | 
val toks = the_chunk (ProofContext.theory_of ctxt) name;  | 
|
142  | 
val (spc, toks') = (case toks of  | 
|
143  | 
        [] => ("", [])
 | 
|
144  | 
      | [tok] => ("", [tok])
 | 
|
145  | 
| tok :: toks' =>  | 
|
146  | 
let  | 
|
147  | 
val (toks'', tok') = split_last toks'  | 
|
148  | 
val toks''' =  | 
|
149  | 
if ML_Lex.kind_of tok' = ML_Lex.Space then toks''  | 
|
150  | 
else toks'  | 
|
151  | 
in  | 
|
152  | 
if ML_Lex.kind_of tok = ML_Lex.Space then  | 
|
153  | 
(tok |> ML_Lex.content_of |> Symbol.explode |>  | 
|
154  | 
take_suffix (equal " ") |> snd |> implode,  | 
|
155  | 
toks''')  | 
|
156  | 
            else ("", tok :: toks''')
 | 
|
157  | 
end)  | 
|
158  | 
val txt = spc ^ implode (transform_cmts toks')  | 
|
159  | 
in  | 
|
| 
57
 
065f472c09ab
Repaired output of marginal comments in ML antiquotation.
 
berghofe 
parents: 
55 
diff
changeset
 | 
160  | 
txt |> split_lines |>  | 
| 
90
 
b071a0b88298
made chunks aware of the gray-option
 
Christian Urban <urbanc@in.tum.de> 
parents: 
57 
diff
changeset
 | 
161  | 
output_list (fn _ => fn s => Pretty.str s) src ctxt  | 
| 24 | 162  | 
end;  | 
163  | 
||
164  | 
val _ = ThyOutput.add_commands  | 
|
165  | 
  [("ML_chunk", ThyOutput.args (Scan.lift Args.name) output_chunk)];
 | 
|
166  | 
||
167  | 
end;  |