| author | haftmann | 
| Fri, 19 Jun 2009 19:45:00 +0200 | |
| changeset 31725 | f08507464b9d | 
| parent 29606 | fedb8be05f24 | 
| child 33035 | 15eab423e573 | 
| permissions | -rw-r--r-- | 
| 24584 | 1  | 
(* Title: Pure/ProofGeneral/pgml.ML  | 
| 23571 | 2  | 
Author: David Aspinall  | 
3  | 
||
| 28038 | 4  | 
PGIP abstraction: PGML  | 
| 23571 | 5  | 
*)  | 
6  | 
||
7  | 
signature PGML =  | 
|
8  | 
sig  | 
|
9  | 
    type pgmlsym = { name: string, content: string }
 | 
|
10  | 
||
11  | 
datatype pgmlatom = Sym of pgmlsym | Str of string  | 
|
12  | 
||
13  | 
datatype pgmlorient = HOVOrient | HOrient | VOrient | HVOrient  | 
|
14  | 
||
| 28038 | 15  | 
datatype pgmlplace = Subscript | Superscript | Above | Below  | 
| 23571 | 16  | 
|
17  | 
datatype pgmldec = Bold | Italic | Error | Warning | Info | Other of string  | 
|
18  | 
||
19  | 
datatype pgmlaction = Toggle | Button | Menu  | 
|
| 28038 | 20  | 
|
21  | 
datatype pgmlterm =  | 
|
| 28020 | 22  | 
             Atoms of { kind: string option, content: pgmlatom list }
 | 
23  | 
           | Box of { orient: pgmlorient option, indent: int option, content: pgmlterm list }
 | 
|
24  | 
           | Break of { mandatory: bool option, indent: int option }
 | 
|
25  | 
           | Subterm of { kind: string option,
 | 
|
26  | 
param: string option,  | 
|
27  | 
place: pgmlplace option,  | 
|
28  | 
name: string option,  | 
|
29  | 
decoration: pgmldec option,  | 
|
30  | 
action: pgmlaction option,  | 
|
31  | 
pos: string option,  | 
|
32  | 
xref: PgipTypes.pgipurl option,  | 
|
33  | 
content: pgmlterm list }  | 
|
34  | 
           | Alt of { kind: string option, content: pgmlterm list }
 | 
|
35  | 
| Embed of XML.tree list  | 
|
36  | 
| Raw of XML.tree  | 
|
| 23571 | 37  | 
|
| 28038 | 38  | 
datatype pgml =  | 
39  | 
             Pgml of { version: string option, systemid: string option,
 | 
|
40  | 
area: PgipTypes.displayarea option,  | 
|
| 28020 | 41  | 
content: pgmlterm list }  | 
| 28038 | 42  | 
|
43  | 
val atom_to_xml : pgmlatom -> XML.tree  | 
|
| 23571 | 44  | 
val pgmlterm_to_xml : pgmlterm -> XML.tree  | 
45  | 
||
| 
23748
 
1ff6b562076f
Track schema changes: add area attribute to pgml packet.  Also add quoted Raw element [hack for Isabelle bottom-up XML production]
 
aspinall 
parents: 
23723 
diff
changeset
 | 
46  | 
val pgml_to_xml : pgml -> XML.tree  | 
| 23571 | 47  | 
end  | 
48  | 
||
49  | 
||
50  | 
structure Pgml : PGML =  | 
|
51  | 
struct  | 
|
52  | 
open PgipTypes  | 
|
53  | 
||
54  | 
    type pgmlsym = { name: string, content: string }
 | 
|
| 
23748
 
1ff6b562076f
Track schema changes: add area attribute to pgml packet.  Also add quoted Raw element [hack for Isabelle bottom-up XML production]
 
aspinall 
parents: 
23723 
diff
changeset
 | 
55  | 
|
| 23571 | 56  | 
datatype pgmlatom = Sym of pgmlsym | Str of string  | 
57  | 
||
58  | 
datatype pgmlorient = HOVOrient | HOrient | VOrient | HVOrient  | 
|
59  | 
||
| 28038 | 60  | 
datatype pgmlplace = Subscript | Superscript | Above | Below  | 
| 23571 | 61  | 
|
62  | 
datatype pgmldec = Bold | Italic | Error | Warning | Info | Other of string  | 
|
63  | 
||
64  | 
datatype pgmlaction = Toggle | Button | Menu  | 
|
| 28038 | 65  | 
|
66  | 
datatype pgmlterm =  | 
|
| 28020 | 67  | 
             Atoms of { kind: string option, content: pgmlatom list }
 | 
68  | 
           | Box of { orient: pgmlorient option, indent: int option, content: pgmlterm list }
 | 
|
69  | 
           | Break of { mandatory: bool option, indent: int option }
 | 
|
70  | 
           | Subterm of { kind: string option,
 | 
|
71  | 
param: string option,  | 
|
72  | 
place: pgmlplace option,  | 
|
73  | 
name: string option,  | 
|
74  | 
decoration: pgmldec option,  | 
|
75  | 
action: pgmlaction option,  | 
|
76  | 
pos: string option,  | 
|
77  | 
xref: PgipTypes.pgipurl option,  | 
|
78  | 
content: pgmlterm list }  | 
|
79  | 
           | Alt of { kind: string option, content: pgmlterm list }
 | 
|
80  | 
| Embed of XML.tree list  | 
|
81  | 
| Raw of XML.tree  | 
|
| 23571 | 82  | 
|
| 28038 | 83  | 
|
84  | 
datatype pgml =  | 
|
85  | 
             Pgml of { version: string option, systemid: string option,
 | 
|
86  | 
area: PgipTypes.displayarea option,  | 
|
| 28020 | 87  | 
content: pgmlterm list }  | 
| 23571 | 88  | 
|
89  | 
fun pgmlorient_to_string HOVOrient = "hov"  | 
|
90  | 
| pgmlorient_to_string HOrient = "h"  | 
|
91  | 
| pgmlorient_to_string VOrient = "v"  | 
|
92  | 
| pgmlorient_to_string HVOrient = "hv"  | 
|
93  | 
||
94  | 
fun pgmlplace_to_string Subscript = "sub"  | 
|
95  | 
| pgmlplace_to_string Superscript = "sup"  | 
|
96  | 
| pgmlplace_to_string Above = "above"  | 
|
97  | 
| pgmlplace_to_string Below = "below"  | 
|
98  | 
||
99  | 
fun pgmldec_to_string Bold = "bold"  | 
|
100  | 
| pgmldec_to_string Italic = "italic"  | 
|
101  | 
| pgmldec_to_string Error = "error"  | 
|
102  | 
| pgmldec_to_string Warning = "warning"  | 
|
103  | 
| pgmldec_to_string Info = "info"  | 
|
104  | 
| pgmldec_to_string (Other s) = "other"  | 
|
105  | 
||
106  | 
fun pgmlaction_to_string Toggle = "toggle"  | 
|
107  | 
| pgmlaction_to_string Button = "button"  | 
|
108  | 
| pgmlaction_to_string Menu = "menu"  | 
|
109  | 
||
| 28038 | 110  | 
(* NOTE: we assume strings are already XML escaped here, for convenience in Isabelle;  | 
111  | 
would be better not to *) (* FIXME !??? *)  | 
|
| 28039 | 112  | 
    fun atom_to_xml (Sym {name, content}) = XML.Elem ("sym", attr "name" name, [XML.Text content])
 | 
| 28038 | 113  | 
| atom_to_xml (Str content) = XML.Text content;  | 
114  | 
||
115  | 
    fun pgmlterm_to_xml (Atoms {kind, content}) =
 | 
|
116  | 
        XML.Elem("atom", opt_attr "kind" kind, map atom_to_xml content)
 | 
|
| 23571 | 117  | 
|
118  | 
      | pgmlterm_to_xml (Box {orient, indent, content}) =
 | 
|
| 28038 | 119  | 
        XML.Elem("box",
 | 
| 28020 | 120  | 
opt_attr_map pgmlorient_to_string "orient" orient @  | 
121  | 
opt_attr_map int_to_pgstring "indent" indent,  | 
|
122  | 
map pgmlterm_to_xml content)  | 
|
| 23571 | 123  | 
|
124  | 
      | pgmlterm_to_xml (Break {mandatory, indent}) =
 | 
|
| 28020 | 125  | 
        XML.Elem("break",
 | 
126  | 
opt_attr_map bool_to_pgstring "mandatory" mandatory @  | 
|
127  | 
opt_attr_map int_to_pgstring "indent" indent, [])  | 
|
| 23571 | 128  | 
|
129  | 
      | pgmlterm_to_xml (Subterm {kind, param, place, name, decoration, action, pos, xref, content}) =
 | 
|
| 28038 | 130  | 
        XML.Elem("subterm",
 | 
| 28020 | 131  | 
opt_attr "kind" kind @  | 
132  | 
opt_attr "param" param @  | 
|
133  | 
opt_attr_map pgmlplace_to_string "place" place @  | 
|
| 28038 | 134  | 
opt_attr "name" name @  | 
| 28020 | 135  | 
opt_attr_map pgmldec_to_string "decoration" decoration @  | 
| 28038 | 136  | 
opt_attr_map pgmlaction_to_string "action" action @  | 
| 28020 | 137  | 
opt_attr "pos" pos @  | 
138  | 
opt_attr_map string_of_pgipurl "xref" xref,  | 
|
139  | 
map pgmlterm_to_xml content)  | 
|
| 23571 | 140  | 
|
| 28038 | 141  | 
      | pgmlterm_to_xml (Alt {kind, content}) =
 | 
| 28020 | 142  | 
        XML.Elem("alt", opt_attr "kind" kind, map pgmlterm_to_xml content)
 | 
| 23571 | 143  | 
|
144  | 
      | pgmlterm_to_xml (Embed xmls) = XML.Elem("embed", [], xmls)
 | 
|
145  | 
||
| 
23748
 
1ff6b562076f
Track schema changes: add area attribute to pgml packet.  Also add quoted Raw element [hack for Isabelle bottom-up XML production]
 
aspinall 
parents: 
23723 
diff
changeset
 | 
146  | 
| pgmlterm_to_xml (Raw xml) = xml  | 
| 23571 | 147  | 
|
| 
23748
 
1ff6b562076f
Track schema changes: add area attribute to pgml packet.  Also add quoted Raw element [hack for Isabelle bottom-up XML production]
 
aspinall 
parents: 
23723 
diff
changeset
 | 
148  | 
|
| 28038 | 149  | 
datatype pgml =  | 
150  | 
             Pgml of { version: string option, systemid: string option,
 | 
|
151  | 
area: PgipTypes.displayarea option,  | 
|
| 28020 | 152  | 
content: pgmlterm list }  | 
| 
23748
 
1ff6b562076f
Track schema changes: add area attribute to pgml packet.  Also add quoted Raw element [hack for Isabelle bottom-up XML production]
 
aspinall 
parents: 
23723 
diff
changeset
 | 
153  | 
|
| 28038 | 154  | 
    fun pgml_to_xml (Pgml {version,systemid,area,content}) =
 | 
| 28020 | 155  | 
        XML.Elem("pgml",
 | 
| 28038 | 156  | 
opt_attr "version" version @  | 
| 28020 | 157  | 
opt_attr "systemid" systemid @  | 
158  | 
Option.getOpt(Option.map PgipTypes.attrs_of_displayarea area,[]),  | 
|
159  | 
map pgmlterm_to_xml content)  | 
|
| 23571 | 160  | 
end  |