author | wenzelm |
Wed, 04 Dec 2013 18:59:20 +0100 | |
changeset 54667 | 4dd08fe126ba |
parent 53171 | a5e54d4d9081 |
child 56029 | 8bedca4bd5a3 |
permissions | -rw-r--r-- |
22107 | 1 |
(* Title: Pure/Thy/term_style.ML |
2 |
Author: Florian Haftmann, TU Muenchen |
|
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 | 5 |
*) |
6 |
||
7 |
signature TERM_STYLE = |
|
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 |
22107 | 11 |
end; |
12 |
||
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
|
13 |
structure Term_Style: TERM_STYLE = |
22107 | 14 |
struct |
15 |
||
16 |
(* style data *) |
|
17 |
||
23655
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
23577
diff
changeset
|
18 |
fun err_dup_style name = |
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
23577
diff
changeset
|
19 |
error ("Duplicate declaration of antiquote style: " ^ quote name); |
22107 | 20 |
|
42016 | 21 |
structure Styles = Theory_Data |
22846 | 22 |
( |
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
|
23 |
type T = ((Proof.context -> term -> term) parser * stamp) Symtab.table; |
22107 | 24 |
val empty = Symtab.empty; |
25 |
val extend = I; |
|
33522 | 26 |
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
|
27 |
handle Symtab.DUP dup => err_dup_style dup; |
22846 | 28 |
); |
22107 | 29 |
|
30 |
||
31 |
(* accessors *) |
|
32 |
||
33 |
fun the_style thy name = |
|
42016 | 34 |
(case Symtab.lookup (Styles.get thy) name of |
22107 | 35 |
NONE => error ("Unknown antiquote style: " ^ quote name) |
36 |
| SOME (style, _) => style); |
|
37 |
||
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
|
38 |
fun setup name style thy = |
42016 | 39 |
Styles.map (Symtab.update_new (name, (style, stamp ()))) thy |
23655
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
23577
diff
changeset
|
40 |
handle Symtab.DUP _ => err_dup_style name; |
22107 | 41 |
|
42 |
||
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
|
43 |
(* 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
|
44 |
|
36950 | 45 |
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
|
46 |
>> (fn x as ((name, _), _) => fst (Args.context_syntax "style" |
42360 | 47 |
(Scan.lift (the_style (Proof_Context.theory_of ctxt) 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
|
48 |
(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
|
49 |
|
50637 | 50 |
val parse = Args.context :|-- (fn ctxt => Scan.lift |
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
|
51 |
(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
|
52 |
>> 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
|
53 |
|| 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
|
54 |
|
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 |
|
22107 | 56 |
(* predefined styles *) |
57 |
||
32891 | 58 |
fun style_lhs_rhs proj = Scan.succeed (fn ctxt => fn t => |
22107 | 59 |
let |
35625 | 60 |
val concl = |
42360 | 61 |
Object_Logic.drop_judgment (Proof_Context.theory_of ctxt) (Logic.strip_imp_concl t) |
22107 | 62 |
in |
50637 | 63 |
(case concl of |
64 |
(_ $ l $ r) => proj (l, r) |
|
65 |
| _ => 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
|
66 |
end); |
22107 | 67 |
|
50638 | 68 |
val style_prem = Parse.nat >> (fn i => fn ctxt => fn t => |
32891 | 69 |
let |
70 |
val prems = Logic.strip_imp_prems t; |
|
71 |
in |
|
72 |
if i <= length prems then nth prems (i - 1) |
|
50637 | 73 |
else |
74 |
error ("Not enough premises for prem " ^ string_of_int i ^ |
|
75 |
" in propositon: " ^ Syntax.string_of_term ctxt t) |
|
32891 | 76 |
end); |
77 |
||
53021
d0fa3f446b9d
discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents:
50638
diff
changeset
|
78 |
fun sub_symbols (d :: s :: ss) = |
41685
e29ea98a76ce
term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents:
36950
diff
changeset
|
79 |
if Symbol.is_ascii_digit d andalso not (String.isPrefix ("\\<^") s) |
53021
d0fa3f446b9d
discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents:
50638
diff
changeset
|
80 |
then d :: "\\<^sub>" :: sub_symbols (s :: ss) |
41685
e29ea98a76ce
term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents:
36950
diff
changeset
|
81 |
else d :: s :: ss |
53021
d0fa3f446b9d
discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents:
50638
diff
changeset
|
82 |
| sub_symbols cs = cs; |
41685
e29ea98a76ce
term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents:
36950
diff
changeset
|
83 |
|
53021
d0fa3f446b9d
discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents:
50638
diff
changeset
|
84 |
val sub_name = implode o rev o sub_symbols o rev o Symbol.explode; |
41685
e29ea98a76ce
term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents:
36950
diff
changeset
|
85 |
|
53021
d0fa3f446b9d
discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents:
50638
diff
changeset
|
86 |
fun sub_term (Free (n, T)) = Free (sub_name n, T) |
d0fa3f446b9d
discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents:
50638
diff
changeset
|
87 |
| sub_term (Var ((n, idx), T)) = |
d0fa3f446b9d
discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents:
50638
diff
changeset
|
88 |
if idx <> 0 then Var ((sub_name (n ^ string_of_int idx), 0), T) |
d0fa3f446b9d
discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents:
50638
diff
changeset
|
89 |
else Var ((sub_name n, 0), T) |
d0fa3f446b9d
discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents:
50638
diff
changeset
|
90 |
| sub_term (t $ u) = sub_term t $ sub_term u |
d0fa3f446b9d
discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents:
50638
diff
changeset
|
91 |
| sub_term (Abs (n, T, b)) = Abs (sub_name n, T, sub_term b) |
d0fa3f446b9d
discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents:
50638
diff
changeset
|
92 |
| sub_term t = t; |
41685
e29ea98a76ce
term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents:
36950
diff
changeset
|
93 |
|
53171 | 94 |
val _ = Theory.setup |
32891 | 95 |
(setup "lhs" (style_lhs_rhs fst) #> |
96 |
setup "rhs" (style_lhs_rhs snd) #> |
|
97 |
setup "prem" style_prem #> |
|
98 |
setup "concl" (Scan.succeed (K Logic.strip_imp_concl)) #> |
|
53171 | 99 |
setup "sub" (Scan.succeed (K sub_term))); |
22107 | 100 |
|
101 |
end; |