src/Pure/Thy/export_theory.ML
author wenzelm
Sun, 16 Sep 2018 22:45:34 +0200
changeset 69003 a015f1d3ba0c
parent 68997 4278947ba336
child 69019 a6ba77af6486
permissions -rw-r--r--
export plain infix syntax;

(*  Title:      Pure/Thy/export_theory.ML
    Author:     Makarius

Export foundational theory content.
*)

signature EXPORT_THEORY =
sig
  val setup_presentation: (Thy_Info.presentation_context -> theory -> unit) -> unit
  val export_body: theory -> string -> XML.body -> unit
end;

structure Export_Theory: EXPORT_THEORY =
struct

(* names for bound variables *)

local
  fun declare_names (Abs (_, _, b)) = declare_names b
    | declare_names (t $ u) = declare_names t #> declare_names u
    | declare_names (Const (c, _)) = Name.declare (Long_Name.base_name c)
    | declare_names (Free (x, _)) = Name.declare x
    | declare_names _ = I;

  fun variant_abs bs (Abs (x, T, t)) =
        let
          val names = fold Name.declare bs (declare_names t Name.context);
          val x' = #1 (Name.variant x names);
          val t' = variant_abs (x' :: bs) t;
        in Abs (x', T, t') end
    | variant_abs bs (t $ u) = variant_abs bs t $ variant_abs bs u
    | variant_abs _ t = t;
in
  val named_bounds = variant_abs [];
end;


(* general setup *)

fun setup_presentation f =
  Theory.setup (Thy_Info.add_presentation (fn context => fn thy =>
    if Options.bool (#options context) "export_theory" then f context thy else ()));

fun export_body thy name body =
  Export.export thy ("theory/" ^ name) (Buffer.chunks (YXML.buffer_body body Buffer.empty));


(* presentation *)

val _ = setup_presentation (fn {adjust_pos, ...} => fn thy =>
  let
    val parents = Theory.parents_of thy;
    val rep_tsig = Type.rep_tsig (Sign.tsig_of thy);

    val thy_ctxt = Proof_Context.init_global thy;


    (* entities *)

    fun entity_markup space name =
      let
        val xname = Name_Space.extern_shortest thy_ctxt space name;
        val {serial, pos, ...} = Name_Space.the_entry space name;
        val props =
          Position.offset_properties_of (adjust_pos pos) @
          Position.id_properties_of pos @
          Markup.serial_properties serial;
      in (Markup.entityN, (Markup.nameN, name) :: (Markup.xnameN, xname) :: props) end;

    fun export_entities export_name export get_space decls =
      let val elems =
        let
          val parent_spaces = map get_space parents;
          val space = get_space thy;
        in
          (decls, []) |-> fold (fn (name, decl) =>
            if exists (fn space => Name_Space.declared space name) parent_spaces then I
            else
              (case export name decl of
                NONE => I
              | SOME body =>
                  cons (#serial (Name_Space.the_entry space name),
                    XML.Elem (entity_markup space name, body))))
          |> sort (int_ord o apply2 #1) |> map #2
        end;
      in if null elems then () else export_body thy export_name elems end;


    (* infix syntax *)

    fun get_infix_const ctxt = Syntax.get_infix (Proof_Context.syn_of ctxt) o Lexicon.mark_const;
    fun get_infix_type ctxt = Syntax.get_infix (Proof_Context.syn_of ctxt) o Lexicon.mark_type;

    fun encode_infix {assoc, delim, pri} =
      let
        val ass =
          (case assoc of
            Syntax_Ext.No_Assoc => 0
          | Syntax_Ext.Left_Assoc => 1
          | Syntax_Ext.Right_Assoc => 2);
        open XML.Encode Term_XML.Encode;
      in triple int string int (ass, delim, pri) end;


    (* types *)

    val encode_type =
      let open XML.Encode Term_XML.Encode
      in triple (option encode_infix) (list string) (option typ) end;

    fun export_type c (Type.LogicalType n) =
          SOME (encode_type (get_infix_type thy_ctxt c, Name.invent Name.context Name.aT n, NONE))
      | export_type c (Type.Abbreviation (args, U, false)) =
          SOME (encode_type (get_infix_type thy_ctxt c, args, SOME U))
      | export_type _ _ = NONE;

    val _ =
      export_entities "types" export_type Sign.type_space
        (Name_Space.dest_table (#types rep_tsig));


    (* consts *)

    val encode_const =
      let open XML.Encode Term_XML.Encode in
        pair (option encode_infix) (pair (list string) (pair typ (option (term o named_bounds))))
      end;

    fun export_const c (T, abbrev) =
      let
        val syntax = get_infix_const thy_ctxt c;
        val T' = T |> Logic.unvarifyT_global |> Type.strip_sorts;
        val abbrev' = abbrev |> Option.map (Logic.unvarify_global #> map_types Type.strip_sorts);
        val args = map (#1 o dest_TFree) (Consts.typargs (Sign.consts_of thy) (c, T'));
      in SOME (encode_const (syntax, (args, (T', abbrev')))) end;

    val _ =
      export_entities "consts" export_const Sign.const_space
        (#constants (Consts.dest (Sign.consts_of thy)));


    (* axioms and facts *)

    fun standard_prop_of raw_thm =
      let
        val thm = raw_thm
          |> Thm.transfer thy
          |> Thm.check_hyps (Context.Theory thy)
          |> Thm.strip_shyps;
        val prop = thm
          |> Thm.full_prop_of
          |> Term_Subst.zero_var_indexes;
      in (Thm.extra_shyps thm, prop) end;

    fun encode_prop (Ss, prop) =
      let
        val prop' = Logic.unvarify_global (named_bounds prop);
        val typargs = rev (Term.add_tfrees prop' []);
        val sorts = Name.invent (Name.make_context (map #1 typargs)) Name.aT (length Ss) ~~ Ss;
        val args = rev (Term.add_frees prop' []);
      in
        (sorts @ typargs, args, prop') |>
          let open XML.Encode Term_XML.Encode
          in triple (list (pair string sort)) (list (pair string typ)) term end
      end;

    val encode_fact = XML.Encode.list encode_prop o map standard_prop_of;

    val _ =
      export_entities "axioms" (fn _ => fn t => SOME (encode_prop ([], t))) Theory.axiom_space
        (Theory.axioms_of thy);
    val _ =
      export_entities "facts" (K (SOME o encode_fact)) (Facts.space_of o Global_Theory.facts_of)
        (Facts.dest_static true [] (Global_Theory.facts_of thy));


    (* type classes *)

    val encode_class =
      let open XML.Encode Term_XML.Encode
      in pair (list (pair string typ)) (list (term o named_bounds)) end;

    fun export_class name =
      (case try (Axclass.get_info thy) name of
        NONE => ([], [])
      | SOME {params, axioms, ...} =>
          (params, map (Logic.unvarify_global o Thm.full_prop_of) axioms))
      |> encode_class |> SOME;

    val _ =
      export_entities "classes" (fn name => fn () => export_class name)
        Sign.class_space (map (rpair ()) (Graph.keys (Sorts.classes_of (#2 (#classes rep_tsig)))));


    (* sort algebra *)

    val {classrel, arities} =
      Sorts.dest_algebra (map (#2 o #classes o Type.rep_tsig o Sign.tsig_of) parents)
        (#2 (#classes rep_tsig));

    val encode_classrel =
      let open XML.Encode
      in list (pair string (list string)) end;

    val encode_arities =
      let open XML.Encode Term_XML.Encode
      in list (triple string (list sort) string) end;

    val _ = if null classrel then () else export_body thy "classrel" (encode_classrel classrel);
    val _ = if null arities then () else export_body thy "arities" (encode_arities arities);


    (* locales *)

    fun encode_locale ({type_params, params, asm, defs}: Locale.content) =
      let
        val args = map #1 params;
        val typargs = type_params @ rev (fold Term.add_tfrees (map Free args @ the_list asm) []);
        val encode =
          let open XML.Encode Term_XML.Encode in
            pair (list (pair string sort))
              (pair (list (pair string typ))
                (pair (option term) (list term)))
          end;
      in encode (typargs, (args, (asm, defs))) end;

    val _ =
      export_entities "locales" (fn _ => SOME o encode_locale) Locale.locale_space
        (Locale.dest_locales thy);


    (* parents *)

    val _ =
      export_body thy "parents"
        (XML.Encode.string (cat_lines (map Context.theory_long_name parents)));

  in () end);

end;