src/Pure/Thy/export_theory.ML
author wenzelm
Fri, 28 Sep 2018 21:16:24 +0200
changeset 69077 11529ae45786
parent 69076 90cce2f79e77
child 69083 6f8ae6ddc26b
permissions -rw-r--r--
more approximative prefix syntax, including binder;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
68154
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Thy/export_theory.ML
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
     3
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
     4
Export foundational theory content and locale/class structure.
68154
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
     5
*)
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
     6
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
     7
signature EXPORT_THEORY =
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
     8
sig
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
     9
  val setup_presentation: (Thy_Info.presentation_context -> theory -> unit) -> unit
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
    10
  val export_body: theory -> string -> XML.body -> unit
68154
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
    11
end;
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
    12
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
    13
structure Export_Theory: EXPORT_THEORY =
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
    14
struct
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
    15
69077
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    16
(* approximative syntax *)
69076
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
    17
69077
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    18
val get_syntax = Syntax.get_approx o Proof_Context.syn_of;
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    19
fun get_syntax_type ctxt = get_syntax ctxt o Lexicon.mark_type;
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    20
fun get_syntax_const ctxt = get_syntax ctxt o Lexicon.mark_const;
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    21
fun get_syntax_fixed ctxt = get_syntax ctxt o Lexicon.mark_fixed;
69076
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
    22
69077
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    23
fun get_syntax_param ctxt loc x =
69076
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
    24
  let val thy = Proof_Context.theory_of ctxt in
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
    25
    if Class.is_class thy loc then
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
    26
      (case AList.lookup (op =) (Class.these_params thy [loc]) x of
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
    27
        NONE => NONE
69077
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    28
      | SOME (_, (c, _)) => get_syntax_const ctxt c)
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    29
    else get_syntax_fixed ctxt x
69076
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
    30
  end;
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
    31
69077
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    32
val encode_syntax =
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    33
  XML.Encode.variant
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    34
   [fn NONE => ([], []),
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    35
    fn SOME (Syntax.Prefix delim) => ([delim], []),
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    36
    fn SOME (Syntax.Infix {assoc, delim, pri}) =>
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    37
      let
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    38
        val ass =
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    39
          (case assoc of
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    40
            Printer.No_Assoc => 0
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    41
          | Printer.Left_Assoc => 1
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    42
          | Printer.Right_Assoc => 2);
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    43
        open XML.Encode Term_XML.Encode;
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    44
      in ([], triple int string int (ass, delim, pri)) end];
69076
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
    45
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
    46
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    47
(* standardization of variables: only frees and named bounds *)
68725
367e60d9aa1b explicit names for bound variables;
wenzelm
parents: 68724
diff changeset
    48
367e60d9aa1b explicit names for bound variables;
wenzelm
parents: 68724
diff changeset
    49
local
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    50
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    51
fun declare_names (Abs (_, _, b)) = declare_names b
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    52
  | declare_names (t $ u) = declare_names t #> declare_names u
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    53
  | declare_names (Const (c, _)) = Name.declare (Long_Name.base_name c)
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    54
  | declare_names (Free (x, _)) = Name.declare x
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    55
  | declare_names _ = I;
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    56
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    57
fun variant_abs bs (Abs (x, T, t)) =
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    58
      let
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    59
        val names = fold Name.declare bs (declare_names t Name.context);
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    60
        val x' = #1 (Name.variant x names);
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    61
        val t' = variant_abs (x' :: bs) t;
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    62
      in Abs (x', T, t') end
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    63
  | variant_abs bs (t $ u) = variant_abs bs t $ variant_abs bs u
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    64
  | variant_abs _ t = t;
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    65
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    66
in
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    67
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    68
fun standard_vars used =
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    69
  let
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    70
    fun zero_var_indexes tm =
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    71
      Term_Subst.instantiate (Term_Subst.zero_var_indexes_inst used [tm]) tm;
68725
367e60d9aa1b explicit names for bound variables;
wenzelm
parents: 68724
diff changeset
    72
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    73
    fun unvarifyT ty = ty |> Term.map_atyps
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    74
      (fn TVar ((a, _), S) => TFree (a, S)
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    75
        | T as TFree (a, _) =>
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    76
            if Name.is_declared used a then T
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    77
            else raise TYPE (Logic.bad_fixed a, [ty], []));
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    78
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    79
    fun unvarify tm = tm |> Term.map_aterms
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    80
      (fn Var ((x, _), T) => Free (x, T)
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    81
        | t as Free (x, _) =>
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    82
            if Name.is_declared used x then t
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    83
            else raise TERM (Logic.bad_fixed x, [tm])
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    84
        | t => t);
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    85
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    86
  in zero_var_indexes #> map_types unvarifyT #> unvarify #> variant_abs [] end;
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    87
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    88
val standard_vars_global = standard_vars Name.context;
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    89
68725
367e60d9aa1b explicit names for bound variables;
wenzelm
parents: 68724
diff changeset
    90
end;
367e60d9aa1b explicit names for bound variables;
wenzelm
parents: 68724
diff changeset
    91
367e60d9aa1b explicit names for bound variables;
wenzelm
parents: 68724
diff changeset
    92
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    93
(* free variables: not declared in the context *)
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    94
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    95
val is_free = not oo Name.is_declared;
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    96
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    97
fun add_frees used =
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    98
  fold_aterms (fn Free (x, T) => is_free used x ? insert (op =) (x, T) | _ => I);
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
    99
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   100
fun add_tfrees used =
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   101
  (fold_types o fold_atyps) (fn TFree (a, S) => is_free used a ? insert (op =) (a, S) | _ => I);
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   102
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   103
69029
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   104
(* locale content *)
69019
a6ba77af6486 export semi-unfolded locale axioms;
wenzelm
parents: 69003
diff changeset
   105
69034
855c3c501b09 obsolete (see aec64b88e708);
wenzelm
parents: 69029
diff changeset
   106
fun locale_content thy loc =
69019
a6ba77af6486 export semi-unfolded locale axioms;
wenzelm
parents: 69003
diff changeset
   107
  let
69076
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
   108
    val loc_ctxt = Locale.init loc thy;
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
   109
    val args = Locale.params_of thy loc
69077
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
   110
      |> map (fn ((x, T), _) => ((x, T), get_syntax_param loc_ctxt loc x));
69029
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   111
    val axioms =
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   112
      let
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   113
        val (intro1, intro2) = Locale.intros_of thy loc;
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   114
        fun intros_tac ctxt = Method.try_intros_tac ctxt (the_list intro1 @ the_list intro2);
69076
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
   115
        val inst = Expression.Named (args |> map (fn ((x, T), _) => (x, Free (x, T))));
69029
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   116
        val res =
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   117
          Proof_Context.init_global thy
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   118
          |> Interpretation.interpretation ([(loc, (("", false), (inst, [])))], [])
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   119
          |> Proof.refine (Method.Basic (METHOD o intros_tac))
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   120
          |> Seq.filter_results
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   121
          |> try Seq.hd;
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   122
      in
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   123
        (case res of
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   124
          SOME st => Thm.prems_of (#goal (Proof.goal st))
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   125
        | NONE => raise Fail ("Cannot unfold locale " ^ quote loc))
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   126
      end;
69076
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
   127
    val typargs = rev (fold Term.add_tfrees (map (Free o #1) args @ axioms) []);
69029
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   128
  in {typargs = typargs, args = args, axioms = axioms} end;
69019
a6ba77af6486 export semi-unfolded locale axioms;
wenzelm
parents: 69003
diff changeset
   129
69069
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   130
fun locale_dependency_subst thy (dep: Locale.locale_dependency) =
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   131
  let
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   132
    val (type_params, params) = Locale.parameters_of thy (#source dep);
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   133
    (* FIXME proper type_params wrt. locale_content (!?!) *)
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   134
    val typargs = fold (Term.add_tfreesT o #2 o #1) params type_params;
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   135
    val substT =
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   136
      typargs |> map_filter (fn v =>
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   137
        let
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   138
          val T = TFree v;
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   139
          val T' = Morphism.typ (#morphism dep) T;
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   140
        in if T = T' then NONE else SOME (v, T') end);
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   141
    val subst =
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   142
      params |> map_filter (fn (v, _) =>
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   143
        let
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   144
          val t = Free v;
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   145
          val t' = Morphism.term (#morphism dep) t;
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   146
        in if t aconv t' then NONE else SOME (v, t') end);
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   147
  in (substT, subst) end;
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   148
69019
a6ba77af6486 export semi-unfolded locale axioms;
wenzelm
parents: 69003
diff changeset
   149
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   150
(* general setup *)
68154
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
   151
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   152
fun setup_presentation f =
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   153
  Theory.setup (Thy_Info.add_presentation (fn context => fn thy =>
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   154
    if Options.bool (#options context) "export_theory" then f context thy else ()));
68154
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
   155
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   156
fun export_body thy name body =
68230
9bee37c2ac2b more scalable;
wenzelm
parents: 68208
diff changeset
   157
  Export.export thy ("theory/" ^ name) (Buffer.chunks (YXML.buffer_body body Buffer.empty));
68165
a7a2174ac014 more exports;
wenzelm
parents: 68154
diff changeset
   158
68154
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
   159
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   160
(* presentation *)
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   161
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   162
val _ = setup_presentation (fn {adjust_pos, ...} => fn thy =>
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   163
  let
68900
1145b25c53de more robust: parents are strict in Export_Theory.read_theory and thus approximate "commit" of exports;
wenzelm
parents: 68864
diff changeset
   164
    val parents = Theory.parents_of thy;
68264
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   165
    val rep_tsig = Type.rep_tsig (Sign.tsig_of thy);
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   166
69003
a015f1d3ba0c export plain infix syntax;
wenzelm
parents: 68997
diff changeset
   167
    val thy_ctxt = Proof_Context.init_global thy;
68997
4278947ba336 more exports;
wenzelm
parents: 68900
diff changeset
   168
68264
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   169
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   170
    (* entities *)
68154
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
   171
69069
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   172
    fun make_entity_markup name xname pos serial =
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   173
      let
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   174
        val props =
68830
44ec6fdaacf8 retain original id, which is command_id/exec_id for PIDE;
wenzelm
parents: 68727
diff changeset
   175
          Position.offset_properties_of (adjust_pos pos) @
44ec6fdaacf8 retain original id, which is command_id/exec_id for PIDE;
wenzelm
parents: 68727
diff changeset
   176
          Position.id_properties_of pos @
44ec6fdaacf8 retain original id, which is command_id/exec_id for PIDE;
wenzelm
parents: 68727
diff changeset
   177
          Markup.serial_properties serial;
68997
4278947ba336 more exports;
wenzelm
parents: 68900
diff changeset
   178
      in (Markup.entityN, (Markup.nameN, name) :: (Markup.xnameN, xname) :: props) end;
68154
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
   179
69069
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   180
    fun entity_markup space name =
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   181
      let
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   182
        val xname = Name_Space.extern_shortest thy_ctxt space name;
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   183
        val {serial, pos, ...} = Name_Space.the_entry space name;
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   184
      in make_entity_markup name xname pos serial end;
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   185
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   186
    fun export_entities export_name export get_space decls =
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   187
      let val elems =
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   188
        let
68206
dedf1a70d1fa export more theory and session structure;
wenzelm
parents: 68201
diff changeset
   189
          val parent_spaces = map get_space parents;
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   190
          val space = get_space thy;
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   191
        in
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   192
          (decls, []) |-> fold (fn (name, decl) =>
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   193
            if exists (fn space => Name_Space.declared space name) parent_spaces then I
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   194
            else
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   195
              (case export name decl of
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   196
                NONE => I
68724
7fafadbf16c7 export in foundational order;
wenzelm
parents: 68539
diff changeset
   197
              | SOME body =>
7fafadbf16c7 export in foundational order;
wenzelm
parents: 68539
diff changeset
   198
                  cons (#serial (Name_Space.the_entry space name),
7fafadbf16c7 export in foundational order;
wenzelm
parents: 68539
diff changeset
   199
                    XML.Elem (entity_markup space name, body))))
7fafadbf16c7 export in foundational order;
wenzelm
parents: 68539
diff changeset
   200
          |> sort (int_ord o apply2 #1) |> map #2
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   201
        end;
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   202
      in if null elems then () else export_body thy export_name elems end;
68165
a7a2174ac014 more exports;
wenzelm
parents: 68154
diff changeset
   203
a7a2174ac014 more exports;
wenzelm
parents: 68154
diff changeset
   204
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   205
    (* types *)
68165
a7a2174ac014 more exports;
wenzelm
parents: 68154
diff changeset
   206
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   207
    val encode_type =
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   208
      let open XML.Encode Term_XML.Encode
69077
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
   209
      in triple encode_syntax (list string) (option typ) end;
68165
a7a2174ac014 more exports;
wenzelm
parents: 68154
diff changeset
   210
69003
a015f1d3ba0c export plain infix syntax;
wenzelm
parents: 68997
diff changeset
   211
    fun export_type c (Type.LogicalType n) =
69077
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
   212
          SOME (encode_type (get_syntax_type thy_ctxt c, Name.invent Name.context Name.aT n, NONE))
69003
a015f1d3ba0c export plain infix syntax;
wenzelm
parents: 68997
diff changeset
   213
      | export_type c (Type.Abbreviation (args, U, false)) =
69077
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
   214
          SOME (encode_type (get_syntax_type thy_ctxt c, args, SOME U))
69003
a015f1d3ba0c export plain infix syntax;
wenzelm
parents: 68997
diff changeset
   215
      | export_type _ _ = NONE;
68165
a7a2174ac014 more exports;
wenzelm
parents: 68154
diff changeset
   216
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   217
    val _ =
69003
a015f1d3ba0c export plain infix syntax;
wenzelm
parents: 68997
diff changeset
   218
      export_entities "types" export_type Sign.type_space
68264
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   219
        (Name_Space.dest_table (#types rep_tsig));
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   220
68173
7ed88a534bb6 more uniform types vs. consts;
wenzelm
parents: 68172
diff changeset
   221
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   222
    (* consts *)
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   223
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   224
    val encode_const =
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   225
      let open XML.Encode Term_XML.Encode
69077
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
   226
      in pair encode_syntax (pair (list string) (pair typ (option term))) end;
68165
a7a2174ac014 more exports;
wenzelm
parents: 68154
diff changeset
   227
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   228
    fun export_const c (T, abbrev) =
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   229
      let
69077
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
   230
        val syntax = get_syntax_const thy_ctxt c;
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   231
        val T' = T |> Logic.unvarifyT_global |> Type.strip_sorts;
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   232
        val abbrev' = abbrev |> Option.map (standard_vars_global #> map_types Type.strip_sorts);
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   233
        val args = map (#1 o dest_TFree) (Consts.typargs (Sign.consts_of thy) (c, T'));
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   234
      in encode_const (syntax, (args, (T', abbrev'))) end;
68165
a7a2174ac014 more exports;
wenzelm
parents: 68154
diff changeset
   235
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   236
    val _ =
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   237
      export_entities "consts" (SOME oo export_const) Sign.const_space
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   238
        (#constants (Consts.dest (Sign.consts_of thy)));
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   239
68208
d9f2cf4fc002 more exports;
wenzelm
parents: 68206
diff changeset
   240
68232
4b93573ac5b4 export facts;
wenzelm
parents: 68231
diff changeset
   241
    (* axioms and facts *)
68208
d9f2cf4fc002 more exports;
wenzelm
parents: 68206
diff changeset
   242
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   243
    fun prop_of raw_thm =
68727
ec0b2833cfc4 export shyps as regular typargs;
wenzelm
parents: 68726
diff changeset
   244
      let
ec0b2833cfc4 export shyps as regular typargs;
wenzelm
parents: 68726
diff changeset
   245
        val thm = raw_thm
ec0b2833cfc4 export shyps as regular typargs;
wenzelm
parents: 68726
diff changeset
   246
          |> Thm.transfer thy
ec0b2833cfc4 export shyps as regular typargs;
wenzelm
parents: 68726
diff changeset
   247
          |> Thm.check_hyps (Context.Theory thy)
ec0b2833cfc4 export shyps as regular typargs;
wenzelm
parents: 68726
diff changeset
   248
          |> Thm.strip_shyps;
ec0b2833cfc4 export shyps as regular typargs;
wenzelm
parents: 68726
diff changeset
   249
        val prop = thm
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   250
          |> Thm.full_prop_of;
68727
ec0b2833cfc4 export shyps as regular typargs;
wenzelm
parents: 68726
diff changeset
   251
      in (Thm.extra_shyps thm, prop) end;
68539
6740e3611a86 disallow hyps in export;
wenzelm
parents: 68295
diff changeset
   252
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   253
    fun encode_prop used (Ss, raw_prop) =
68208
d9f2cf4fc002 more exports;
wenzelm
parents: 68206
diff changeset
   254
      let
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   255
        val prop = standard_vars used raw_prop;
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   256
        val args = rev (add_frees used prop []);
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   257
        val typargs = rev (add_tfrees used prop []);
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   258
        val used' = fold (Name.declare o #1) typargs used;
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   259
        val sorts = Name.invent used' Name.aT (length Ss) ~~ Ss;
68726
782d6b89fb19 more uniform facts: single vs. multi;
wenzelm
parents: 68725
diff changeset
   260
      in
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   261
        (sorts @ typargs, args, prop) |>
68727
ec0b2833cfc4 export shyps as regular typargs;
wenzelm
parents: 68726
diff changeset
   262
          let open XML.Encode Term_XML.Encode
ec0b2833cfc4 export shyps as regular typargs;
wenzelm
parents: 68726
diff changeset
   263
          in triple (list (pair string sort)) (list (pair string typ)) term end
68726
782d6b89fb19 more uniform facts: single vs. multi;
wenzelm
parents: 68725
diff changeset
   264
      end;
68208
d9f2cf4fc002 more exports;
wenzelm
parents: 68206
diff changeset
   265
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   266
    fun encode_axiom used t = encode_prop used ([], t);
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   267
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   268
    val encode_fact_single = encode_prop Name.context o prop_of;
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   269
    val encode_fact_multi = XML.Encode.list (encode_prop Name.context) o map prop_of;
68208
d9f2cf4fc002 more exports;
wenzelm
parents: 68206
diff changeset
   270
68232
4b93573ac5b4 export facts;
wenzelm
parents: 68231
diff changeset
   271
    val _ =
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   272
      export_entities "axioms" (fn _ => fn t => SOME (encode_axiom Name.context t))
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   273
        Theory.axiom_space (Theory.axioms_of thy);
68235
a3bd410db5b2 standardize implicit variables: non-zero indexes do occur occasionally, e.g. via RS;
wenzelm
parents: 68232
diff changeset
   274
    val _ =
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   275
      export_entities "facts" (K (SOME o encode_fact_multi))
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   276
        (Facts.space_of o Global_Theory.facts_of)
68232
4b93573ac5b4 export facts;
wenzelm
parents: 68231
diff changeset
   277
        (Facts.dest_static true [] (Global_Theory.facts_of thy));
4b93573ac5b4 export facts;
wenzelm
parents: 68231
diff changeset
   278
68264
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   279
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   280
    (* type classes *)
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   281
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   282
    val encode_class =
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   283
      let open XML.Encode Term_XML.Encode
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   284
      in pair (list (pair string typ)) (list encode_fact_single) end;
68264
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   285
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   286
    fun export_class name =
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   287
      (case try (Axclass.get_info thy) name of
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   288
        NONE => ([], [])
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   289
      | SOME {params, axioms, ...} => (params, axioms))
68264
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   290
      |> encode_class |> SOME;
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   291
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   292
    val _ =
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   293
      export_entities "classes" (fn name => fn () => export_class name)
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   294
        Sign.class_space (map (rpair ()) (Graph.keys (Sorts.classes_of (#2 (#classes rep_tsig)))));
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   295
68295
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   296
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   297
    (* sort algebra *)
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   298
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   299
    val {classrel, arities} =
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   300
      Sorts.dest_algebra (map (#2 o #classes o Type.rep_tsig o Sign.tsig_of) parents)
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   301
        (#2 (#classes rep_tsig));
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   302
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   303
    val encode_classrel =
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   304
      let open XML.Encode
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   305
      in list (pair string (list string)) end;
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   306
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   307
    val encode_arities =
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   308
      let open XML.Encode Term_XML.Encode
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   309
      in list (triple string (list sort) string) end;
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   310
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   311
    val _ = if null classrel then () else export_body thy "classrel" (encode_classrel classrel);
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   312
    val _ = if null arities then () else export_body thy "arities" (encode_arities arities);
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   313
68862
47e9912c53c3 export locale content;
wenzelm
parents: 68830
diff changeset
   314
47e9912c53c3 export locale content;
wenzelm
parents: 68830
diff changeset
   315
    (* locales *)
47e9912c53c3 export locale content;
wenzelm
parents: 68830
diff changeset
   316
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   317
    fun encode_locale used =
69076
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
   318
      let open XML.Encode Term_XML.Encode in
69077
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
   319
        triple (list (pair string sort)) (list (pair (pair string typ) encode_syntax))
69076
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
   320
          (list (encode_axiom used))
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
   321
      end;
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   322
69029
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   323
    fun export_locale loc =
68864
1dacce27bc25 clarified signature: proper typargs;
wenzelm
parents: 68862
diff changeset
   324
      let
69034
855c3c501b09 obsolete (see aec64b88e708);
wenzelm
parents: 69029
diff changeset
   325
        val {typargs, args, axioms} = locale_content thy loc;
69076
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
   326
        val used = fold Name.declare (map #1 typargs @ map (#1 o #1) args) Name.context;
69027
5ea3f424e787 clarified errors;
wenzelm
parents: 69023
diff changeset
   327
      in encode_locale used (typargs, args, axioms) end
5ea3f424e787 clarified errors;
wenzelm
parents: 69023
diff changeset
   328
      handle ERROR msg =>
5ea3f424e787 clarified errors;
wenzelm
parents: 69023
diff changeset
   329
        cat_error msg ("The error(s) above occurred in locale " ^
5ea3f424e787 clarified errors;
wenzelm
parents: 69023
diff changeset
   330
          quote (Locale.markup_name thy_ctxt loc));
68862
47e9912c53c3 export locale content;
wenzelm
parents: 68830
diff changeset
   331
47e9912c53c3 export locale content;
wenzelm
parents: 68830
diff changeset
   332
    val _ =
69029
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   333
      export_entities "locales" (fn loc => fn () => SOME (export_locale loc))
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   334
        Locale.locale_space
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
   335
        (map (rpair ()) (Locale.get_locales thy));
68862
47e9912c53c3 export locale content;
wenzelm
parents: 68830
diff changeset
   336
68900
1145b25c53de more robust: parents are strict in Export_Theory.read_theory and thus approximate "commit" of exports;
wenzelm
parents: 68864
diff changeset
   337
69069
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   338
    (* locale dependencies *)
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   339
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   340
    fun encode_locale_dependency (dep: Locale.locale_dependency) =
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   341
      (#source dep, (#target dep, (#prefix dep, locale_dependency_subst thy dep))) |>
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   342
        let
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   343
          open XML.Encode Term_XML.Encode;
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   344
          val encode_subst =
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   345
            pair (list (pair (pair string sort) typ)) (list (pair (pair string typ) term));
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   346
        in pair string (pair string (pair (list (pair string bool)) encode_subst)) end;
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   347
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   348
    val _ =
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   349
      (case Locale.dest_dependencies parents thy of
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   350
        [] => ()
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   351
      | deps =>
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   352
          deps |> map_index (fn (i, dep) =>
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   353
            let
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   354
              val xname = string_of_int (i + 1);
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   355
              val name = Long_Name.implode [Context.theory_name thy, xname];
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   356
              val body = encode_locale_dependency dep;
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   357
            in XML.Elem (make_entity_markup name xname (#pos dep) (#serial dep), body) end)
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   358
          |> export_body thy "locale_dependencies");
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   359
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   360
68900
1145b25c53de more robust: parents are strict in Export_Theory.read_theory and thus approximate "commit" of exports;
wenzelm
parents: 68864
diff changeset
   361
    (* parents *)
1145b25c53de more robust: parents are strict in Export_Theory.read_theory and thus approximate "commit" of exports;
wenzelm
parents: 68864
diff changeset
   362
1145b25c53de more robust: parents are strict in Export_Theory.read_theory and thus approximate "commit" of exports;
wenzelm
parents: 68864
diff changeset
   363
    val _ =
1145b25c53de more robust: parents are strict in Export_Theory.read_theory and thus approximate "commit" of exports;
wenzelm
parents: 68864
diff changeset
   364
      export_body thy "parents"
1145b25c53de more robust: parents are strict in Export_Theory.read_theory and thus approximate "commit" of exports;
wenzelm
parents: 68864
diff changeset
   365
        (XML.Encode.string (cat_lines (map Context.theory_long_name parents)));
1145b25c53de more robust: parents are strict in Export_Theory.read_theory and thus approximate "commit" of exports;
wenzelm
parents: 68864
diff changeset
   366
68295
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   367
  in () end);
68165
a7a2174ac014 more exports;
wenzelm
parents: 68154
diff changeset
   368
a7a2174ac014 more exports;
wenzelm
parents: 68154
diff changeset
   369
end;