src/Pure/Syntax/symbol_font.ML
author wenzelm
Mon, 16 Dec 1996 10:04:12 +0100
changeset 2407 f733d555b3d0
parent 2362 4744b27cdf89
child 2419 98a571307d5e
permissions -rw-r--r--
added write_charnames';
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Syntax/symbol_font.ML
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
     3
    Author:     Markus Wenzel, TU Muenchen
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
     4
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
     5
The Isabelle symbol font.
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
     6
*)
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
     7
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
     8
signature SYMBOL_FONT =
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
     9
sig
2208
39002438a79c tuned some char names;
wenzelm
parents: 2207
diff changeset
    10
  val char: string -> string option
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    11
  val name: string -> string option
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    12
  val read_charnames: string list -> string list
2407
f733d555b3d0 added write_charnames';
wenzelm
parents: 2362
diff changeset
    13
  val write_charnames:  string list -> string list	(*normal backslashes*)
f733d555b3d0 added write_charnames';
wenzelm
parents: 2362
diff changeset
    14
  val write_charnames':  string list -> string list	(*escaped backslashes*)
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    15
end;
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    16
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    17
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    18
structure SymbolFont : SYMBOL_FONT =
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    19
struct
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    20
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    21
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    22
(** font encoding vector **)
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    23
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    24
(* tables *)
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    25
2208
39002438a79c tuned some char names;
wenzelm
parents: 2207
diff changeset
    26
val enc_start = 161;
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    27
val enc_end = 255;
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    28
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    29
val enc_vector =
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    30
[
2208
39002438a79c tuned some char names;
wenzelm
parents: 2207
diff changeset
    31
      "Gamma", "Delta", "Theta", "Lambda", "Pi", "Sigma", "Phi",
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    32
  "Psi", "Omega", "alpha", "beta", "gamma", "delta", "epsilon", "zeta",
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    33
  "eta", "theta", "kappa", "lambda", "mu", "nu", "xi", "pi",
2255
f9126d306a02 improved some symbol names;
wenzelm
parents: 2208
diff changeset
    34
  "rho", "sigma", "tau", "phi", "chi", "psi", "omega", "not",
f9126d306a02 improved some symbol names;
wenzelm
parents: 2208
diff changeset
    35
  "and", "or", "forall", "exists", "And", "lceil", "rceil", "lfloor",
f9126d306a02 improved some symbol names;
wenzelm
parents: 2208
diff changeset
    36
  "rfloor", "lparr", "rparr", "lbrakk", "rbrakk", "empty", "in", "subseteq",
f9126d306a02 improved some symbol names;
wenzelm
parents: 2208
diff changeset
    37
  "inter", "union", "Inter", "Union", "sqinter", "squnion", "Sqinter", "Squnion",
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    38
  "bottom", "doteq", "equiv", "noteq", "sqsubset", "sqsubseteq", "prec", "preceq",
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    39
  "succ", "succeq", "sim", "simeq", "le", "ge", "leftarrow", "midarrow",
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    40
  "rightarrow", "Leftarrow", "Midarrow", "Rightarrow", "rrightarrow", "mapsto", "leadsto", "up",
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    41
  "down", "notin", "times", "oplus", "ominus", "otimes", "oslash", "natural",
2208
39002438a79c tuned some char names;
wenzelm
parents: 2207
diff changeset
    42
  "infinity", "box", "diamond", "circ", "bullet", "parallel", "tick", "copyright"
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    43
];
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    44
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    45
val enc_rel = enc_vector ~~ (map chr (enc_start upto enc_end));
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    46
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    47
val enc_tab = Symtab.make enc_rel;
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    48
val dec_tab = Symtab.make (map swap enc_rel);
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    49
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    50
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    51
(* chars and names *)
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    52
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    53
fun char name = Symtab.lookup (enc_tab, name);
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    54
fun name char = Symtab.lookup (dec_tab, char);
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    55
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    56
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    57
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    58
(** input and output of symbols **)
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    59
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    60
(* read_charnames *)
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    61
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    62
local
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    63
  open Scanner;
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    64
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    65
  fun scan_symbol ("\\" :: "<" :: cs) =
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    66
        let
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    67
          val (name, cs') = (scan_id -- $$ ">" >> #1) cs handle LEXICAL_ERROR
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    68
            => error ("Malformed symbolic char specification: \"\\<" ^ implode cs ^ "\"");
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    69
          val c = the (char name) handle OPTION _
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    70
            => error ("Unknown symbolic char name: " ^ quote name);
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    71
        in (c, cs') end
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    72
    | scan_symbol _ = raise LEXICAL_ERROR;
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    73
in
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    74
  val read_charnames = #1 o repeat (scan_symbol || scan_one (K true));
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    75
end;
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    76
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    77
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    78
(* write_charnames *)
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    79
2407
f733d555b3d0 added write_charnames';
wenzelm
parents: 2362
diff changeset
    80
fun write_char prfx ch =
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    81
  (case name ch of
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    82
    None => ch
2407
f733d555b3d0 added write_charnames';
wenzelm
parents: 2362
diff changeset
    83
  | Some nm => prfx ^ "\\<" ^ nm ^ ">");
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    84
2407
f733d555b3d0 added write_charnames';
wenzelm
parents: 2362
diff changeset
    85
val write_charnames = map (write_char "");
f733d555b3d0 added write_charnames';
wenzelm
parents: 2362
diff changeset
    86
val write_charnames' = map (write_char "\\");
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    87
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    88
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    89
end;