src/Pure/Thy/term_style.ML
author wenzelm
Sat, 15 May 2010 23:16:32 +0200
changeset 36950 75b8f26f2f07
parent 35625 9c818cab0dd0
child 41685 e29ea98a76ce
permissions -rw-r--r--
refer directly to structure Keyword and Parse; eliminated old-style structure aliases K and P;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Thy/term_style.ML
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
     3
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
     4
Styles for term printing.
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
     5
*)
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
     6
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
     7
signature TERM_STYLE =
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
     8
sig
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
     9
  val setup: string -> (Proof.context -> term -> term) parser -> theory -> theory
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    10
  val parse: (term -> term) context_parser
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    11
  val parse_bare: (term -> term) context_parser
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    12
end;
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    13
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    14
structure Term_Style: TERM_STYLE =
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    15
struct
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    16
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    17
(* style data *)
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    18
23655
d2d1138e0ddc replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents: 23577
diff changeset
    19
fun err_dup_style name =
d2d1138e0ddc replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents: 23577
diff changeset
    20
  error ("Duplicate declaration of antiquote style: " ^ quote name);
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    21
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33184
diff changeset
    22
structure StyleData = Theory_Data
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22107
diff changeset
    23
(
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    24
  type T = ((Proof.context -> term -> term) parser * stamp) Symtab.table;
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    25
  val empty = Symtab.empty;
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    26
  val extend = I;
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33184
diff changeset
    27
  fun merge data : T = Symtab.merge (eq_snd (op =)) data
23655
d2d1138e0ddc replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents: 23577
diff changeset
    28
    handle Symtab.DUP dup => err_dup_style dup;
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22107
diff changeset
    29
);
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    30
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    31
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    32
(* accessors *)
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    33
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    34
fun the_style thy name =
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    35
  (case Symtab.lookup (StyleData.get thy) name of
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    36
    NONE => error ("Unknown antiquote style: " ^ quote name)
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    37
  | SOME (style, _) => style);
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    38
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    39
fun setup name style thy =
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    40
  StyleData.map (Symtab.update_new (name, (style, stamp ()))) thy
23655
d2d1138e0ddc replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents: 23577
diff changeset
    41
    handle Symtab.DUP _ => err_dup_style name;
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    42
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    43
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    44
(* style parsing *)
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    45
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 35625
diff changeset
    46
fun parse_single ctxt = Parse.position (Parse.xname -- Args.parse)
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    47
  >> (fn x as ((name, _), _) => fst (Args.context_syntax "style"
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    48
       (Scan.lift (the_style (ProofContext.theory_of ctxt) name))
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    49
         (Args.src x) ctxt |>> (fn f => f ctxt)));
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    50
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    51
val parse = Args.context :|-- (fn ctxt => Scan.lift 
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    52
  (Args.parens (parse_single ctxt ::: Scan.repeat (Args.$$$ "," |-- parse_single ctxt))
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    53
      >> fold I
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    54
  || Scan.succeed I));
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    55
33184
de8cc01e8d9e legacy warnings for old-style term styles
haftmann
parents: 32891
diff changeset
    56
val parse_bare = Args.context :|-- (fn ctxt => (Output.legacy_feature "Old-style antiquotation style.";
de8cc01e8d9e legacy warnings for old-style term styles
haftmann
parents: 32891
diff changeset
    57
  Scan.lift Args.liberal_name
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    58
  >> (fn name => fst (Args.context_syntax "style"
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    59
       (Scan.lift (the_style (ProofContext.theory_of ctxt) name))
33184
de8cc01e8d9e legacy warnings for old-style term styles
haftmann
parents: 32891
diff changeset
    60
          (Args.src (("style", []), Position.none)) ctxt |>> (fn f => f ctxt)))));
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    61
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    62
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    63
(* predefined styles *)
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    64
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    65
fun style_lhs_rhs proj = Scan.succeed (fn ctxt => fn t =>
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    66
  let
35625
9c818cab0dd0 modernized structure Object_Logic;
wenzelm
parents: 33522
diff changeset
    67
    val concl =
9c818cab0dd0 modernized structure Object_Logic;
wenzelm
parents: 33522
diff changeset
    68
      Object_Logic.drop_judgment (ProofContext.theory_of ctxt) (Logic.strip_imp_concl t)
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    69
  in
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    70
    case concl of (_ $ l $ r) => proj (l, r)
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 23655
diff changeset
    71
    | _ => error ("Binary operator expected in term: " ^ Syntax.string_of_term ctxt concl)
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    72
  end);
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    73
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    74
val style_prem = Args.name >> (fn raw_i => fn ctxt => fn t =>
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    75
  let
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    76
    val i = (the o Int.fromString) raw_i;
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    77
    val prems = Logic.strip_imp_prems t;
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    78
  in
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    79
    if i <= length prems then nth prems (i - 1)
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    80
    else error ("Not enough premises for prem " ^ string_of_int i ^
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    81
      " in propositon: " ^ Syntax.string_of_term ctxt t)
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    82
  end);
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    83
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    84
fun style_parm_premise i = Scan.succeed (fn ctxt => fn t =>
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    85
  let
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    86
    val i_str = string_of_int i;
33184
de8cc01e8d9e legacy warnings for old-style term styles
haftmann
parents: 32891
diff changeset
    87
    val _ = Output.legacy_feature (quote ("prem" ^ i_str)
de8cc01e8d9e legacy warnings for old-style term styles
haftmann
parents: 32891
diff changeset
    88
      ^ " term style encountered; use explicit argument syntax "
de8cc01e8d9e legacy warnings for old-style term styles
haftmann
parents: 32891
diff changeset
    89
      ^ quote ("prem " ^ i_str) ^ " instead.");
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    90
    val prems = Logic.strip_imp_prems t;
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    91
  in
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    92
    if i <= length prems then nth prems (i - 1)
33184
de8cc01e8d9e legacy warnings for old-style term styles
haftmann
parents: 32891
diff changeset
    93
    else error ("Not enough premises for prem" ^ i_str ^
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 23655
diff changeset
    94
      " in propositon: " ^ Syntax.string_of_term ctxt t)
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    95
  end);
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    96
26463
9283b4185fdf Context.>> : operate on Context.generic;
wenzelm
parents: 26435
diff changeset
    97
val _ = Context.>> (Context.map_theory
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    98
 (setup "lhs" (style_lhs_rhs fst) #>
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    99
  setup "rhs" (style_lhs_rhs snd) #>
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
   100
  setup "prem" style_prem #>
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
   101
  setup "concl" (Scan.succeed (K Logic.strip_imp_concl)) #>
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
   102
  setup "prem1" (style_parm_premise 1) #>
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
   103
  setup "prem2" (style_parm_premise 2) #>
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
   104
  setup "prem3" (style_parm_premise 3) #>
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
   105
  setup "prem4" (style_parm_premise 4) #>
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
   106
  setup "prem5" (style_parm_premise 5) #>
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
   107
  setup "prem6" (style_parm_premise 6) #>
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
   108
  setup "prem7" (style_parm_premise 7) #>
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
   109
  setup "prem8" (style_parm_premise 8) #>
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
   110
  setup "prem9" (style_parm_premise 9) #>
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
   111
  setup "prem10" (style_parm_premise 10) #>
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
   112
  setup "prem11" (style_parm_premise 11) #>
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
   113
  setup "prem12" (style_parm_premise 12) #>
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
   114
  setup "prem13" (style_parm_premise 13) #>
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
   115
  setup "prem14" (style_parm_premise 14) #>
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
   116
  setup "prem15" (style_parm_premise 15) #>
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
   117
  setup "prem16" (style_parm_premise 16) #>
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
   118
  setup "prem17" (style_parm_premise 17) #>
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
   119
  setup "prem18" (style_parm_premise 18) #>
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
   120
  setup "prem19" (style_parm_premise 19)));
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
   121
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
   122
end;