src/Pure/Isar/term_style.ML
author wenzelm
Tue, 17 May 2005 18:10:43 +0200
changeset 15990 4ef32dcbb44f
parent 15961 24c6b96b4a2f
child 16160 833f4160130e
permissions -rw-r--r--
substantial tuning -- adapted to common conventions;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15918
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
     1
(*  Title:      Pure/Isar/term_style.ML
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
     2
    ID:         $Id$
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
     3
    Author:     Florian Haftmann, TU Muenchen
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
     4
15990
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
     5
Styles for terms, to use with the "term_style" and "thm_style" antiquotations.
15918
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
     6
*)
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
     7
15960
9bd6550dc004 added Proof.context to antiquotation
haftmann
parents: 15918
diff changeset
     8
signature TERM_STYLE =
15918
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
     9
sig
15990
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    10
  val the_style: theory -> string -> (Proof.context -> term -> term)
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    11
  val add_style: string -> (Proof.context -> term -> term) -> theory -> theory
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    12
  val print_styles: theory -> unit
15918
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    13
end;
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    14
15960
9bd6550dc004 added Proof.context to antiquotation
haftmann
parents: 15918
diff changeset
    15
structure TermStyle: TERM_STYLE =
15918
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    16
struct
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    17
15990
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    18
(* style data *)
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    19
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    20
fun err_dup_styles names =
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    21
  error ("Duplicate declaration of antiquote style(s): " ^ commas_quote names);
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    22
15918
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    23
structure StyleArgs =
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    24
struct
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    25
  val name = "Isar/style";
15990
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    26
  type T = ((Proof.context -> term -> term) * stamp) Symtab.table;
15960
9bd6550dc004 added Proof.context to antiquotation
haftmann
parents: 15918
diff changeset
    27
  val empty = Symtab.empty;
15918
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    28
  val copy = I;
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    29
  val prep_ext = I;
15990
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    30
  fun merge tabs = Symtab.merge eq_snd tabs
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    31
    handle Symtab.DUPS dups => err_dup_styles dups;
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    32
  fun print _ tab = Pretty.writeln (Pretty.strs ("antiquote styles:" :: Symtab.keys tab));
15918
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    33
end;
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    34
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    35
structure StyleData = TheoryDataFun(StyleArgs);
15990
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    36
val _ = Context.add_setup [StyleData.init];
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    37
val print_styles = StyleData.print;
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    38
15918
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    39
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    40
(* accessors *)
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    41
15990
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    42
fun the_style thy name =
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    43
  (case Symtab.lookup (StyleData.get thy, name) of
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    44
    NONE => error ("Unknown antiquote style: " ^ quote name)
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    45
  | SOME (style, _) => style);
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    46
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    47
fun add_style name style thy =
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    48
  StyleData.map (curry Symtab.update_new (name, (style, stamp ()))) thy
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    49
    handle Symtab.DUP _ => err_dup_styles [name];
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    50
15918
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    51
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    52
(* predefined styles *)
15990
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    53
15960
9bd6550dc004 added Proof.context to antiquotation
haftmann
parents: 15918
diff changeset
    54
fun style_binargs ctxt t =
9bd6550dc004 added Proof.context to antiquotation
haftmann
parents: 15918
diff changeset
    55
  let val concl = ObjectLogic.drop_judgment (ProofContext.sign_of ctxt) (Logic.strip_imp_concl t) in
9bd6550dc004 added Proof.context to antiquotation
haftmann
parents: 15918
diff changeset
    56
    case concl of (_ $ l $ r) => (l, r)
15990
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    57
    | _ => error ("Binary operator expected in term: " ^ ProofContext.string_of_term ctxt concl)
15960
9bd6550dc004 added Proof.context to antiquotation
haftmann
parents: 15918
diff changeset
    58
  end;
15918
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    59
15990
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    60
val _ = Context.add_setup
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    61
 [add_style "lhs" (fst oo style_binargs),
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    62
  add_style "rhs" (snd oo style_binargs),
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    63
  add_style "conclusion" (K Logic.strip_imp_concl)];
15918
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    64
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    65
end;