src/Pure/Syntax/symbol_font.ML
author wenzelm
Fri, 07 Mar 1997 15:32:16 +0100
changeset 2769 77903c147673
parent 2676 585cd2311a98
child 2913 ce271fa4d8e2
permissions -rw-r--r--
removed lparr, rparr, empty, succeq, ge, rrightarrow; added turnstile, Turnstile, cdot, approx, Colon, bow; renamed tick to surd;
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
2676
585cd2311a98 replaced natural by subset;
wenzelm
parents: 2423
diff changeset
     5
The Isabelle symbol font.  Please keep enc_vector consistent with
585cd2311a98 replaced natural by subset;
wenzelm
parents: 2423
diff changeset
     6
lib/scripts/symbolinput.pl!
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
     7
*)
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
     8
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
     9
signature SYMBOL_FONT =
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    10
sig
2208
39002438a79c tuned some char names;
wenzelm
parents: 2207
diff changeset
    11
  val char: string -> string option
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    12
  val name: string -> string option
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    13
  val read_charnames: string list -> string list
2407
f733d555b3d0 added write_charnames';
wenzelm
parents: 2362
diff changeset
    14
  val write_charnames:  string list -> string list	(*normal backslashes*)
f733d555b3d0 added write_charnames';
wenzelm
parents: 2362
diff changeset
    15
  val write_charnames':  string list -> string list	(*escaped backslashes*)
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    16
end;
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    17
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    18
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    19
structure SymbolFont : SYMBOL_FONT =
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    20
struct
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    21
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    22
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    23
(** font encoding vector **)
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    24
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    25
(* tables *)
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    26
2208
39002438a79c tuned some char names;
wenzelm
parents: 2207
diff changeset
    27
val enc_start = 161;
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    28
val enc_end = 255;
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    29
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    30
val enc_vector =
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    31
[
2208
39002438a79c tuned some char names;
wenzelm
parents: 2207
diff changeset
    32
      "Gamma", "Delta", "Theta", "Lambda", "Pi", "Sigma", "Phi",
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    33
  "Psi", "Omega", "alpha", "beta", "gamma", "delta", "epsilon", "zeta",
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    34
  "eta", "theta", "kappa", "lambda", "mu", "nu", "xi", "pi",
2255
f9126d306a02 improved some symbol names;
wenzelm
parents: 2208
diff changeset
    35
  "rho", "sigma", "tau", "phi", "chi", "psi", "omega", "not",
f9126d306a02 improved some symbol names;
wenzelm
parents: 2208
diff changeset
    36
  "and", "or", "forall", "exists", "And", "lceil", "rceil", "lfloor",
2769
77903c147673 removed lparr, rparr, empty, succeq, ge, rrightarrow;
wenzelm
parents: 2676
diff changeset
    37
  "rfloor", "turnstile", "Turnstile", "lbrakk", "rbrakk", "cdot", "in", "subseteq",
2255
f9126d306a02 improved some symbol names;
wenzelm
parents: 2208
diff changeset
    38
  "inter", "union", "Inter", "Union", "sqinter", "squnion", "Sqinter", "Squnion",
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    39
  "bottom", "doteq", "equiv", "noteq", "sqsubset", "sqsubseteq", "prec", "preceq",
2769
77903c147673 removed lparr, rparr, empty, succeq, ge, rrightarrow;
wenzelm
parents: 2676
diff changeset
    40
  "succ", "approx", "sim", "simeq", "le", "Colon", "leftarrow", "midarrow",
77903c147673 removed lparr, rparr, empty, succeq, ge, rrightarrow;
wenzelm
parents: 2676
diff changeset
    41
  "rightarrow", "Leftarrow", "Midarrow", "Rightarrow", "bow", "mapsto", "leadsto", "up",
2676
585cd2311a98 replaced natural by subset;
wenzelm
parents: 2423
diff changeset
    42
  "down", "notin", "times", "oplus", "ominus", "otimes", "oslash", "subset",
2769
77903c147673 removed lparr, rparr, empty, succeq, ge, rrightarrow;
wenzelm
parents: 2676
diff changeset
    43
  "infinity", "box", "diamond", "circ", "bullet", "parallel", "surd", "copyright"
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    44
];
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    45
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    46
val enc_rel = enc_vector ~~ (map chr (enc_start upto enc_end));
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    47
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    48
val enc_tab = Symtab.make enc_rel;
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    49
val dec_tab = Symtab.make (map swap enc_rel);
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    50
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    51
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    52
(* chars and names *)
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    53
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    54
fun char name = Symtab.lookup (enc_tab, name);
2419
98a571307d5e tuned read and write functions;
wenzelm
parents: 2407
diff changeset
    55
98a571307d5e tuned read and write functions;
wenzelm
parents: 2407
diff changeset
    56
fun name char =
98a571307d5e tuned read and write functions;
wenzelm
parents: 2407
diff changeset
    57
  if ord char < enc_start then None
98a571307d5e tuned read and write functions;
wenzelm
parents: 2407
diff changeset
    58
  else Symtab.lookup (dec_tab, char);
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    59
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    60
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    61
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    62
(** input and output of symbols **)
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    63
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    64
(* read_charnames *)
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    65
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    66
local
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    67
  open Scanner;
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    68
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    69
  fun scan_symbol ("\\" :: "<" :: cs) =
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    70
        let
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    71
          val (name, cs') = (scan_id -- $$ ">" >> #1) cs handle LEXICAL_ERROR
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    72
            => error ("Malformed symbolic char specification: \"\\<" ^ implode cs ^ "\"");
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    73
          val c = the (char name) handle OPTION _
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    74
            => error ("Unknown symbolic char name: " ^ quote name);
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    75
        in (c, cs') end
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    76
    | scan_symbol _ = raise LEXICAL_ERROR;
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    77
in
2419
98a571307d5e tuned read and write functions;
wenzelm
parents: 2407
diff changeset
    78
  fun read_charnames chs =
98a571307d5e tuned read and write functions;
wenzelm
parents: 2407
diff changeset
    79
    if forall (not_equal "\\") chs then chs
98a571307d5e tuned read and write functions;
wenzelm
parents: 2407
diff changeset
    80
    else #1 (repeat (scan_symbol || scan_one (K true)) chs);
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    81
end;
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    82
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    83
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    84
(* write_charnames *)
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    85
2407
f733d555b3d0 added write_charnames';
wenzelm
parents: 2362
diff changeset
    86
fun write_char prfx ch =
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    87
  (case name ch of
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    88
    None => ch
2407
f733d555b3d0 added write_charnames';
wenzelm
parents: 2362
diff changeset
    89
  | Some nm => prfx ^ "\\<" ^ nm ^ ">");
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    90
2419
98a571307d5e tuned read and write functions;
wenzelm
parents: 2407
diff changeset
    91
fun write_chars prfx chs =
98a571307d5e tuned read and write functions;
wenzelm
parents: 2407
diff changeset
    92
  if forall (fn ch => ord ch < enc_start) chs then chs
98a571307d5e tuned read and write functions;
wenzelm
parents: 2407
diff changeset
    93
  else map (write_char prfx) chs;
98a571307d5e tuned read and write functions;
wenzelm
parents: 2407
diff changeset
    94
98a571307d5e tuned read and write functions;
wenzelm
parents: 2407
diff changeset
    95
val write_charnames = write_chars "";
98a571307d5e tuned read and write functions;
wenzelm
parents: 2407
diff changeset
    96
val write_charnames' = write_chars "\\";
2362
4744b27cdf89 added read_charnames, write_charnames;
wenzelm
parents: 2255
diff changeset
    97
2207
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    98
f7f06d4deaa2 added this file;
wenzelm
parents:
diff changeset
    99
end;