src/Pure/Isar/term_style.ML
author wenzelm
Thu, 15 Sep 2005 17:16:56 +0200
changeset 17412 e26cb20ef0cc
parent 17392 a639d580b34b
child 17496 26535df536ae
permissions -rw-r--r--
TableFun/Symtab: curried lookup and update;
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
17221
6cd180204582 curried_lookup/update;
wenzelm
parents: 16541
diff changeset
     5
Styles for terms, to use with the "term_style" and "thm_style"
6cd180204582 curried_lookup/update;
wenzelm
parents: 16541
diff changeset
     6
antiquotations.
15918
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
     7
*)
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
     8
15960
9bd6550dc004 added Proof.context to antiquotation
haftmann
parents: 15918
diff changeset
     9
signature TERM_STYLE =
15918
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    10
sig
15990
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    11
  val the_style: theory -> string -> (Proof.context -> term -> term)
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    12
  val add_style: string -> (Proof.context -> term -> term) -> theory -> theory
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    13
  val print_styles: theory -> unit
15918
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    14
end;
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    15
15960
9bd6550dc004 added Proof.context to antiquotation
haftmann
parents: 15918
diff changeset
    16
structure TermStyle: TERM_STYLE =
15918
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    17
struct
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    18
15990
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    19
(* style data *)
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    20
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    21
fun err_dup_styles names =
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    22
  error ("Duplicate declaration of antiquote style(s): " ^ commas_quote names);
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    23
16424
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 16167
diff changeset
    24
structure StyleData = TheoryDataFun
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 16167
diff changeset
    25
(struct
16541
d539d47cce69 renamed data kind;
wenzelm
parents: 16424
diff changeset
    26
  val name = "Isar/antiquote_style";
15990
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    27
  type T = ((Proof.context -> term -> term) * stamp) Symtab.table;
15960
9bd6550dc004 added Proof.context to antiquotation
haftmann
parents: 15918
diff changeset
    28
  val empty = Symtab.empty;
15918
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    29
  val copy = I;
16424
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 16167
diff changeset
    30
  val extend = I;
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 16167
diff changeset
    31
  fun merge _ tabs = Symtab.merge eq_snd tabs
15990
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    32
    handle Symtab.DUPS dups => err_dup_styles dups;
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    33
  fun print _ tab = Pretty.writeln (Pretty.strs ("antiquote styles:" :: Symtab.keys tab));
16424
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 16167
diff changeset
    34
end);
15918
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    35
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 =
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17392
diff changeset
    43
  (case Symtab.lookup (StyleData.get thy) name of
15990
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 =
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17392
diff changeset
    48
  StyleData.map (Symtab.update_new (name, (style, stamp ()))) thy
15990
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 =
16424
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 16167
diff changeset
    55
  let
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 16167
diff changeset
    56
    val concl = ObjectLogic.drop_judgment (ProofContext.theory_of ctxt)
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 16167
diff changeset
    57
      (Logic.strip_imp_concl t)
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 16167
diff changeset
    58
  in
15960
9bd6550dc004 added Proof.context to antiquotation
haftmann
parents: 15918
diff changeset
    59
    case concl of (_ $ l $ r) => (l, r)
15990
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    60
    | _ => error ("Binary operator expected in term: " ^ ProofContext.string_of_term ctxt concl)
15960
9bd6550dc004 added Proof.context to antiquotation
haftmann
parents: 15918
diff changeset
    61
  end;
15918
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    62
16167
b2e4c4058b71 renamed premise* to prem
haftmann
parents: 16165
diff changeset
    63
fun style_parm_premise i ctxt t =
16424
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 16167
diff changeset
    64
  let val prems = Logic.strip_imp_prems t in
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 16167
diff changeset
    65
    if i <= length prems then List.nth (prems, i - 1)
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 16167
diff changeset
    66
    else error ("Not enough premises for prem" ^ string_of_int i ^
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 16167
diff changeset
    67
      " in propositon: " ^ ProofContext.string_of_term ctxt t)
16160
833f4160130e added premise<i>
nipkow
parents: 15990
diff changeset
    68
  end;
16167
b2e4c4058b71 renamed premise* to prem
haftmann
parents: 16165
diff changeset
    69
15990
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    70
val _ = Context.add_setup
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    71
 [add_style "lhs" (fst oo style_binargs),
4ef32dcbb44f substantial tuning -- adapted to common conventions;
wenzelm
parents: 15961
diff changeset
    72
  add_style "rhs" (snd oo style_binargs),
16167
b2e4c4058b71 renamed premise* to prem
haftmann
parents: 16165
diff changeset
    73
  add_style "prem1" (style_parm_premise 1),
b2e4c4058b71 renamed premise* to prem
haftmann
parents: 16165
diff changeset
    74
  add_style "prem2" (style_parm_premise 2),
b2e4c4058b71 renamed premise* to prem
haftmann
parents: 16165
diff changeset
    75
  add_style "prem3" (style_parm_premise 3),
b2e4c4058b71 renamed premise* to prem
haftmann
parents: 16165
diff changeset
    76
  add_style "prem4" (style_parm_premise 4),
b2e4c4058b71 renamed premise* to prem
haftmann
parents: 16165
diff changeset
    77
  add_style "prem5" (style_parm_premise 5),
b2e4c4058b71 renamed premise* to prem
haftmann
parents: 16165
diff changeset
    78
  add_style "prem6" (style_parm_premise 6),
b2e4c4058b71 renamed premise* to prem
haftmann
parents: 16165
diff changeset
    79
  add_style "prem7" (style_parm_premise 7),
b2e4c4058b71 renamed premise* to prem
haftmann
parents: 16165
diff changeset
    80
  add_style "prem8" (style_parm_premise 8),
b2e4c4058b71 renamed premise* to prem
haftmann
parents: 16165
diff changeset
    81
  add_style "prem9" (style_parm_premise 9),
17392
a639d580b34b added prem10 - prem19
schirmer
parents: 17221
diff changeset
    82
  add_style "prem10" (style_parm_premise 10),
a639d580b34b added prem10 - prem19
schirmer
parents: 17221
diff changeset
    83
  add_style "prem11" (style_parm_premise 11),
a639d580b34b added prem10 - prem19
schirmer
parents: 17221
diff changeset
    84
  add_style "prem12" (style_parm_premise 12),
a639d580b34b added prem10 - prem19
schirmer
parents: 17221
diff changeset
    85
  add_style "prem13" (style_parm_premise 13),
a639d580b34b added prem10 - prem19
schirmer
parents: 17221
diff changeset
    86
  add_style "prem14" (style_parm_premise 14),
a639d580b34b added prem10 - prem19
schirmer
parents: 17221
diff changeset
    87
  add_style "prem15" (style_parm_premise 15),
a639d580b34b added prem10 - prem19
schirmer
parents: 17221
diff changeset
    88
  add_style "prem16" (style_parm_premise 16),
a639d580b34b added prem10 - prem19
schirmer
parents: 17221
diff changeset
    89
  add_style "prem17" (style_parm_premise 17),
a639d580b34b added prem10 - prem19
schirmer
parents: 17221
diff changeset
    90
  add_style "prem18" (style_parm_premise 18),
a639d580b34b added prem10 - prem19
schirmer
parents: 17221
diff changeset
    91
  add_style "prem19" (style_parm_premise 19),
16165
dbe9ee8ffcdd concl antiqutations
haftmann
parents: 16160
diff changeset
    92
  add_style "concl" (K Logic.strip_imp_concl)];
15918
6d6d3fabef02 final implementation of antiquotations styles
haftmann
parents:
diff changeset
    93
end;