src/Pure/Build/export_theory.ML
author wenzelm
Wed, 12 Mar 2025 11:39:00 +0100
changeset 82265 4b875a4c83b0
parent 81512 c1aa8a61ee65
permissions -rw-r--r--
update for release;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
79502
c7a98469c0e7 clarified directories;
wenzelm
parents: 79469
diff changeset
     1
(*  Title:      Pure/Build/export_theory.ML
68154
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
74261
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
     9
  val other_name_space: (theory -> Name_Space.T) -> theory -> theory
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
    10
  val export_enabled: Thy_Info.presentation_context -> bool
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
    11
  val export_body: theory -> string -> XML.body -> unit
68154
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
    12
end;
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
    13
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
    14
structure Export_Theory: EXPORT_THEORY =
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
    15
struct
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
    16
74261
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
    17
(* other name spaces *)
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
    18
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
    19
structure Data = Theory_Data
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
    20
(
74264
04214caeb9ac simplified: uniqueness check happens in export_consumer;
wenzelm
parents: 74261
diff changeset
    21
  type T = (theory -> Name_Space.T) Inttab.table;
04214caeb9ac simplified: uniqueness check happens in export_consumer;
wenzelm
parents: 74261
diff changeset
    22
  val empty = Inttab.empty;
04214caeb9ac simplified: uniqueness check happens in export_consumer;
wenzelm
parents: 74261
diff changeset
    23
  val merge = Inttab.merge (K true);
74261
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
    24
);
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
    25
74264
04214caeb9ac simplified: uniqueness check happens in export_consumer;
wenzelm
parents: 74261
diff changeset
    26
val other_name_spaces = map #2 o Inttab.dest o Data.get;
04214caeb9ac simplified: uniqueness check happens in export_consumer;
wenzelm
parents: 74261
diff changeset
    27
fun other_name_space get_space thy = Data.map (Inttab.update (serial (), get_space)) thy;
74261
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
    28
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
    29
val _ = Theory.setup
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
    30
 (other_name_space Thm.oracle_space #>
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
    31
  other_name_space Global_Theory.fact_space #>
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
    32
  other_name_space (Bundle.bundle_space o Context.Theory) #>
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
    33
  other_name_space (Attrib.attribute_space o Context.Theory) #>
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
    34
  other_name_space (Method.method_space o Context.Theory));
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
    35
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
    36
69077
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    37
(* approximative syntax *)
69076
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
    38
80074
951c371c1cd9 clarified names: discontinue odd convention from 3 decades ago;
wenzelm
parents: 79502
diff changeset
    39
val get_syntax = Syntax.get_approx o Proof_Context.syntax_of;
69077
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    40
fun get_syntax_type ctxt = get_syntax ctxt o Lexicon.mark_type;
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    41
fun get_syntax_const ctxt = get_syntax ctxt o Lexicon.mark_const;
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    42
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
    43
69077
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    44
fun get_syntax_param ctxt loc x =
69076
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
    45
  let val thy = Proof_Context.theory_of ctxt in
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
    46
    if Class.is_class thy loc then
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
    47
      (case AList.lookup (op =) (Class.these_params thy [loc]) x of
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
    48
        NONE => NONE
69077
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    49
      | SOME (_, (c, _)) => get_syntax_const ctxt c)
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    50
    else get_syntax_fixed ctxt x
69076
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
    51
  end;
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
    52
69077
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    53
val encode_syntax =
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    54
  XML.Encode.variant
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    55
   [fn NONE => ([], []),
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    56
    fn SOME (Syntax.Prefix delim) => ([delim], []),
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    57
    fn SOME (Syntax.Infix {assoc, delim, pri}) =>
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    58
      let
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    59
        val ass =
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    60
          (case assoc of
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    61
            Printer.No_Assoc => 0
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    62
          | Printer.Left_Assoc => 1
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    63
          | Printer.Right_Assoc => 2);
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    64
        open XML.Encode Term_XML.Encode;
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
    65
      in ([], triple int string int (ass, delim, pri)) end];
69076
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
    66
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
    67
71209
8508cc7f79aa tuned comment;
wenzelm
parents: 71208
diff changeset
    68
(* locales *)
69019
a6ba77af6486 export semi-unfolded locale axioms;
wenzelm
parents: 69003
diff changeset
    69
69034
855c3c501b09 obsolete (see aec64b88e708);
wenzelm
parents: 69029
diff changeset
    70
fun locale_content thy loc =
69019
a6ba77af6486 export semi-unfolded locale axioms;
wenzelm
parents: 69003
diff changeset
    71
  let
69083
6f8ae6ddc26b more direct locale goal: avoid renaming of type_parameters;
wenzelm
parents: 69077
diff changeset
    72
    val ctxt = Locale.init loc thy;
6f8ae6ddc26b more direct locale goal: avoid renaming of type_parameters;
wenzelm
parents: 69077
diff changeset
    73
    val args =
6f8ae6ddc26b more direct locale goal: avoid renaming of type_parameters;
wenzelm
parents: 69077
diff changeset
    74
      Locale.params_of thy loc
6f8ae6ddc26b more direct locale goal: avoid renaming of type_parameters;
wenzelm
parents: 69077
diff changeset
    75
      |> map (fn ((x, T), _) => ((x, T), get_syntax_param 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
    76
    val axioms =
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
    77
      let
69083
6f8ae6ddc26b more direct locale goal: avoid renaming of type_parameters;
wenzelm
parents: 69077
diff changeset
    78
        val (asm, defs) = Locale.specification_of thy loc;
6f8ae6ddc26b more direct locale goal: avoid renaming of type_parameters;
wenzelm
parents: 69077
diff changeset
    79
        val cprops = map (Thm.cterm_of ctxt) (the_list asm @ defs);
69029
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
    80
        val (intro1, intro2) = Locale.intros_of thy loc;
69083
6f8ae6ddc26b more direct locale goal: avoid renaming of type_parameters;
wenzelm
parents: 69077
diff changeset
    81
        val intros_tac = Method.try_intros_tac ctxt (the_list intro1 @ the_list intro2) [];
69029
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
    82
        val res =
69083
6f8ae6ddc26b more direct locale goal: avoid renaming of type_parameters;
wenzelm
parents: 69077
diff changeset
    83
          Goal.init (Conjunction.mk_conjunction_balanced cprops)
6f8ae6ddc26b more direct locale goal: avoid renaming of type_parameters;
wenzelm
parents: 69077
diff changeset
    84
          |> (ALLGOALS Goal.conjunction_tac THEN intros_tac)
69029
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
    85
          |> 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
    86
      in
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
    87
        (case res of
69083
6f8ae6ddc26b more direct locale goal: avoid renaming of type_parameters;
wenzelm
parents: 69077
diff changeset
    88
          SOME goal => Thm.prems_of goal
69029
aec64b88e708 clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
wenzelm
parents: 69028
diff changeset
    89
        | 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
    90
      end;
74235
dbaed92fd8af tuned signature;
wenzelm
parents: 74234
diff changeset
    91
    val typargs = build_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
    92
  in {typargs = typargs, args = args, axioms = axioms} end;
69019
a6ba77af6486 export semi-unfolded locale axioms;
wenzelm
parents: 69003
diff changeset
    93
69087
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
    94
fun get_locales thy =
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
    95
  Locale.get_locales thy |> map_filter (fn loc =>
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
    96
    if Experiment.is_experiment thy loc then NONE else SOME (loc, ()));
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
    97
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
    98
fun get_dependencies prev_thys thy =
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
    99
  Locale.dest_dependencies prev_thys thy |> map_filter (fn dep =>
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   100
    if Experiment.is_experiment thy (#source dep) orelse
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   101
      Experiment.is_experiment thy (#target dep) then NONE
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   102
    else
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   103
      let
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   104
        val (type_params, params) = Locale.parameters_of thy (#source dep);
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   105
        val typargs = fold (Term.add_tfreesT o #2 o #1) params type_params;
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   106
        val substT =
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   107
          typargs |> map_filter (fn v =>
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   108
            let
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   109
              val T = TFree v;
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   110
              val T' = Morphism.typ (#morphism dep) T;
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   111
            in if T = T' then NONE else SOME (v, T') end);
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   112
        val subst =
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   113
          params |> map_filter (fn (v, _) =>
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   114
            let
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   115
              val t = Free v;
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   116
              val t' = Morphism.term (#morphism dep) t;
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   117
            in if t aconv t' then NONE else SOME (v, t') end);
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   118
      in SOME (dep, (substT, subst)) end);
69069
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   119
69019
a6ba77af6486 export semi-unfolded locale axioms;
wenzelm
parents: 69003
diff changeset
   120
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   121
(* presentation *)
68154
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
   122
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   123
fun export_enabled (context: Thy_Info.presentation_context) =
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   124
  Options.bool (#options context) "export_theory";
68154
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
   125
70990
e5e34bd28257 clarified signature;
wenzelm
parents: 70984
diff changeset
   126
fun export_body thy name body =
70994
ff11af12dfdd clarified signature;
wenzelm
parents: 70992
diff changeset
   127
  if XML.is_empty_body body then ()
74261
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
   128
  else Export.export thy (Path.binding0 (Path.make ("theory" :: space_explode "/" name))) body;
68165
a7a2174ac014 more exports;
wenzelm
parents: 68154
diff changeset
   129
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   130
val _ = (Theory.setup o Thy_Info.add_presentation) (fn context => fn thy =>
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   131
  let
68264
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   132
    val rep_tsig = Type.rep_tsig (Sign.tsig_of thy);
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   133
    val consts = Sign.consts_of thy;
69003
a015f1d3ba0c export plain infix syntax;
wenzelm
parents: 68997
diff changeset
   134
    val thy_ctxt = Proof_Context.init_global thy;
68997
4278947ba336 more exports;
wenzelm
parents: 68900
diff changeset
   135
71249
877316c54ed3 clarified signature;
wenzelm
parents: 71222
diff changeset
   136
    val pos_properties = Thy_Info.adjust_pos_properties context;
877316c54ed3 clarified signature;
wenzelm
parents: 71222
diff changeset
   137
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   138
    val enabled = export_enabled context;
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   139
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   140
80608
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   141
    (* recode *)
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   142
80613
42408be39d6c disable thy_cache for now (amending 0b8922e351a3): avoid crash of AFP/Ramsey-Infinite due to exception THEORY "Duplicate theory name";
wenzelm
parents: 80608
diff changeset
   143
    val thy_cache = thy;  (* FIXME tmp *)
80608
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   144
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   145
    val ztyp_of = ZTerm.ztyp_cache thy_cache;
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   146
    val zterm_of = ZTerm.zterm_of thy_cache;
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   147
    val zproof_of = Proofterm.proof_to_zproof thy_cache;
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   148
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   149
    val encode_ztyp = ZTerm.encode_ztyp;
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   150
    val encode_zterm = ZTerm.encode_zterm {typed_vars = true};
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   151
    val encode_term = encode_zterm o zterm_of;
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   152
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   153
    val encode_standard_zterm = ZTerm.encode_zterm {typed_vars = false};
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   154
    val encode_standard_zproof = ZTerm.encode_zproof {typed_vars = false};
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   155
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   156
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   157
    (* strict parents *)
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   158
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   159
    val parents = Theory.parents_of thy;
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   160
    val _ =
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   161
      Export.export thy \<^path_binding>\<open>theory/parents\<close>
75858
657b2de27454 clarified signature: avoid confusing operation name (amending 8cf14d4ebec4);
wenzelm
parents: 75763
diff changeset
   162
        (XML.Encode.string (cat_lines (map Context.theory_long_name parents) ^ "\n"));
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   163
68264
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   164
71218
73b313432d8a clarified position for spec rule: like entity;
wenzelm
parents: 71215
diff changeset
   165
    (* spec rules *)
73b313432d8a clarified position for spec rule: like entity;
wenzelm
parents: 71215
diff changeset
   166
73b313432d8a clarified position for spec rule: like entity;
wenzelm
parents: 71215
diff changeset
   167
    fun spec_rule_content {pos, name, rough_classification, terms, rules} =
73b313432d8a clarified position for spec rule: like entity;
wenzelm
parents: 71215
diff changeset
   168
      let
73b313432d8a clarified position for spec rule: like entity;
wenzelm
parents: 71215
diff changeset
   169
        val spec =
73b313432d8a clarified position for spec rule: like entity;
wenzelm
parents: 71215
diff changeset
   170
          terms @ map Thm.plain_prop_of rules
73b313432d8a clarified position for spec rule: like entity;
wenzelm
parents: 71215
diff changeset
   171
          |> Term_Subst.zero_var_indexes_list
73b313432d8a clarified position for spec rule: like entity;
wenzelm
parents: 71215
diff changeset
   172
          |> map Logic.unvarify_global;
73b313432d8a clarified position for spec rule: like entity;
wenzelm
parents: 71215
diff changeset
   173
      in
71249
877316c54ed3 clarified signature;
wenzelm
parents: 71222
diff changeset
   174
       {props = pos_properties pos,
71218
73b313432d8a clarified position for spec rule: like entity;
wenzelm
parents: 71215
diff changeset
   175
        name = name,
73b313432d8a clarified position for spec rule: like entity;
wenzelm
parents: 71215
diff changeset
   176
        rough_classification = rough_classification,
74235
dbaed92fd8af tuned signature;
wenzelm
parents: 74234
diff changeset
   177
        typargs = build_rev (fold Term.add_tfrees spec),
dbaed92fd8af tuned signature;
wenzelm
parents: 74234
diff changeset
   178
        args = build_rev (fold Term.add_frees spec),
71218
73b313432d8a clarified position for spec rule: like entity;
wenzelm
parents: 71215
diff changeset
   179
        terms = map (fn t => (t, Term.type_of t)) (take (length terms) spec),
73b313432d8a clarified position for spec rule: like entity;
wenzelm
parents: 71215
diff changeset
   180
        rules = drop (length terms) spec}
73b313432d8a clarified position for spec rule: like entity;
wenzelm
parents: 71215
diff changeset
   181
      end;
73b313432d8a clarified position for spec rule: like entity;
wenzelm
parents: 71215
diff changeset
   182
73b313432d8a clarified position for spec rule: like entity;
wenzelm
parents: 71215
diff changeset
   183
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   184
    (* entities *)
68154
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
   185
69069
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   186
    fun make_entity_markup name xname pos serial =
71249
877316c54ed3 clarified signature;
wenzelm
parents: 71222
diff changeset
   187
      let val props = pos_properties pos @ Markup.serial_properties serial;
68997
4278947ba336 more exports;
wenzelm
parents: 68900
diff changeset
   188
      in (Markup.entityN, (Markup.nameN, name) :: (Markup.xnameN, xname) :: props) end;
68154
42d63ea39161 some export of foundational theory content;
wenzelm
parents:
diff changeset
   189
69069
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   190
    fun entity_markup space name =
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   191
      let
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   192
        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
   193
        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
   194
      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
   195
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   196
    fun export_entities export_name get_space decls export =
70601
79831e40e2be more scalable: avoid huge intermediate XML elems;
wenzelm
parents: 70597
diff changeset
   197
      let
79831e40e2be more scalable: avoid huge intermediate XML elems;
wenzelm
parents: 70597
diff changeset
   198
        val parent_spaces = map get_space parents;
79831e40e2be more scalable: avoid huge intermediate XML elems;
wenzelm
parents: 70597
diff changeset
   199
        val space = get_space thy;
79831e40e2be more scalable: avoid huge intermediate XML elems;
wenzelm
parents: 70597
diff changeset
   200
      in
74234
4f2bd13edce3 clarified signature;
wenzelm
parents: 74114
diff changeset
   201
        build (decls |> fold (fn (name, decl) =>
70601
79831e40e2be more scalable: avoid huge intermediate XML elems;
wenzelm
parents: 70597
diff changeset
   202
          if exists (fn space => Name_Space.declared space name) parent_spaces then I
79831e40e2be more scalable: avoid huge intermediate XML elems;
wenzelm
parents: 70597
diff changeset
   203
          else
79831e40e2be more scalable: avoid huge intermediate XML elems;
wenzelm
parents: 70597
diff changeset
   204
            (case export name decl of
79831e40e2be more scalable: avoid huge intermediate XML elems;
wenzelm
parents: 70597
diff changeset
   205
              NONE => I
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   206
            | SOME make_body =>
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   207
                let
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   208
                  val i = #serial (Name_Space.the_entry space name);
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   209
                  val body = if enabled then make_body () else [];
74234
4f2bd13edce3 clarified signature;
wenzelm
parents: 74114
diff changeset
   210
                in cons (i, XML.Elem (entity_markup space name, body)) end)))
70601
79831e40e2be more scalable: avoid huge intermediate XML elems;
wenzelm
parents: 70597
diff changeset
   211
        |> sort (int_ord o apply2 #1) |> map #2
79831e40e2be more scalable: avoid huge intermediate XML elems;
wenzelm
parents: 70597
diff changeset
   212
        |> export_body thy export_name
79831e40e2be more scalable: avoid huge intermediate XML elems;
wenzelm
parents: 70597
diff changeset
   213
      end;
68165
a7a2174ac014 more exports;
wenzelm
parents: 68154
diff changeset
   214
a7a2174ac014 more exports;
wenzelm
parents: 68154
diff changeset
   215
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   216
    (* types *)
68165
a7a2174ac014 more exports;
wenzelm
parents: 68154
diff changeset
   217
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   218
    val encode_type =
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   219
      let open XML.Encode Term_XML.Encode
69077
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
   220
      in triple encode_syntax (list string) (option typ) end;
68165
a7a2174ac014 more exports;
wenzelm
parents: 68154
diff changeset
   221
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   222
    val _ =
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   223
      export_entities "types" Sign.type_space (Name_Space.dest_table (#types rep_tsig))
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   224
        (fn c =>
79469
deb50d396ff7 tuned signature;
wenzelm
parents: 79411
diff changeset
   225
          (fn Type.Logical_Type n =>
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   226
                SOME (fn () =>
81507
08574da77b4a clarified signature: shorten common cases;
wenzelm
parents: 81503
diff changeset
   227
                  encode_type (get_syntax_type thy_ctxt c, Name.invent_global_types n, NONE))
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   228
            | Type.Abbreviation (args, U, false) =>
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   229
                SOME (fn () =>
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   230
                  encode_type (get_syntax_type thy_ctxt c, args, SOME U))
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   231
            | _ => NONE));
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   232
68173
7ed88a534bb6 more uniform types vs. consts;
wenzelm
parents: 68172
diff changeset
   233
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   234
    (* consts *)
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   235
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   236
    val encode_const =
71222
2bc39c80a95d clarified export of consts: recursion is accessible via spec_rules;
wenzelm
parents: 71218
diff changeset
   237
      let open XML.Encode Term_XML.Encode
80608
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   238
      in pair encode_syntax (pair (list string) (pair typ (pair (option encode_zterm) bool))) end;
68165
a7a2174ac014 more exports;
wenzelm
parents: 68154
diff changeset
   239
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   240
    val _ =
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   241
      export_entities "consts" Sign.const_space (#constants (Consts.dest consts))
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   242
        (fn c => fn (T, abbrev) =>
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   243
          SOME (fn () =>
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   244
            let
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   245
              val syntax = get_syntax_const thy_ctxt c;
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   246
              val U = Logic.unvarifyT_global T;
79411
700d4f16b5f2 minor performance tuning: proper Same.operation;
wenzelm
parents: 77979
diff changeset
   247
              val U0 = Term.strip_sortsT U;
80608
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   248
              fun trim_abbrev t =
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   249
                ZTerm.standard_vars Name.context (zterm_of t, NONE) |> #prop |> ZTerm.strip_sorts;
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   250
              val abbrev' = Option.map trim_abbrev abbrev;
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   251
              val args = map (#1 o dest_TFree) (Consts.typargs consts (c, U0));
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   252
              val propositional = Object_Logic.is_propositional thy_ctxt (Term.body_type U0);
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   253
            in encode_const (syntax, (args, (U0, (abbrev', propositional)))) end));
68201
dee993b88a7b misc tuning and clarification;
wenzelm
parents: 68199
diff changeset
   254
68208
d9f2cf4fc002 more exports;
wenzelm
parents: 68206
diff changeset
   255
70579
5a8e3e4b3760 clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents: 70541
diff changeset
   256
    (* axioms *)
68208
d9f2cf4fc002 more exports;
wenzelm
parents: 68206
diff changeset
   257
70534
fb876ebbf5a7 export facts with reconstructed proof term (if possible), but its PThm boxes need to be collected separately;
wenzelm
parents: 70529
diff changeset
   258
    fun standard_prop used extra_shyps raw_prop raw_proof =
68727
ec0b2833cfc4 export shyps as regular typargs;
wenzelm
parents: 68726
diff changeset
   259
      let
80608
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   260
        val {typargs, args, prop, proof} =
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   261
          ZTerm.standard_vars used (zterm_of raw_prop, Option.map zproof_of raw_proof);
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   262
        val is_free = not o Name.is_declared used;
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   263
        val args' = args |> filter (is_free o #1);
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   264
        val typargs' = typargs |> filter (is_free o #1);
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   265
        val used_typargs = fold (Name.declare o #1) typargs' used;
81512
c1aa8a61ee65 tuned: more direct use of Name.context operations;
wenzelm
parents: 81507
diff changeset
   266
        val sorts = Name.invent_types used_typargs extra_shyps;
80608
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   267
      in ((sorts @ typargs', args', prop), proof) end;
70534
fb876ebbf5a7 export facts with reconstructed proof term (if possible), but its PThm boxes need to be collected separately;
wenzelm
parents: 70529
diff changeset
   268
71202
785610ad6bfa export spec rules;
wenzelm
parents: 71179
diff changeset
   269
    fun standard_prop_of thm =
77869
1156aa9db7f5 backout 4a174bea55e2;
wenzelm
parents: 77730
diff changeset
   270
      standard_prop Name.context (Thm.extra_shyps thm) (Thm.full_prop_of thm);
71202
785610ad6bfa export spec rules;
wenzelm
parents: 71179
diff changeset
   271
70534
fb876ebbf5a7 export facts with reconstructed proof term (if possible), but its PThm boxes need to be collected separately;
wenzelm
parents: 70529
diff changeset
   272
    val encode_prop =
fb876ebbf5a7 export facts with reconstructed proof term (if possible), but its PThm boxes need to be collected separately;
wenzelm
parents: 70529
diff changeset
   273
      let open XML.Encode Term_XML.Encode
80608
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   274
      in triple (list (pair string sort)) (list (pair string encode_ztyp)) encode_zterm end;
70534
fb876ebbf5a7 export facts with reconstructed proof term (if possible), but its PThm boxes need to be collected separately;
wenzelm
parents: 70529
diff changeset
   275
fb876ebbf5a7 export facts with reconstructed proof term (if possible), but its PThm boxes need to be collected separately;
wenzelm
parents: 70529
diff changeset
   276
    fun encode_axiom used prop =
fb876ebbf5a7 export facts with reconstructed proof term (if possible), but its PThm boxes need to be collected separately;
wenzelm
parents: 70529
diff changeset
   277
      encode_prop (#1 (standard_prop used [] prop NONE));
68208
d9f2cf4fc002 more exports;
wenzelm
parents: 68206
diff changeset
   278
70579
5a8e3e4b3760 clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents: 70541
diff changeset
   279
    val _ =
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   280
      export_entities "axioms" Theory.axiom_space (Theory.all_axioms_of thy)
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   281
        (fn _ => fn prop => SOME (fn () => encode_axiom Name.context prop));
70579
5a8e3e4b3760 clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents: 70541
diff changeset
   282
5a8e3e4b3760 clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents: 70541
diff changeset
   283
70914
05c4c6a99b3f option to export standardized proof terms (not scalable);
wenzelm
parents: 70909
diff changeset
   284
    (* theorems and proof terms *)
05c4c6a99b3f option to export standardized proof terms (not scalable);
wenzelm
parents: 70909
diff changeset
   285
70895
2a318149b01b proof boxes based on proof digest (not proof term): thus it works with prune_proofs;
wenzelm
parents: 70892
diff changeset
   286
    val clean_thm = Thm.check_hyps (Context.Theory thy) #> Thm.strip_shyps;
71660
4269db8981b8 avoid vacous type variable, due to "potentially redundant" shyps in Thm.unconstrainT;
wenzelm
parents: 71249
diff changeset
   287
    val prep_thm = clean_thm #> Thm.unconstrainT #> Thm.strip_shyps;
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   288
70892
aadb5c23af24 clarified proof_boxes (requires prune_proofs=false);
wenzelm
parents: 70884
diff changeset
   289
    val lookup_thm_id = Global_Theory.lookup_thm_id thy;
aadb5c23af24 clarified proof_boxes (requires prune_proofs=false);
wenzelm
parents: 70884
diff changeset
   290
70915
bd4d37edfee4 clarified expand_proof/expand_name: allow more detailed control via thm_header;
wenzelm
parents: 70914
diff changeset
   291
    fun expand_name thm_id (header: Proofterm.thm_header) =
80590
505f97165f52 clarified thm_header command_pos vs. thm_pos;
wenzelm
parents: 80585
diff changeset
   292
      if #serial header = #serial thm_id then Thm_Name.none
505f97165f52 clarified thm_header command_pos vs. thm_pos;
wenzelm
parents: 80585
diff changeset
   293
      else the_default Thm_Name.none (lookup_thm_id (Proofterm.thm_header_id header));
70915
bd4d37edfee4 clarified expand_proof/expand_name: allow more detailed control via thm_header;
wenzelm
parents: 70914
diff changeset
   294
70579
5a8e3e4b3760 clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents: 70541
diff changeset
   295
    fun entity_markup_thm (serial, (name, i)) =
5a8e3e4b3760 clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents: 70541
diff changeset
   296
      let
74261
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
   297
        val space = Global_Theory.fact_space thy;
70579
5a8e3e4b3760 clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents: 70541
diff changeset
   298
        val xname = Name_Space.extern_shortest thy_ctxt space name;
5a8e3e4b3760 clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents: 70541
diff changeset
   299
        val {pos, ...} = Name_Space.the_entry space name;
5a8e3e4b3760 clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents: 70541
diff changeset
   300
      in make_entity_markup (Thm_Name.print (name, i)) (Thm_Name.print (xname, i)) pos serial end;
5a8e3e4b3760 clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents: 70541
diff changeset
   301
70915
bd4d37edfee4 clarified expand_proof/expand_name: allow more detailed control via thm_header;
wenzelm
parents: 70914
diff changeset
   302
    fun encode_thm thm_id raw_thm =
70597
a896257a3f07 export thm_deps;
wenzelm
parents: 70579
diff changeset
   303
      let
80313
a828e47c867c clarified data representation: prefer explicit type Thm_Name;
wenzelm
parents: 80295
diff changeset
   304
        val deps = map #2 (Thm_Deps.thm_deps thy [raw_thm]);
71202
785610ad6bfa export spec rules;
wenzelm
parents: 71179
diff changeset
   305
        val thm = prep_thm raw_thm;
70914
05c4c6a99b3f option to export standardized proof terms (not scalable);
wenzelm
parents: 70909
diff changeset
   306
05c4c6a99b3f option to export standardized proof terms (not scalable);
wenzelm
parents: 70909
diff changeset
   307
        val proof0 =
70983
52a62342c9ce clarified signature;
wenzelm
parents: 70981
diff changeset
   308
          if Proofterm.export_standard_enabled () then
71088
4b45d592ce29 clarified modules;
wenzelm
parents: 71018
diff changeset
   309
            Proof_Syntax.standard_proof_of
4b45d592ce29 clarified modules;
wenzelm
parents: 71018
diff changeset
   310
              {full = true, expand_name = SOME o expand_name thm_id} thm
70914
05c4c6a99b3f option to export standardized proof terms (not scalable);
wenzelm
parents: 70909
diff changeset
   311
          else if Proofterm.export_enabled () then Thm.reconstruct_proof_of thm
05c4c6a99b3f option to export standardized proof terms (not scalable);
wenzelm
parents: 70909
diff changeset
   312
          else MinProof;
71202
785610ad6bfa export spec rules;
wenzelm
parents: 71179
diff changeset
   313
        val (prop, SOME proof) = standard_prop_of thm (SOME proof0);
71018
d32ed8927a42 more robust expose_proofs corresponding to register_proofs/consolidate_theory;
wenzelm
parents: 71015
diff changeset
   314
        val _ = Thm.expose_proofs thy [thm];
70597
a896257a3f07 export thm_deps;
wenzelm
parents: 70579
diff changeset
   315
      in
71015
bb49abc2ecbb determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
wenzelm
parents: 70994
diff changeset
   316
        (prop, deps, proof) |>
80608
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   317
          let open XML.Encode Term_XML.Encode
0b8922e351a3 clarified export: replaced Proofterm.standard_vars by ZTerm.standard_vars;
wenzelm
parents: 80590
diff changeset
   318
          in triple encode_prop (list Thm_Name.encode) encode_standard_zproof end
70597
a896257a3f07 export thm_deps;
wenzelm
parents: 70579
diff changeset
   319
      end;
68208
d9f2cf4fc002 more exports;
wenzelm
parents: 68206
diff changeset
   320
80590
505f97165f52 clarified thm_header command_pos vs. thm_pos;
wenzelm
parents: 80585
diff changeset
   321
    fun export_thm (thm_id, (thm_name, _)) =
70601
79831e40e2be more scalable: avoid huge intermediate XML elems;
wenzelm
parents: 70597
diff changeset
   322
      let
70904
caf91f9b847b proper treatment of self thm_id;
wenzelm
parents: 70896
diff changeset
   323
        val markup = entity_markup_thm (#serial thm_id, thm_name);
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   324
        val body =
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   325
          if enabled then
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   326
            Global_Theory.get_thm_name thy (thm_name, Position.none)
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   327
            |> encode_thm thm_id
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   328
          else [];
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   329
      in XML.Elem (markup, body) end;
70579
5a8e3e4b3760 clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents: 70541
diff changeset
   330
70992
e7dfc505de1b more direct output of XML material -- bypass Buffer.T;
wenzelm
parents: 70991
diff changeset
   331
    val _ = export_body thy "thms" (map export_thm (Global_Theory.dest_thm_names thy));
68232
4b93573ac5b4 export facts;
wenzelm
parents: 68231
diff changeset
   332
68264
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   333
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   334
    (* type classes *)
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   335
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   336
    val encode_class =
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   337
      let open XML.Encode Term_XML.Encode
70534
fb876ebbf5a7 export facts with reconstructed proof term (if possible), but its PThm boxes need to be collected separately;
wenzelm
parents: 70529
diff changeset
   338
      in pair (list (pair string typ)) (list (encode_axiom Name.context)) end;
68264
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   339
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   340
    val _ =
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   341
      export_entities "classes" Sign.class_space
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   342
        (map (rpair ()) (Graph.keys (Sorts.classes_of (#2 (#classes rep_tsig)))))
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   343
        (fn name => fn () => SOME (fn () =>
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   344
          (case try (Axclass.get_info thy) name of
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   345
            NONE => ([], [])
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   346
          | SOME {params, axioms, ...} => (params, map (Thm.plain_prop_of o clean_thm) axioms))
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   347
          |> encode_class));
68264
bb9a3be6952a more exports;
wenzelm
parents: 68235
diff changeset
   348
68295
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   349
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   350
    (* sort algebra *)
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   351
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   352
    val _ =
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   353
      if enabled then
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   354
        let
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   355
          val prop = encode_axiom Name.context o Logic.varify_global;
68295
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   356
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   357
          val encode_classrel =
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   358
            let open XML.Encode
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   359
            in list (pair prop (pair string string)) end;
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   360
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   361
          val encode_arities =
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   362
            let open XML.Encode Term_XML.Encode
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   363
            in list (pair prop (triple string (list sort) string)) end;
70384
8ce08b154aa1 clarified export of sort algebra: avoid logical operations in Isabelle/Scala;
wenzelm
parents: 70015
diff changeset
   364
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   365
          val export_classrel =
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   366
            maps (fn (c, cs) => map (pair c) cs) #> map (`Logic.mk_classrel) #> encode_classrel;
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   367
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   368
          val export_arities = map (`Logic.mk_arity) #> encode_arities;
68295
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   369
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   370
          val {classrel, arities} =
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   371
            Sorts.dest_algebra (map (#2 o #classes o Type.rep_tsig o Sign.tsig_of) parents)
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   372
              (#2 (#classes rep_tsig));
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   373
        in
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   374
          if null classrel then () else export_body thy "classrel" (export_classrel classrel);
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   375
          if null arities then () else export_body thy "arities" (export_arities arities)
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   376
        end
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   377
      else ();
68295
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   378
68862
47e9912c53c3 export locale content;
wenzelm
parents: 68830
diff changeset
   379
47e9912c53c3 export locale content;
wenzelm
parents: 68830
diff changeset
   380
    (* locales *)
47e9912c53c3 export locale content;
wenzelm
parents: 68830
diff changeset
   381
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   382
    fun encode_locale used =
69076
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
   383
      let open XML.Encode Term_XML.Encode in
69077
11529ae45786 more approximative prefix syntax, including binder;
wenzelm
parents: 69076
diff changeset
   384
        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
   385
          (list (encode_axiom used))
90cce2f79e77 proper syntax for locale vs. class parameters;
wenzelm
parents: 69071
diff changeset
   386
      end;
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 69019
diff changeset
   387
68862
47e9912c53c3 export locale content;
wenzelm
parents: 68830
diff changeset
   388
    val _ =
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   389
      export_entities "locales" Locale.locale_space (get_locales thy)
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   390
        (fn loc => fn () => SOME (fn () =>
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   391
          let
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   392
            val {typargs, args, axioms} = locale_content thy loc;
81503
60e61a934a80 tuned signature;
wenzelm
parents: 80613
diff changeset
   393
            val used = Name.build_context (fold Name.declare (map #1 typargs @ map (#1 o #1) args));
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   394
          in encode_locale used (typargs, args, axioms) end
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   395
          handle ERROR msg =>
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   396
            cat_error msg ("The error(s) above occurred in locale " ^
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   397
              quote (Locale.markup_name thy_ctxt loc))));
68862
47e9912c53c3 export locale content;
wenzelm
parents: 68830
diff changeset
   398
68900
1145b25c53de more robust: parents are strict in Export_Theory.read_theory and thus approximate "commit" of exports;
wenzelm
parents: 68864
diff changeset
   399
69069
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   400
    (* locale dependencies *)
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   401
69087
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   402
    fun encode_locale_dependency (dep: Locale.locale_dependency, subst) =
06017b2c4552 suppress aux. locales from command 'experiment' -- avoid crash of theory Dict_Construction.Test_Dict_Construction (AFP);
wenzelm
parents: 69086
diff changeset
   403
      (#source dep, (#target dep, (#prefix dep, subst))) |>
69069
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   404
        let
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   405
          open XML.Encode Term_XML.Encode;
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   406
          val encode_subst =
70784
799437173553 Term_XML.Encode/Decode.term uses Const "typargs";
wenzelm
parents: 70605
diff changeset
   407
            pair (list (pair (pair string sort) typ)) (list (pair (pair string typ) (term consts)));
69069
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   408
        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
   409
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   410
    val _ =
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   411
      if enabled then
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   412
        get_dependencies parents thy |> map_index (fn (i, dep) =>
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   413
          let
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   414
            val xname = string_of_int (i + 1);
77889
5db014c36f42 clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
wenzelm
parents: 77869
diff changeset
   415
            val name = Long_Name.implode [Context.theory_base_name thy, xname];
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   416
            val markup = make_entity_markup name xname (#pos (#1 dep)) (#serial (#1 dep));
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   417
            val body = encode_locale_dependency dep;
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   418
          in XML.Elem (markup, body) end)
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   419
        |> export_body thy "locale_dependencies"
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   420
      else ();
69069
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   421
b9aca3b9619f export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents: 69034
diff changeset
   422
70920
1e0ad25c94c8 export constdefs according to defs.ML;
wenzelm
parents: 70919
diff changeset
   423
    (* constdefs *)
1e0ad25c94c8 export constdefs according to defs.ML;
wenzelm
parents: 70919
diff changeset
   424
1e0ad25c94c8 export constdefs according to defs.ML;
wenzelm
parents: 70919
diff changeset
   425
    val _ =
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   426
      if enabled then
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   427
        let
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   428
          val constdefs =
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   429
            Defs.dest_constdefs (map Theory.defs_of (Theory.parents_of thy)) (Theory.defs_of thy)
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   430
            |> sort_by #1;
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   431
          val encode =
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   432
            let open XML.Encode
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   433
            in list (pair string string) end;
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   434
        in if null constdefs then () else export_body thy "constdefs" (encode constdefs) end
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   435
      else ();
70920
1e0ad25c94c8 export constdefs according to defs.ML;
wenzelm
parents: 70919
diff changeset
   436
1e0ad25c94c8 export constdefs according to defs.ML;
wenzelm
parents: 70919
diff changeset
   437
71202
785610ad6bfa export spec rules;
wenzelm
parents: 71179
diff changeset
   438
    (* spec rules *)
785610ad6bfa export spec rules;
wenzelm
parents: 71179
diff changeset
   439
71208
5e0050eb64f2 clarified export of spec rules: more like locale;
wenzelm
parents: 71207
diff changeset
   440
    val encode_specs =
5e0050eb64f2 clarified export of spec rules: more like locale;
wenzelm
parents: 71207
diff changeset
   441
      let open XML.Encode Term_XML.Encode in
5e0050eb64f2 clarified export of spec rules: more like locale;
wenzelm
parents: 71207
diff changeset
   442
        list (fn {props, name, rough_classification, typargs, args, terms, rules} =>
5e0050eb64f2 clarified export of spec rules: more like locale;
wenzelm
parents: 71207
diff changeset
   443
          pair properties (pair string (pair Spec_Rules.encode_rough_classification
5e0050eb64f2 clarified export of spec rules: more like locale;
wenzelm
parents: 71207
diff changeset
   444
            (pair (list (pair string sort)) (pair (list (pair string typ))
71211
7d522732b7f2 more informative export;
wenzelm
parents: 71210
diff changeset
   445
              (pair (list (pair encode_term typ)) (list encode_term))))))
71208
5e0050eb64f2 clarified export of spec rules: more like locale;
wenzelm
parents: 71207
diff changeset
   446
              (props, (name, (rough_classification, (typargs, (args, (terms, rules)))))))
71202
785610ad6bfa export spec rules;
wenzelm
parents: 71179
diff changeset
   447
      end;
785610ad6bfa export spec rules;
wenzelm
parents: 71179
diff changeset
   448
785610ad6bfa export spec rules;
wenzelm
parents: 71179
diff changeset
   449
    val _ =
74114
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   450
      if enabled then
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   451
        (case Spec_Rules.dest_theory thy of
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   452
          [] => ()
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   453
        | spec_rules =>
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   454
            export_body thy "spec_rules" (encode_specs (map spec_rule_content spec_rules)))
700e5bd59c7d clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
wenzelm
parents: 71660
diff changeset
   455
      else ();
68900
1145b25c53de more robust: parents are strict in Export_Theory.read_theory and thus approximate "commit" of exports;
wenzelm
parents: 68864
diff changeset
   456
74261
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
   457
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
   458
    (* other entities *)
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
   459
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
   460
    fun export_other get_space =
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
   461
      let
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
   462
        val space = get_space thy;
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
   463
        val export_name = "other/" ^ Name_Space.kind_of space;
77979
a12c48fbf10f back to more elementary concept of aliases as adhoc change of accesses, but now with "suppress" information (see also 31ea5c1f874d);
wenzelm
parents: 77970
diff changeset
   464
        val decls = Name_Space.get_names space |> map (rpair ());
74261
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
   465
      in export_entities export_name get_space decls (fn _ => fn () => SOME (K [])) end;
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
   466
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
   467
    val other_spaces = other_name_spaces thy;
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
   468
    val other_kinds = map (fn get_space => Name_Space.kind_of (get_space thy)) other_spaces;
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
   469
    val _ =
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
   470
      if null other_kinds then ()
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
   471
      else
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
   472
        Export.export thy \<^path_binding>\<open>theory/other_kinds\<close>
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
   473
          (XML.Encode.string (cat_lines other_kinds));
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
   474
    val _ = List.app export_other other_spaces;
d28a51dd9da6 export other entities, e.g. relevant for formal document output;
wenzelm
parents: 74235
diff changeset
   475
68295
781a98696638 export sort algebra;
wenzelm
parents: 68264
diff changeset
   476
  in () end);
68165
a7a2174ac014 more exports;
wenzelm
parents: 68154
diff changeset
   477
a7a2174ac014 more exports;
wenzelm
parents: 68154
diff changeset
   478
end;