author | wenzelm |
Tue, 23 Dec 2008 00:56:03 +0100 | |
changeset 29152 | 89b0803404d7 |
parent 28708 | a1a436f09ec6 |
child 30161 | c26e515f1c29 |
permissions | -rw-r--r-- |
28060 | 1 |
(* Title: Tools/code/code_printer.ML |
2 |
ID: $Id$ |
|
3 |
Author: Florian Haftmann, TU Muenchen |
|
4 |
||
5 |
Generic operations for pretty printing of target language code. |
|
6 |
*) |
|
7 |
||
8 |
signature CODE_PRINTER = |
|
9 |
sig |
|
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28064
diff
changeset
|
10 |
val nerror: thm -> string -> 'a |
28060 | 11 |
|
12 |
val @@ : 'a * 'a -> 'a list |
|
13 |
val @| : 'a list * 'a -> 'a list |
|
14 |
val str: string -> Pretty.T |
|
15 |
val concat: Pretty.T list -> Pretty.T |
|
16 |
val brackets: Pretty.T list -> Pretty.T |
|
17 |
val semicolon: Pretty.T list -> Pretty.T |
|
18 |
val enum_default: string -> string -> string -> string -> Pretty.T list -> Pretty.T |
|
19 |
||
20 |
type lrx |
|
21 |
val L: lrx |
|
22 |
val R: lrx |
|
23 |
val X: lrx |
|
24 |
type fixity |
|
25 |
val BR: fixity |
|
26 |
val NOBR: fixity |
|
27 |
val INFX: int * lrx -> fixity |
|
28 |
val APP: fixity |
|
29 |
val brackify: fixity -> Pretty.T list -> Pretty.T |
|
30 |
val brackify_infix: int * lrx -> fixity -> Pretty.T list -> Pretty.T |
|
31 |
||
32 |
type itype = Code_Thingol.itype |
|
33 |
type iterm = Code_Thingol.iterm |
|
34 |
type const = Code_Thingol.const |
|
35 |
type dict = Code_Thingol.dict |
|
36 |
type tyco_syntax |
|
37 |
type const_syntax |
|
38 |
val parse_infix: ('a -> 'b) -> lrx * int -> string |
|
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28064
diff
changeset
|
39 |
-> int * ((fixity -> 'b -> Pretty.T) |
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28064
diff
changeset
|
40 |
-> fixity -> 'a list -> Pretty.T) |
28060 | 41 |
val parse_syntax: ('a -> 'b) -> OuterParse.token list |
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28064
diff
changeset
|
42 |
-> (int * ((fixity -> 'b -> Pretty.T) |
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28064
diff
changeset
|
43 |
-> fixity -> 'a list -> Pretty.T)) option * OuterParse.token list |
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28064
diff
changeset
|
44 |
val simple_const_syntax: (int * ((fixity -> iterm -> Pretty.T) |
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28064
diff
changeset
|
45 |
-> fixity -> (iterm * itype) list -> Pretty.T)) option -> const_syntax option |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28690
diff
changeset
|
46 |
val gen_pr_app: (thm -> Code_Name.var_ctxt -> const * iterm list -> Pretty.T list) |
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28690
diff
changeset
|
47 |
-> (thm -> Code_Name.var_ctxt -> fixity -> iterm -> Pretty.T) |
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28690
diff
changeset
|
48 |
-> (string -> const_syntax option) -> Code_Thingol.naming |
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28690
diff
changeset
|
49 |
-> thm -> Code_Name.var_ctxt -> fixity -> const * iterm list -> Pretty.T |
28060 | 50 |
val gen_pr_bind: ((string option * Pretty.T option) * itype -> Pretty.T) |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28690
diff
changeset
|
51 |
-> (thm -> Code_Name.var_ctxt -> fixity -> iterm -> Pretty.T) |
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28064
diff
changeset
|
52 |
-> thm -> fixity |
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28064
diff
changeset
|
53 |
-> (string option * iterm option) * itype -> Code_Name.var_ctxt -> Pretty.T * Code_Name.var_ctxt |
28060 | 54 |
|
28064 | 55 |
type literals |
56 |
val Literals: { literal_char: string -> string, literal_string: string -> string, |
|
57 |
literal_numeral: bool -> int -> string, literal_list: Pretty.T list -> Pretty.T, infix_cons: int * string } |
|
58 |
-> literals |
|
59 |
val literal_char: literals -> string -> string |
|
60 |
val literal_string: literals -> string -> string |
|
61 |
val literal_numeral: literals -> bool -> int -> string |
|
62 |
val literal_list: literals -> Pretty.T list -> Pretty.T |
|
63 |
val infix_cons: literals -> int * string |
|
28060 | 64 |
end; |
65 |
||
66 |
structure Code_Printer : CODE_PRINTER = |
|
67 |
struct |
|
68 |
||
69 |
open Code_Thingol; |
|
70 |
||
71 |
fun nerror thm s = error (s ^ ",\nin equation " ^ Display.string_of_thm thm); |
|
72 |
||
73 |
(** assembling text pieces **) |
|
74 |
||
75 |
infixr 5 @@; |
|
76 |
infixr 5 @|; |
|
77 |
fun x @@ y = [x, y]; |
|
78 |
fun xs @| y = xs @ [y]; |
|
79 |
val str = PrintMode.setmp [] Pretty.str; |
|
80 |
val concat = Pretty.block o Pretty.breaks; |
|
81 |
val brackets = Pretty.enclose "(" ")" o Pretty.breaks; |
|
82 |
fun semicolon ps = Pretty.block [concat ps, str ";"]; |
|
83 |
fun enum_default default sep opn cls [] = str default |
|
84 |
| enum_default default sep opn cls xs = Pretty.enum sep opn cls xs; |
|
85 |
||
86 |
||
87 |
(** syntax printer **) |
|
88 |
||
89 |
(* binding priorities *) |
|
90 |
||
91 |
datatype lrx = L | R | X; |
|
92 |
||
93 |
datatype fixity = |
|
94 |
BR |
|
95 |
| NOBR |
|
96 |
| INFX of (int * lrx); |
|
97 |
||
98 |
val APP = INFX (~1, L); |
|
99 |
||
100 |
fun fixity_lrx L L = false |
|
101 |
| fixity_lrx R R = false |
|
102 |
| fixity_lrx _ _ = true; |
|
103 |
||
104 |
fun fixity NOBR _ = false |
|
105 |
| fixity _ NOBR = false |
|
106 |
| fixity (INFX (pr, lr)) (INFX (pr_ctxt, lr_ctxt)) = |
|
107 |
pr < pr_ctxt |
|
108 |
orelse pr = pr_ctxt |
|
109 |
andalso fixity_lrx lr lr_ctxt |
|
110 |
orelse pr_ctxt = ~1 |
|
111 |
| fixity BR (INFX _) = false |
|
112 |
| fixity _ _ = true; |
|
113 |
||
114 |
fun gen_brackify _ [p] = p |
|
115 |
| gen_brackify true (ps as _::_) = Pretty.enclose "(" ")" ps |
|
116 |
| gen_brackify false (ps as _::_) = Pretty.block ps; |
|
117 |
||
118 |
fun brackify fxy_ctxt = |
|
119 |
gen_brackify (fixity BR fxy_ctxt) o Pretty.breaks; |
|
120 |
||
121 |
fun brackify_infix infx fxy_ctxt = |
|
122 |
gen_brackify (fixity (INFX infx) fxy_ctxt) o Pretty.breaks; |
|
123 |
||
124 |
||
125 |
(* generic syntax *) |
|
126 |
||
127 |
type tyco_syntax = int * ((fixity -> itype -> Pretty.T) |
|
128 |
-> fixity -> itype list -> Pretty.T); |
|
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28064
diff
changeset
|
129 |
type const_syntax = int * ((Code_Name.var_ctxt -> fixity -> iterm -> Pretty.T) |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28690
diff
changeset
|
130 |
-> Code_Thingol.naming -> thm -> Code_Name.var_ctxt -> fixity -> (iterm * itype) list -> Pretty.T); |
28060 | 131 |
|
132 |
fun simple_const_syntax x = (Option.map o apsnd) |
|
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28690
diff
changeset
|
133 |
(fn pretty => fn pr => fn naming => fn thm => fn vars => pretty (pr vars)) x; |
28060 | 134 |
|
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28690
diff
changeset
|
135 |
fun gen_pr_app pr_app pr_term syntax_const naming thm vars fxy (app as ((c, (_, tys)), ts)) = |
28060 | 136 |
case syntax_const c |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28690
diff
changeset
|
137 |
of NONE => brackify fxy (pr_app thm vars app) |
28060 | 138 |
| SOME (k, pr) => |
139 |
let |
|
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28690
diff
changeset
|
140 |
fun pr' fxy ts = pr (pr_term thm) naming thm vars fxy (ts ~~ curry Library.take k tys); |
28060 | 141 |
in if k = length ts |
142 |
then pr' fxy ts |
|
143 |
else if k < length ts |
|
144 |
then case chop k ts of (ts1, ts2) => |
|
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28690
diff
changeset
|
145 |
brackify fxy (pr' APP ts1 :: map (pr_term thm vars BR) ts2) |
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28690
diff
changeset
|
146 |
else pr_term thm vars fxy (Code_Thingol.eta_expand k app) |
28060 | 147 |
end; |
148 |
||
149 |
fun gen_pr_bind pr_bind pr_term thm (fxy : fixity) ((v, pat), ty : itype) vars = |
|
150 |
let |
|
151 |
val vs = case pat |
|
152 |
of SOME pat => Code_Thingol.fold_varnames (insert (op =)) pat [] |
|
153 |
| NONE => []; |
|
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28064
diff
changeset
|
154 |
val vars' = Code_Name.intro_vars (the_list v) vars; |
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28064
diff
changeset
|
155 |
val vars'' = Code_Name.intro_vars vs vars'; |
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28064
diff
changeset
|
156 |
val v' = Option.map (Code_Name.lookup_var vars') v; |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28690
diff
changeset
|
157 |
val pat' = Option.map (pr_term thm vars'' fxy) pat; |
28060 | 158 |
in (pr_bind ((v', pat'), ty), vars'') end; |
159 |
||
160 |
||
161 |
(* mixfix syntax *) |
|
162 |
||
163 |
datatype 'a mixfix = |
|
164 |
Arg of fixity |
|
165 |
| Pretty of Pretty.T; |
|
166 |
||
167 |
fun mk_mixfix prep_arg (fixity_this, mfx) = |
|
168 |
let |
|
169 |
fun is_arg (Arg _) = true |
|
170 |
| is_arg _ = false; |
|
171 |
val i = (length o filter is_arg) mfx; |
|
172 |
fun fillin _ [] [] = |
|
173 |
[] |
|
174 |
| fillin pr (Arg fxy :: mfx) (a :: args) = |
|
175 |
(pr fxy o prep_arg) a :: fillin pr mfx args |
|
176 |
| fillin pr (Pretty p :: mfx) args = |
|
177 |
p :: fillin pr mfx args; |
|
178 |
in |
|
179 |
(i, fn pr => fn fixity_ctxt => fn args => |
|
180 |
gen_brackify (fixity fixity_this fixity_ctxt) (fillin pr mfx args)) |
|
181 |
end; |
|
182 |
||
183 |
fun parse_infix prep_arg (x, i) s = |
|
184 |
let |
|
185 |
val l = case x of L => INFX (i, L) | _ => INFX (i, X); |
|
186 |
val r = case x of R => INFX (i, R) | _ => INFX (i, X); |
|
187 |
in |
|
188 |
mk_mixfix prep_arg (INFX (i, x), |
|
189 |
[Arg l, (Pretty o Pretty.brk) 1, (Pretty o str) s, (Pretty o Pretty.brk) 1, Arg r]) |
|
190 |
end; |
|
191 |
||
192 |
fun parse_mixfix prep_arg s = |
|
193 |
let |
|
194 |
val sym_any = Scan.one Symbol.is_regular; |
|
195 |
val parse = Scan.optional ($$ "!" >> K true) false -- Scan.repeat ( |
|
196 |
($$ "(" -- $$ "_" -- $$ ")" >> K (Arg NOBR)) |
|
197 |
|| ($$ "_" >> K (Arg BR)) |
|
198 |
|| ($$ "/" |-- Scan.repeat ($$ " ") >> (Pretty o Pretty.brk o length)) |
|
199 |
|| (Scan.repeat1 |
|
200 |
( $$ "'" |-- sym_any |
|
201 |
|| Scan.unless ($$ "_" || $$ "/" || $$ "(" |-- $$ "_" |-- $$ ")") |
|
202 |
sym_any) >> (Pretty o str o implode))); |
|
203 |
in case Scan.finite Symbol.stopper parse (Symbol.explode s) |
|
204 |
of ((_, p as [_]), []) => mk_mixfix prep_arg (NOBR, p) |
|
205 |
| ((b, p as _ :: _ :: _), []) => mk_mixfix prep_arg (if b then NOBR else BR, p) |
|
206 |
| _ => Scan.!! |
|
207 |
(the_default ("malformed mixfix annotation: " ^ quote s) o snd) Scan.fail () |
|
208 |
end; |
|
209 |
||
210 |
val (infixK, infixlK, infixrK) = ("infix", "infixl", "infixr"); |
|
211 |
||
212 |
fun parse_syntax prep_arg xs = |
|
213 |
Scan.option (( |
|
214 |
((OuterParse.$$$ infixK >> K X) |
|
215 |
|| (OuterParse.$$$ infixlK >> K L) |
|
216 |
|| (OuterParse.$$$ infixrK >> K R)) |
|
217 |
-- OuterParse.nat >> parse_infix prep_arg |
|
218 |
|| Scan.succeed (parse_mixfix prep_arg)) |
|
219 |
-- OuterParse.string |
|
220 |
>> (fn (parse, s) => parse s)) xs; |
|
221 |
||
222 |
val _ = List.app OuterKeyword.keyword [infixK, infixlK, infixrK]; |
|
223 |
||
28064 | 224 |
|
225 |
(** pretty literals **) |
|
226 |
||
227 |
datatype literals = Literals of { |
|
228 |
literal_char: string -> string, |
|
229 |
literal_string: string -> string, |
|
230 |
literal_numeral: bool -> int -> string, |
|
231 |
literal_list: Pretty.T list -> Pretty.T, |
|
232 |
infix_cons: int * string |
|
233 |
}; |
|
234 |
||
235 |
fun dest_Literals (Literals lits) = lits; |
|
236 |
||
237 |
val literal_char = #literal_char o dest_Literals; |
|
238 |
val literal_string = #literal_string o dest_Literals; |
|
239 |
val literal_numeral = #literal_numeral o dest_Literals; |
|
240 |
val literal_list = #literal_list o dest_Literals; |
|
241 |
val infix_cons = #infix_cons o dest_Literals; |
|
242 |
||
28060 | 243 |
end; (*struct*) |