| author | haftmann | 
| Fri, 01 Oct 2010 10:25:36 +0200 | |
| changeset 39811 | 0659e84bdc5f | 
| parent 39666 | 4f628ee48fd7 | 
| child 40402 | b646316f8b3c | 
| permissions | -rw-r--r-- | 
| 
7726
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
1  | 
(* Title: Pure/Thy/latex.ML  | 
| 
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
Author: Markus Wenzel, TU Muenchen  | 
| 
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
|
| 15801 | 4  | 
LaTeX presentation elements -- based on outer lexical syntax.  | 
| 
7726
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
signature LATEX =  | 
| 
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
sig  | 
| 14924 | 9  | 
val output_known_symbols: (string -> bool) * (string -> bool) ->  | 
10  | 
Symbol.symbol list -> string  | 
|
| 11719 | 11  | 
val output_symbols: Symbol.symbol list -> string  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
30642 
diff
changeset
 | 
12  | 
val output_basic: Token.T -> string  | 
| 17081 | 13  | 
val output_markup: string -> string -> string  | 
14  | 
val output_markup_env: string -> string -> string  | 
|
15  | 
val output_verbatim: string -> string  | 
|
16  | 
val markup_true: string  | 
|
17  | 
val markup_false: string  | 
|
18  | 
val begin_delim: string -> string  | 
|
19  | 
val end_delim: string -> string  | 
|
20  | 
val begin_tag: string -> string  | 
|
21  | 
val end_tag: string -> string  | 
|
| 9707 | 22  | 
val tex_trailer: string  | 
| 9916 | 23  | 
val isabelle_file: string -> string -> string  | 
| 14924 | 24  | 
val symbol_source: (string -> bool) * (string -> bool) ->  | 
25  | 
string -> Symbol.symbol list -> string  | 
|
| 
7726
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
26  | 
val theory_entry: string -> string  | 
| 9135 | 27  | 
val modes: string list  | 
| 
7726
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
28  | 
end;  | 
| 
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
29  | 
|
| 
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
30  | 
structure Latex: LATEX =  | 
| 
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
31  | 
struct  | 
| 
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
32  | 
|
| 
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
33  | 
|
| 
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
34  | 
(* symbol output *)  | 
| 
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
35  | 
|
| 7900 | 36  | 
local  | 
37  | 
||
| 
7726
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
38  | 
val output_chr = fn  | 
| 9505 | 39  | 
" " => "\\ " |  | 
| 
7726
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
40  | 
"\n" => "\\isanewline\n" |  | 
| 9668 | 41  | 
  "!" => "{\\isacharbang}" |
 | 
42  | 
  "\"" => "{\\isachardoublequote}" |
 | 
|
| 9663 | 43  | 
  "#" => "{\\isacharhash}" |
 | 
44  | 
  "$" => "{\\isachardollar}" |
 | 
|
45  | 
  "%" => "{\\isacharpercent}" |
 | 
|
46  | 
  "&" => "{\\isacharampersand}" |
 | 
|
47  | 
  "'" => "{\\isacharprime}" |
 | 
|
48  | 
  "(" => "{\\isacharparenleft}" |
 | 
|
49  | 
  ")" => "{\\isacharparenright}" |
 | 
|
50  | 
  "*" => "{\\isacharasterisk}" |
 | 
|
| 9668 | 51  | 
  "+" => "{\\isacharplus}" |
 | 
52  | 
  "," => "{\\isacharcomma}" |
 | 
|
| 9663 | 53  | 
  "-" => "{\\isacharminus}" |
 | 
| 9668 | 54  | 
  "." => "{\\isachardot}" |
 | 
55  | 
  "/" => "{\\isacharslash}" |
 | 
|
56  | 
  ":" => "{\\isacharcolon}" |
 | 
|
57  | 
  ";" => "{\\isacharsemicolon}" |
 | 
|
| 9663 | 58  | 
  "<" => "{\\isacharless}" |
 | 
| 9668 | 59  | 
  "=" => "{\\isacharequal}" |
 | 
| 9663 | 60  | 
  ">" => "{\\isachargreater}" |
 | 
| 9668 | 61  | 
  "?" => "{\\isacharquery}" |
 | 
62  | 
  "@" => "{\\isacharat}" |
 | 
|
| 9663 | 63  | 
  "[" => "{\\isacharbrackleft}" |
 | 
64  | 
  "\\" => "{\\isacharbackslash}" |
 | 
|
65  | 
  "]" => "{\\isacharbrackright}" |
 | 
|
66  | 
  "^" => "{\\isacharcircum}" |
 | 
|
67  | 
  "_" => "{\\isacharunderscore}" |
 | 
|
| 9668 | 68  | 
  "`" => "{\\isacharbackquote}" |
 | 
| 9663 | 69  | 
  "{" => "{\\isacharbraceleft}" |
 | 
70  | 
  "|" => "{\\isacharbar}" |
 | 
|
71  | 
  "}" => "{\\isacharbraceright}" |
 | 
|
72  | 
  "~" => "{\\isachartilde}" |
 | 
|
| 14992 | 73  | 
  c => if Symbol.is_ascii_digit c then enclose "{\\isadigit{" "}}" c else c;
 | 
| 
7726
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
74  | 
|
| 14924 | 75  | 
val output_chrs = translate_string output_chr;  | 
76  | 
||
77  | 
fun output_known_sym (known_sym, known_ctrl) sym =  | 
|
| 14874 | 78  | 
(case Symbol.decode sym of  | 
79  | 
Symbol.Char s => output_chr s  | 
|
| 
37533
 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 
wenzelm 
parents: 
36959 
diff
changeset
 | 
80  | 
| Symbol.UTF8 s => s  | 
| 14924 | 81  | 
  | Symbol.Sym s => if known_sym s then enclose "{\\isasym" "}" s else output_chrs sym
 | 
82  | 
| Symbol.Ctrl s => if known_ctrl s then enclose "\\isactrl" " " s else output_chrs sym  | 
|
| 14874 | 83  | 
| Symbol.Raw s => s);  | 
| 7900 | 84  | 
|
85  | 
in  | 
|
86  | 
||
| 14924 | 87  | 
val output_known_symbols = implode oo (map o output_known_sym);  | 
88  | 
val output_symbols = output_known_symbols (K true, K true);  | 
|
| 8192 | 89  | 
val output_syms = output_symbols o Symbol.explode;  | 
| 7900 | 90  | 
|
| 
27874
 
f0364f1c0ecf
antiquotes: proper SymbolPos decoding, adapted Antiquote.read/Antiq;
 
wenzelm 
parents: 
27809 
diff
changeset
 | 
91  | 
val output_syms_antiq =  | 
| 30589 | 92  | 
(fn Antiquote.Text ss => output_symbols (map Symbol_Pos.symbol ss)  | 
| 
27874
 
f0364f1c0ecf
antiquotes: proper SymbolPos decoding, adapted Antiquote.read/Antiq;
 
wenzelm 
parents: 
27809 
diff
changeset
 | 
93  | 
| Antiquote.Antiq (ss, _) =>  | 
| 30573 | 94  | 
enclose "%\n\\isaantiq\n" "%\n\\endisaantiq\n" (output_symbols (map Symbol_Pos.symbol ss))  | 
| 27344 | 95  | 
    | Antiquote.Open _ => "{\\isaantiqopen}"
 | 
| 
27874
 
f0364f1c0ecf
antiquotes: proper SymbolPos decoding, adapted Antiquote.read/Antiq;
 
wenzelm 
parents: 
27809 
diff
changeset
 | 
96  | 
    | Antiquote.Close _ => "{\\isaantiqclose}");
 | 
| 
22648
 
8c6b4f7548e3
output_basic: added isaantiq markup (only inside verbatim tokens);
 
wenzelm 
parents: 
19265 
diff
changeset
 | 
97  | 
|
| 7900 | 98  | 
end;  | 
| 
7726
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
99  | 
|
| 
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
100  | 
|
| 
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
101  | 
(* token output *)  | 
| 
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
102  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
30642 
diff
changeset
 | 
103  | 
val invisible_token = Token.keyword_with (fn s => s = ";") orf Token.is_kind Token.Comment;  | 
| 7756 | 104  | 
|
| 17081 | 105  | 
fun output_basic tok =  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
30642 
diff
changeset
 | 
106  | 
let val s = Token.content_of tok in  | 
| 17081 | 107  | 
if invisible_token tok then ""  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
30642 
diff
changeset
 | 
108  | 
else if Token.is_kind Token.Command tok then  | 
| 17081 | 109  | 
      "\\isacommand{" ^ output_syms s ^ "}"
 | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
30642 
diff
changeset
 | 
110  | 
else if Token.is_kind Token.Keyword tok andalso Syntax.is_ascii_identifier s then  | 
| 17081 | 111  | 
      "\\isakeyword{" ^ output_syms s ^ "}"
 | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
30642 
diff
changeset
 | 
112  | 
else if Token.is_kind Token.String tok then  | 
| 17169 | 113  | 
      enclose "{\\isachardoublequoteopen}" "{\\isachardoublequoteclose}" (output_syms s)
 | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
30642 
diff
changeset
 | 
114  | 
else if Token.is_kind Token.AltString tok then  | 
| 17169 | 115  | 
      enclose "{\\isacharbackquoteopen}" "{\\isacharbackquoteclose}" (output_syms s)
 | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
30642 
diff
changeset
 | 
116  | 
else if Token.is_kind Token.Verbatim tok then  | 
| 
27874
 
f0364f1c0ecf
antiquotes: proper SymbolPos decoding, adapted Antiquote.read/Antiq;
 
wenzelm 
parents: 
27809 
diff
changeset
 | 
117  | 
let  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
30642 
diff
changeset
 | 
118  | 
val (txt, pos) = Token.source_position_of tok;  | 
| 30642 | 119  | 
val ants = Antiquote.read (Symbol_Pos.explode (txt, pos), pos);  | 
| 
27874
 
f0364f1c0ecf
antiquotes: proper SymbolPos decoding, adapted Antiquote.read/Antiq;
 
wenzelm 
parents: 
27809 
diff
changeset
 | 
120  | 
val out = implode (map output_syms_antiq ants);  | 
| 
 
f0364f1c0ecf
antiquotes: proper SymbolPos decoding, adapted Antiquote.read/Antiq;
 
wenzelm 
parents: 
27809 
diff
changeset
 | 
121  | 
      in enclose "{\\isacharverbatimopen}" "{\\isacharverbatimclose}" out end
 | 
| 17081 | 122  | 
else output_syms s  | 
123  | 
end;  | 
|
124  | 
||
125  | 
fun output_markup cmd txt = "%\n\\isamarkup" ^ cmd ^ "{" ^ Symbol.strip_blanks txt ^ "%\n}\n";
 | 
|
| 7745 | 126  | 
|
| 17081 | 127  | 
fun output_markup_env cmd txt =  | 
128  | 
  "%\n\\begin{isamarkup" ^ cmd ^ "}%\n" ^
 | 
|
129  | 
Symbol.strip_blanks txt ^  | 
|
130  | 
  "%\n\\end{isamarkup" ^ cmd ^ "}%\n";
 | 
|
| 7756 | 131  | 
|
| 17081 | 132  | 
fun output_verbatim txt = "%\n" ^ Symbol.strip_blanks txt ^ "\n";  | 
| 
7726
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
133  | 
|
| 17081 | 134  | 
val markup_true = "\\isamarkuptrue%\n";  | 
135  | 
val markup_false = "\\isamarkupfalse%\n";  | 
|
| 
7726
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
136  | 
|
| 17081 | 137  | 
val begin_delim = enclose "%\n\\isadelim" "\n";  | 
138  | 
val end_delim = enclose "%\n\\endisadelim" "\n";  | 
|
139  | 
val begin_tag = enclose "%\n\\isatag" "\n";  | 
|
140  | 
fun end_tag tg = enclose "%\n\\endisatag" "\n" tg ^ enclose "{\\isafold" "}%\n" tg;
 | 
|
| 11860 | 141  | 
|
142  | 
||
| 
7726
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
143  | 
(* theory presentation *)  | 
| 
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
144  | 
|
| 9707 | 145  | 
val tex_trailer =  | 
146  | 
"%%% Local Variables:\n\  | 
|
| 9135 | 147  | 
\%%% mode: latex\n\  | 
148  | 
\%%% TeX-master: \"root\"\n\  | 
|
| 9144 | 149  | 
\%%% End:\n";  | 
| 8192 | 150  | 
|
| 9916 | 151  | 
fun isabelle_file name txt =  | 
152  | 
  "%\n\\begin{isabellebody}%\n\
 | 
|
| 10247 | 153  | 
  \\\def\\isabellecontext{" ^ output_syms name ^ "}%\n" ^ txt ^
 | 
| 9916 | 154  | 
  "\\end{isabellebody}%\n" ^ tex_trailer;
 | 
| 9707 | 155  | 
|
| 14924 | 156  | 
fun symbol_source known name syms = isabelle_file name  | 
157  | 
  ("\\isamarkupheader{" ^ output_known_symbols known (Symbol.explode name) ^ "}%\n" ^
 | 
|
158  | 
output_known_symbols known syms);  | 
|
| 
7726
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
159  | 
|
| 
9038
 
63d20536971f
session.tex: nsert blank lines in order to guarantee new paragraphs
 
wenzelm 
parents: 
8965 
diff
changeset
 | 
160  | 
fun theory_entry name = "\\input{" ^ name ^ ".tex}\n\n";
 | 
| 
7726
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
161  | 
|
| 
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
162  | 
|
| 8460 | 163  | 
(* print mode *)  | 
164  | 
||
| 8965 | 165  | 
val latexN = "latex";  | 
| 17218 | 166  | 
val modes = [latexN, Symbol.xsymbolsN];  | 
| 8965 | 167  | 
|
| 8460 | 168  | 
fun latex_output str =  | 
169  | 
let val syms = Symbol.explode str  | 
|
| 23621 | 170  | 
in (output_symbols syms, Symbol.length syms) end;  | 
| 19265 | 171  | 
|
| 23621 | 172  | 
fun latex_markup (s, _) =  | 
173  | 
  if s = Markup.keywordN then ("\\isakeyword{", "}")
 | 
|
174  | 
  else if s = Markup.commandN then ("\\isacommand{", "}")
 | 
|
| 29325 | 175  | 
else Markup.no_output;  | 
| 10955 | 176  | 
|
| 23621 | 177  | 
fun latex_indent "" _ = ""  | 
178  | 
  | latex_indent s _ = enclose "\\isaindent{" "}" s;
 | 
|
179  | 
||
180  | 
val _ = Output.add_mode latexN latex_output Symbol.encode_raw;  | 
|
| 23703 | 181  | 
val _ = Markup.add_mode latexN latex_markup;  | 
182  | 
val _ = Pretty.add_mode latexN latex_indent;  | 
|
| 8460 | 183  | 
|
| 
7726
 
2c7fc0ba1e12
Simple LaTeX presentation primitives (based on outer lexical syntax).
 
wenzelm 
parents:  
diff
changeset
 | 
184  | 
end;  |