obua@19064: signature XML_CONV = obua@19064: sig obua@19064: type 'a input = XML.tree -> 'a obua@19064: type 'a output = 'a -> XML.tree obua@19064: obua@19064: exception ParseException of string obua@19064: obua@19064: val xml_of_typ: typ output obua@19064: val typ_of_xml: typ input obua@19064: obua@19064: val xml_of_term: term output obua@19064: val term_of_xml : term input obua@19064: obua@19064: val xml_of_mixfix: mixfix output obua@19064: val mixfix_of_xml: mixfix input obua@19064: obua@19064: val xml_of_proof: Proofterm.proof output obua@19064: obua@19064: val xml_of_bool: bool output obua@19064: val bool_of_xml: bool input wenzelm@32960: obua@19064: val xml_of_int: int output obua@19064: val int_of_xml: int input obua@19064: obua@19064: val xml_of_string: string output obua@19064: val string_of_xml: string input obua@19064: obua@19064: val xml_of_list: 'a output -> 'a list output obua@19064: val list_of_xml: 'a input -> 'a list input obua@19064: obua@19064: val xml_of_tuple : 'a output -> 'a list output obua@19064: val tuple_of_xml: 'a input -> int -> 'a list input obua@19064: obua@19064: val xml_of_option: 'a output -> 'a option output obua@19064: val option_of_xml: 'a input -> 'a option input obua@19064: obua@19064: val xml_of_pair: 'a output -> 'b output -> ('a * 'b) output obua@19064: val pair_of_xml: 'a input -> 'b input -> ('a * 'b) input obua@19064: obua@19064: val xml_of_triple: 'a output -> 'b output -> 'c output -> ('a * 'b * 'c) output obua@19064: val triple_of_xml: 'a input -> 'b input -> 'c input -> ('a * 'b * 'c) input obua@19064: obua@19064: val xml_of_unit: unit output obua@19064: val unit_of_xml: unit input obua@19064: obua@19064: val xml_of_quadruple: 'a output -> 'b output -> 'c output -> 'd output -> ('a * 'b * 'c * 'd) output obua@19064: val quadruple_of_xml: 'a input -> 'b input -> 'c input -> 'd input -> ('a * 'b * 'c * 'd) input obua@19064: obua@19064: val xml_of_quintuple: 'a output -> 'b output -> 'c output -> 'd output -> 'e output -> ('a * 'b * 'c * 'd * 'e) output obua@19064: val quintuple_of_xml: 'a input -> 'b input -> 'c input -> 'd input -> 'e input -> ('a * 'b * 'c * 'd * 'e) input obua@19064: obua@19064: val wrap : string -> 'a output -> 'a output obua@19064: val unwrap : ('a -> 'b) -> 'a input -> 'b input obua@19064: val wrap_empty : string output obua@19064: val unwrap_empty : 'a -> 'a input obua@19064: val name_of_wrap : XML.tree -> string obua@19064: obua@19064: val write_to_file: 'a output -> string -> 'a -> unit obua@19064: val read_from_file: 'a input -> string -> 'a obua@19089: obua@19089: val serialize_to_file : 'a output -> string -> 'a -> unit obua@19089: val deserialize_from_file : 'a input -> string -> 'a obua@19064: end; obua@19064: obua@19064: structure XMLConv : XML_CONV = obua@19064: struct obua@19064: obua@19064: type 'a input = XML.tree -> 'a obua@19064: type 'a output = 'a -> XML.tree obua@19064: exception ParseException of string obua@19064: obua@19064: fun parse_err s = raise ParseException s obua@19064: obua@19064: fun xml_of_class name = XML.Elem ("class", [("name", name)], []); obua@19064: obua@19064: fun class_of_xml (XML.Elem ("class", [("name", name)], [])) = name obua@19064: | class_of_xml _ = parse_err "class" obua@19064: obua@19064: fun xml_of_typ (TVar ((s, i), S)) = XML.Elem ("TVar", obua@19064: ("name", s) :: (if i=0 then [] else [("index", string_of_int i)]), obua@19064: map xml_of_class S) obua@19064: | xml_of_typ (TFree (s, S)) = obua@19064: XML.Elem ("TFree", [("name", s)], map xml_of_class S) obua@19064: | xml_of_typ (Type (s, Ts)) = obua@19064: XML.Elem ("Type", [("name", s)], map xml_of_typ Ts); obua@19064: obua@19064: fun intofstr s i = obua@19064: case Int.fromString i of wenzelm@32960: NONE => parse_err (s^", invalid index: "^i) obua@19064: | SOME i => i obua@19064: obua@19064: fun typ_of_xml (XML.Elem ("TVar", [("name", s)], cs)) = TVar ((s,0), map class_of_xml cs) obua@19064: | typ_of_xml (XML.Elem ("TVar", [("name", s), ("index", i)], cs)) = obua@19064: TVar ((s, intofstr "TVar" i), map class_of_xml cs) obua@19064: | typ_of_xml (XML.Elem ("TFree", [("name", s)], cs)) = TFree (s, map class_of_xml cs) obua@19064: | typ_of_xml (XML.Elem ("Type", [("name", s)], ts)) = Type (s, map typ_of_xml ts) obua@19064: | typ_of_xml _ = parse_err "type" obua@19064: obua@19064: fun xml_of_term (Bound i) = obua@19064: XML.Elem ("Bound", [("index", string_of_int i)], []) obua@19064: | xml_of_term (Free (s, T)) = obua@19064: XML.Elem ("Free", [("name", s)], [xml_of_typ T]) obua@19064: | xml_of_term (Var ((s, i), T)) = XML.Elem ("Var", obua@19064: ("name", s) :: (if i=0 then [] else [("index", string_of_int i)]), obua@19064: [xml_of_typ T]) obua@19064: | xml_of_term (Const (s, T)) = obua@19064: XML.Elem ("Const", [("name", s)], [xml_of_typ T]) obua@19064: | xml_of_term (t $ u) = obua@19064: XML.Elem ("App", [], [xml_of_term t, xml_of_term u]) obua@19064: | xml_of_term (Abs (s, T, t)) = obua@19064: XML.Elem ("Abs", [("vname", s)], [xml_of_typ T, xml_of_term t]); obua@19064: obua@19064: fun term_of_xml (XML.Elem ("Bound", [("index", i)], [])) = Bound (intofstr "Bound" i) obua@19064: | term_of_xml (XML.Elem ("Free", [("name", s)], [T])) = Free (s, typ_of_xml T) obua@19064: | term_of_xml (XML.Elem ("Var",[("name", s)], [T])) = Var ((s,0), typ_of_xml T) obua@19064: | term_of_xml (XML.Elem ("Var",[("name", s), ("index", i)], [T])) = Var ((s,intofstr "Var" i), typ_of_xml T) obua@19064: | term_of_xml (XML.Elem ("Const", [("name", s)], [T])) = Const (s, typ_of_xml T) obua@19064: | term_of_xml (XML.Elem ("App", [], [t, u])) = (term_of_xml t) $ (term_of_xml u) obua@19064: | term_of_xml (XML.Elem ("Abs", [("vname", s)], [T, t])) = Abs (s, typ_of_xml T, term_of_xml t) obua@19064: | term_of_xml _ = parse_err "term" obua@19064: obua@19064: fun xml_of_opttypes NONE = [] obua@19064: | xml_of_opttypes (SOME Ts) = [XML.Elem ("types", [], map xml_of_typ Ts)]; obua@19064: obua@19064: (* FIXME: the t argument of PThm and PAxm is actually redundant, since *) obua@19064: (* it can be looked up in the theorem database. Thus, it could be *) obua@19064: (* omitted from the xml representation. *) obua@19064: obua@19064: fun xml_of_proof (PBound i) = obua@19064: XML.Elem ("PBound", [("index", string_of_int i)], []) obua@19064: | xml_of_proof (Abst (s, optT, prf)) = obua@19064: XML.Elem ("Abst", [("vname", s)], obua@19064: (case optT of NONE => [] | SOME T => [xml_of_typ T]) @ obua@19064: [xml_of_proof prf]) obua@19064: | xml_of_proof (AbsP (s, optt, prf)) = obua@19064: XML.Elem ("AbsP", [("vname", s)], obua@19064: (case optt of NONE => [] | SOME t => [xml_of_term t]) @ obua@19064: [xml_of_proof prf]) obua@19064: | xml_of_proof (prf % optt) = obua@19064: XML.Elem ("Appt", [], obua@19064: xml_of_proof prf :: obua@19064: (case optt of NONE => [] | SOME t => [xml_of_term t])) obua@19064: | xml_of_proof (prf %% prf') = obua@19064: XML.Elem ("AppP", [], [xml_of_proof prf, xml_of_proof prf']) obua@19064: | xml_of_proof (Hyp t) = XML.Elem ("Hyp", [], [xml_of_term t]) wenzelm@28811: | xml_of_proof (PThm (_, ((s, t, optTs), _))) = obua@19064: XML.Elem ("PThm", [("name", s)], xml_of_term t :: xml_of_opttypes optTs) obua@19064: | xml_of_proof (PAxm (s, t, optTs)) = obua@19064: XML.Elem ("PAxm", [("name", s)], xml_of_term t :: xml_of_opttypes optTs) obua@19064: | xml_of_proof (Oracle (s, t, optTs)) = obua@19064: XML.Elem ("Oracle", [("name", s)], xml_of_term t :: xml_of_opttypes optTs) wenzelm@28811: | xml_of_proof MinProof = XML.Elem ("MinProof", [], []); obua@19064: obua@19064: fun xml_of_bool b = XML.Elem ("bool", [("value", if b then "true" else "false")], []) obua@19064: fun xml_of_int i = XML.Elem ("int", [("value", string_of_int i)], []) obua@19064: fun xml_of_string s = XML.Elem ("string", [("value", s)], []) obua@19064: fun xml_of_unit () = XML.Elem ("unit", [], []) obua@19064: fun xml_of_list output ls = XML.Elem ("list", [], map output ls) obua@19064: fun xml_of_tuple output ls = XML.Elem ("tuple", [], map output ls) obua@19064: fun xml_of_option output opt = XML.Elem ("option", [], case opt of NONE => [] | SOME x => [output x]) obua@19064: fun wrap s output x = XML.Elem (s, [], [output x]) obua@19064: fun wrap_empty s = XML.Elem (s, [], []) obua@19064: fun xml_of_pair output1 output2 (x1, x2) = XML.Elem ("pair", [], [output1 x1, output2 x2]) obua@19064: fun xml_of_triple output1 output2 output3 (x1, x2, x3) = XML.Elem ("triple", [], [output1 x1, output2 x2, output3 x3]) obua@19064: fun xml_of_quadruple output1 output2 output3 output4 (x1, x2, x3, x4) = obua@19064: XML.Elem ("quadruple", [], [output1 x1, output2 x2, output3 x3, output4 x4]) obua@19064: fun xml_of_quintuple output1 output2 output3 output4 output5 (x1, x2, x3, x4, x5) = obua@19064: XML.Elem ("quintuple", [], [output1 x1, output2 x2, output3 x3, output4 x4, output5 x5]) wenzelm@32960: obua@19064: fun bool_of_xml (XML.Elem ("bool", [("value", "true")], [])) = true obua@19064: | bool_of_xml (XML.Elem ("bool", [("value", "false")], [])) = false obua@19064: | bool_of_xml _ = parse_err "bool" obua@19064: obua@19064: fun int_of_xml (XML.Elem ("int", [("value", i)], [])) = intofstr "int" i obua@19064: | int_of_xml _ = parse_err "int" obua@19064: obua@19064: fun unit_of_xml (XML.Elem ("unit", [], [])) = () obua@19064: | unit_of_xml _ = parse_err "unit" obua@19064: obua@19064: fun list_of_xml input (XML.Elem ("list", [], ls)) = map input ls obua@19064: | list_of_xml _ _ = parse_err "list" obua@19064: obua@19064: fun tuple_of_xml input i (XML.Elem ("tuple", [], ls)) = obua@19064: if i = length ls then map input ls else parse_err "tuple" obua@19064: | tuple_of_xml _ _ _ = parse_err "tuple" obua@19064: obua@19064: fun option_of_xml input (XML.Elem ("option", [], [])) = NONE obua@19064: | option_of_xml input (XML.Elem ("option", [], [opt])) = SOME (input opt) obua@19064: | option_of_xml _ _ = parse_err "option" obua@19064: obua@19064: fun pair_of_xml input1 input2 (XML.Elem ("pair", [], [x1, x2])) = (input1 x1, input2 x2) obua@19064: | pair_of_xml _ _ _ = parse_err "pair" obua@19064: obua@19064: fun triple_of_xml input1 input2 input3 (XML.Elem ("triple", [], [x1, x2, x3])) = (input1 x1, input2 x2, input3 x3) obua@19064: | triple_of_xml _ _ _ _ = parse_err "triple" obua@19064: obua@19064: fun quadruple_of_xml input1 input2 input3 input4 (XML.Elem ("quadruple", [], [x1, x2, x3, x4])) = obua@19064: (input1 x1, input2 x2, input3 x3, input4 x4) obua@19064: | quadruple_of_xml _ _ _ _ _ = parse_err "quadruple" obua@19064: obua@19064: fun quintuple_of_xml input1 input2 input3 input4 input5 (XML.Elem ("quintuple", [], [x1, x2, x3, x4, x5])) = obua@19064: (input1 x1, input2 x2, input3 x3, input4 x4, input5 x5) obua@19064: | quintuple_of_xml _ _ _ _ _ _ = parse_err "quintuple" obua@19064: obua@19064: fun unwrap f input (XML.Elem (_, [], [x])) = f (input x) obua@19064: | unwrap _ _ _ = parse_err "unwrap" obua@19064: obua@19064: fun unwrap_empty x (XML.Elem (_, [], [])) = x obua@19064: | unwrap_empty _ _ = parse_err "unwrap_unit" obua@19064: obua@19064: fun name_of_wrap (XML.Elem (name, _, _)) = name obua@19064: | name_of_wrap _ = parse_err "name_of_wrap" obua@19064: obua@19064: fun string_of_xml (XML.Elem ("string", [("value", s)], [])) = s obua@19064: | string_of_xml _ = parse_err "string" obua@19064: obua@19064: fun xml_of_mixfix NoSyn = wrap_empty "nosyn" obua@19064: | xml_of_mixfix Structure = wrap_empty "structure" obua@19064: | xml_of_mixfix (Mixfix args) = wrap "mixfix" (xml_of_triple xml_of_string (xml_of_list xml_of_int) xml_of_int) args obua@19064: | xml_of_mixfix (Delimfix s) = wrap "delimfix" xml_of_string s wenzelm@35130: | xml_of_mixfix (Infix args) = wrap "infix" (xml_of_pair xml_of_string xml_of_int) args wenzelm@35130: | xml_of_mixfix (Infixl args) = wrap "infixl" (xml_of_pair xml_of_string xml_of_int) args wenzelm@35130: | xml_of_mixfix (Infixr args) = wrap "infixr" (xml_of_pair xml_of_string xml_of_int) args obua@19064: | xml_of_mixfix (Binder args) = wrap "binder" (xml_of_triple xml_of_string xml_of_int xml_of_int) args wenzelm@32960: obua@19064: fun mixfix_of_xml e = obua@19064: (case name_of_wrap e of wenzelm@32960: "nosyn" => unwrap_empty NoSyn obua@19064: | "structure" => unwrap_empty Structure obua@19064: | "mixfix" => unwrap Mixfix (triple_of_xml string_of_xml (list_of_xml int_of_xml) int_of_xml) obua@19064: | "delimfix" => unwrap Delimfix string_of_xml wenzelm@35130: | "infix" => unwrap Infix (pair_of_xml string_of_xml int_of_xml) wenzelm@35130: | "infixl" => unwrap Infixl (pair_of_xml string_of_xml int_of_xml) wenzelm@35130: | "infixr" => unwrap Infixr (pair_of_xml string_of_xml int_of_xml) obua@19064: | "binder" => unwrap Binder (triple_of_xml string_of_xml int_of_xml int_of_xml) obua@19064: | _ => parse_err "mixfix" obua@19064: ) e obua@19064: obua@19064: wenzelm@21858: fun to_file f output fname x = File.write (Path.explode fname) (f (output x)) obua@19064: obua@19089: fun from_file f input fname = obua@19064: let wenzelm@32960: val _ = writeln "read_from_file enter" wenzelm@32960: val s = File.read (Path.explode fname) wenzelm@32960: val _ = writeln "done: read file" wenzelm@32960: val tree = timeit (fn () => f s) wenzelm@32960: val _ = writeln "done: tree" wenzelm@32960: val x = timeit (fn () => input tree) wenzelm@32960: val _ = writeln "done: input" obua@19064: in wenzelm@32960: x obua@19064: end obua@19064: obua@19089: fun write_to_file x = to_file XML.string_of_tree x obua@19089: fun read_from_file x = timeit (fn () => (writeln "read_from_file"; from_file XML.tree_of_string x)) obua@19089: obua@19089: fun serialize_to_file x = to_file XML.encoded_string_of_tree x obua@19089: fun deserialize_from_file x = timeit (fn () => (writeln "deserialize_from_file"; from_file XML.tree_of_encoded_string x)) obua@19089: obua@19064: end; obua@19064: obua@19064: structure XMLConvOutput = obua@19064: struct obua@19064: open XMLConv obua@19064: obua@19064: val typ = xml_of_typ obua@19064: val term = xml_of_term obua@19064: val mixfix = xml_of_mixfix obua@19064: val proof = xml_of_proof obua@19064: val bool = xml_of_bool obua@19064: val int = xml_of_int obua@19064: val string = xml_of_string obua@19064: val unit = xml_of_unit obua@19064: val list = xml_of_list obua@19064: val pair = xml_of_pair obua@19064: val option = xml_of_option obua@19064: val triple = xml_of_triple obua@19064: val quadruple = xml_of_quadruple obua@19064: val quintuple = xml_of_quintuple obua@19064: obua@19064: end obua@19064: obua@19064: structure XMLConvInput = obua@19064: struct obua@19064: open XMLConv obua@19064: obua@19064: val typ = typ_of_xml obua@19064: val term = term_of_xml obua@19064: val mixfix = mixfix_of_xml obua@19064: val bool = bool_of_xml obua@19064: val int = int_of_xml obua@19064: val string = string_of_xml obua@19064: val unit = unit_of_xml obua@19064: val list = list_of_xml obua@19064: val pair = pair_of_xml obua@19064: val option = option_of_xml obua@19064: val triple = triple_of_xml obua@19064: val quadruple = quadruple_of_xml obua@19064: val quintuple = quintuple_of_xml obua@19064: obua@19064: end obua@19064: