src/Pure/Build/export_theory.ML
changeset 79502 c7a98469c0e7
parent 79469 deb50d396ff7
child 80074 951c371c1cd9
equal deleted inserted replaced
79501:bce98b5dfec6 79502:c7a98469c0e7
       
     1 (*  Title:      Pure/Build/export_theory.ML
       
     2     Author:     Makarius
       
     3 
       
     4 Export foundational theory content and locale/class structure.
       
     5 *)
       
     6 
       
     7 signature EXPORT_THEORY =
       
     8 sig
       
     9   val other_name_space: (theory -> Name_Space.T) -> theory -> theory
       
    10   val export_enabled: Thy_Info.presentation_context -> bool
       
    11   val export_body: theory -> string -> XML.body -> unit
       
    12 end;
       
    13 
       
    14 structure Export_Theory: EXPORT_THEORY =
       
    15 struct
       
    16 
       
    17 (* other name spaces *)
       
    18 
       
    19 fun err_dup_kind kind = error ("Duplicate name space kind " ^ quote kind);
       
    20 
       
    21 structure Data = Theory_Data
       
    22 (
       
    23   type T = (theory -> Name_Space.T) Inttab.table;
       
    24   val empty = Inttab.empty;
       
    25   val merge = Inttab.merge (K true);
       
    26 );
       
    27 
       
    28 val other_name_spaces = map #2 o Inttab.dest o Data.get;
       
    29 fun other_name_space get_space thy = Data.map (Inttab.update (serial (), get_space)) thy;
       
    30 
       
    31 val _ = Theory.setup
       
    32  (other_name_space Thm.oracle_space #>
       
    33   other_name_space Global_Theory.fact_space #>
       
    34   other_name_space (Bundle.bundle_space o Context.Theory) #>
       
    35   other_name_space (Attrib.attribute_space o Context.Theory) #>
       
    36   other_name_space (Method.method_space o Context.Theory));
       
    37 
       
    38 
       
    39 (* approximative syntax *)
       
    40 
       
    41 val get_syntax = Syntax.get_approx o Proof_Context.syn_of;
       
    42 fun get_syntax_type ctxt = get_syntax ctxt o Lexicon.mark_type;
       
    43 fun get_syntax_const ctxt = get_syntax ctxt o Lexicon.mark_const;
       
    44 fun get_syntax_fixed ctxt = get_syntax ctxt o Lexicon.mark_fixed;
       
    45 
       
    46 fun get_syntax_param ctxt loc x =
       
    47   let val thy = Proof_Context.theory_of ctxt in
       
    48     if Class.is_class thy loc then
       
    49       (case AList.lookup (op =) (Class.these_params thy [loc]) x of
       
    50         NONE => NONE
       
    51       | SOME (_, (c, _)) => get_syntax_const ctxt c)
       
    52     else get_syntax_fixed ctxt x
       
    53   end;
       
    54 
       
    55 val encode_syntax =
       
    56   XML.Encode.variant
       
    57    [fn NONE => ([], []),
       
    58     fn SOME (Syntax.Prefix delim) => ([delim], []),
       
    59     fn SOME (Syntax.Infix {assoc, delim, pri}) =>
       
    60       let
       
    61         val ass =
       
    62           (case assoc of
       
    63             Printer.No_Assoc => 0
       
    64           | Printer.Left_Assoc => 1
       
    65           | Printer.Right_Assoc => 2);
       
    66         open XML.Encode Term_XML.Encode;
       
    67       in ([], triple int string int (ass, delim, pri)) end];
       
    68 
       
    69 
       
    70 (* free variables: not declared in the context *)
       
    71 
       
    72 val is_free = not oo Name.is_declared;
       
    73 
       
    74 fun add_frees used =
       
    75   fold_aterms (fn Free (x, T) => is_free used x ? insert (op =) (x, T) | _ => I);
       
    76 
       
    77 fun add_tfrees used =
       
    78   (fold_types o fold_atyps) (fn TFree (a, S) => is_free used a ? insert (op =) (a, S) | _ => I);
       
    79 
       
    80 
       
    81 (* locales *)
       
    82 
       
    83 fun locale_content thy loc =
       
    84   let
       
    85     val ctxt = Locale.init loc thy;
       
    86     val args =
       
    87       Locale.params_of thy loc
       
    88       |> map (fn ((x, T), _) => ((x, T), get_syntax_param ctxt loc x));
       
    89     val axioms =
       
    90       let
       
    91         val (asm, defs) = Locale.specification_of thy loc;
       
    92         val cprops = map (Thm.cterm_of ctxt) (the_list asm @ defs);
       
    93         val (intro1, intro2) = Locale.intros_of thy loc;
       
    94         val intros_tac = Method.try_intros_tac ctxt (the_list intro1 @ the_list intro2) [];
       
    95         val res =
       
    96           Goal.init (Conjunction.mk_conjunction_balanced cprops)
       
    97           |> (ALLGOALS Goal.conjunction_tac THEN intros_tac)
       
    98           |> try Seq.hd;
       
    99       in
       
   100         (case res of
       
   101           SOME goal => Thm.prems_of goal
       
   102         | NONE => raise Fail ("Cannot unfold locale " ^ quote loc))
       
   103       end;
       
   104     val typargs = build_rev (fold Term.add_tfrees (map (Free o #1) args @ axioms));
       
   105   in {typargs = typargs, args = args, axioms = axioms} end;
       
   106 
       
   107 fun get_locales thy =
       
   108   Locale.get_locales thy |> map_filter (fn loc =>
       
   109     if Experiment.is_experiment thy loc then NONE else SOME (loc, ()));
       
   110 
       
   111 fun get_dependencies prev_thys thy =
       
   112   Locale.dest_dependencies prev_thys thy |> map_filter (fn dep =>
       
   113     if Experiment.is_experiment thy (#source dep) orelse
       
   114       Experiment.is_experiment thy (#target dep) then NONE
       
   115     else
       
   116       let
       
   117         val (type_params, params) = Locale.parameters_of thy (#source dep);
       
   118         val typargs = fold (Term.add_tfreesT o #2 o #1) params type_params;
       
   119         val substT =
       
   120           typargs |> map_filter (fn v =>
       
   121             let
       
   122               val T = TFree v;
       
   123               val T' = Morphism.typ (#morphism dep) T;
       
   124             in if T = T' then NONE else SOME (v, T') end);
       
   125         val subst =
       
   126           params |> map_filter (fn (v, _) =>
       
   127             let
       
   128               val t = Free v;
       
   129               val t' = Morphism.term (#morphism dep) t;
       
   130             in if t aconv t' then NONE else SOME (v, t') end);
       
   131       in SOME (dep, (substT, subst)) end);
       
   132 
       
   133 
       
   134 (* presentation *)
       
   135 
       
   136 fun export_enabled (context: Thy_Info.presentation_context) =
       
   137   Options.bool (#options context) "export_theory";
       
   138 
       
   139 fun export_body thy name body =
       
   140   if XML.is_empty_body body then ()
       
   141   else Export.export thy (Path.binding0 (Path.make ("theory" :: space_explode "/" name))) body;
       
   142 
       
   143 val _ = (Theory.setup o Thy_Info.add_presentation) (fn context => fn thy =>
       
   144   let
       
   145     val rep_tsig = Type.rep_tsig (Sign.tsig_of thy);
       
   146     val consts = Sign.consts_of thy;
       
   147     val thy_ctxt = Proof_Context.init_global thy;
       
   148 
       
   149     val pos_properties = Thy_Info.adjust_pos_properties context;
       
   150 
       
   151     val enabled = export_enabled context;
       
   152 
       
   153 
       
   154     (* strict parents *)
       
   155 
       
   156     val parents = Theory.parents_of thy;
       
   157     val _ =
       
   158       Export.export thy \<^path_binding>\<open>theory/parents\<close>
       
   159         (XML.Encode.string (cat_lines (map Context.theory_long_name parents) ^ "\n"));
       
   160 
       
   161 
       
   162     (* spec rules *)
       
   163 
       
   164     fun spec_rule_content {pos, name, rough_classification, terms, rules} =
       
   165       let
       
   166         val spec =
       
   167           terms @ map Thm.plain_prop_of rules
       
   168           |> Term_Subst.zero_var_indexes_list
       
   169           |> map Logic.unvarify_global;
       
   170       in
       
   171        {props = pos_properties pos,
       
   172         name = name,
       
   173         rough_classification = rough_classification,
       
   174         typargs = build_rev (fold Term.add_tfrees spec),
       
   175         args = build_rev (fold Term.add_frees spec),
       
   176         terms = map (fn t => (t, Term.type_of t)) (take (length terms) spec),
       
   177         rules = drop (length terms) spec}
       
   178       end;
       
   179 
       
   180 
       
   181     (* entities *)
       
   182 
       
   183     fun make_entity_markup name xname pos serial =
       
   184       let val props = pos_properties pos @ Markup.serial_properties serial;
       
   185       in (Markup.entityN, (Markup.nameN, name) :: (Markup.xnameN, xname) :: props) end;
       
   186 
       
   187     fun entity_markup space name =
       
   188       let
       
   189         val xname = Name_Space.extern_shortest thy_ctxt space name;
       
   190         val {serial, pos, ...} = Name_Space.the_entry space name;
       
   191       in make_entity_markup name xname pos serial end;
       
   192 
       
   193     fun export_entities export_name get_space decls export =
       
   194       let
       
   195         val parent_spaces = map get_space parents;
       
   196         val space = get_space thy;
       
   197       in
       
   198         build (decls |> fold (fn (name, decl) =>
       
   199           if exists (fn space => Name_Space.declared space name) parent_spaces then I
       
   200           else
       
   201             (case export name decl of
       
   202               NONE => I
       
   203             | SOME make_body =>
       
   204                 let
       
   205                   val i = #serial (Name_Space.the_entry space name);
       
   206                   val body = if enabled then make_body () else [];
       
   207                 in cons (i, XML.Elem (entity_markup space name, body)) end)))
       
   208         |> sort (int_ord o apply2 #1) |> map #2
       
   209         |> export_body thy export_name
       
   210       end;
       
   211 
       
   212 
       
   213     (* types *)
       
   214 
       
   215     val encode_type =
       
   216       let open XML.Encode Term_XML.Encode
       
   217       in triple encode_syntax (list string) (option typ) end;
       
   218 
       
   219     val _ =
       
   220       export_entities "types" Sign.type_space (Name_Space.dest_table (#types rep_tsig))
       
   221         (fn c =>
       
   222           (fn Type.Logical_Type n =>
       
   223                 SOME (fn () =>
       
   224                   encode_type (get_syntax_type thy_ctxt c, Name.invent Name.context Name.aT n, NONE))
       
   225             | Type.Abbreviation (args, U, false) =>
       
   226                 SOME (fn () =>
       
   227                   encode_type (get_syntax_type thy_ctxt c, args, SOME U))
       
   228             | _ => NONE));
       
   229 
       
   230 
       
   231     (* consts *)
       
   232 
       
   233     val encode_term = Term_XML.Encode.term consts;
       
   234 
       
   235     val encode_const =
       
   236       let open XML.Encode Term_XML.Encode
       
   237       in pair encode_syntax (pair (list string) (pair typ (pair (option encode_term) bool))) end;
       
   238 
       
   239     val _ =
       
   240       export_entities "consts" Sign.const_space (#constants (Consts.dest consts))
       
   241         (fn c => fn (T, abbrev) =>
       
   242           SOME (fn () =>
       
   243             let
       
   244               val syntax = get_syntax_const thy_ctxt c;
       
   245               val U = Logic.unvarifyT_global T;
       
   246               val U0 = Term.strip_sortsT U;
       
   247               val trim_abbrev = Proofterm.standard_vars_term Name.context #> Term.strip_sorts;
       
   248               val abbrev' = Option.map trim_abbrev abbrev;
       
   249               val args = map (#1 o dest_TFree) (Consts.typargs consts (c, U0));
       
   250               val propositional = Object_Logic.is_propositional thy_ctxt (Term.body_type U0);
       
   251             in encode_const (syntax, (args, (U0, (abbrev', propositional)))) end));
       
   252 
       
   253 
       
   254     (* axioms *)
       
   255 
       
   256     fun standard_prop used extra_shyps raw_prop raw_proof =
       
   257       let
       
   258         val (prop, proof) = Proofterm.standard_vars used (raw_prop, raw_proof);
       
   259         val args = rev (add_frees used prop []);
       
   260         val typargs = rev (add_tfrees used prop []);
       
   261         val used_typargs = fold (Name.declare o #1) typargs used;
       
   262         val sorts = Name.invent used_typargs Name.aT (length extra_shyps) ~~ extra_shyps;
       
   263       in ((sorts @ typargs, args, prop), proof) end;
       
   264 
       
   265     fun standard_prop_of thm =
       
   266       standard_prop Name.context (Thm.extra_shyps thm) (Thm.full_prop_of thm);
       
   267 
       
   268     val encode_prop =
       
   269       let open XML.Encode Term_XML.Encode
       
   270       in triple (list (pair string sort)) (list (pair string typ)) encode_term end;
       
   271 
       
   272     fun encode_axiom used prop =
       
   273       encode_prop (#1 (standard_prop used [] prop NONE));
       
   274 
       
   275     val _ =
       
   276       export_entities "axioms" Theory.axiom_space (Theory.all_axioms_of thy)
       
   277         (fn _ => fn prop => SOME (fn () => encode_axiom Name.context prop));
       
   278 
       
   279 
       
   280     (* theorems and proof terms *)
       
   281 
       
   282     val clean_thm = Thm.check_hyps (Context.Theory thy) #> Thm.strip_shyps;
       
   283     val prep_thm = clean_thm #> Thm.unconstrainT #> Thm.strip_shyps;
       
   284 
       
   285     val lookup_thm_id = Global_Theory.lookup_thm_id thy;
       
   286 
       
   287     fun expand_name thm_id (header: Proofterm.thm_header) =
       
   288       if #serial header = #serial thm_id then ""
       
   289       else
       
   290         (case lookup_thm_id (Proofterm.thm_header_id header) of
       
   291           NONE => ""
       
   292         | SOME thm_name => Thm_Name.print thm_name);
       
   293 
       
   294     fun entity_markup_thm (serial, (name, i)) =
       
   295       let
       
   296         val space = Global_Theory.fact_space thy;
       
   297         val xname = Name_Space.extern_shortest thy_ctxt space name;
       
   298         val {pos, ...} = Name_Space.the_entry space name;
       
   299       in make_entity_markup (Thm_Name.print (name, i)) (Thm_Name.print (xname, i)) pos serial end;
       
   300 
       
   301     fun encode_thm thm_id raw_thm =
       
   302       let
       
   303         val deps = map (Thm_Name.print o #2) (Thm_Deps.thm_deps thy [raw_thm]);
       
   304         val thm = prep_thm raw_thm;
       
   305 
       
   306         val proof0 =
       
   307           if Proofterm.export_standard_enabled () then
       
   308             Proof_Syntax.standard_proof_of
       
   309               {full = true, expand_name = SOME o expand_name thm_id} thm
       
   310           else if Proofterm.export_enabled () then Thm.reconstruct_proof_of thm
       
   311           else MinProof;
       
   312         val (prop, SOME proof) = standard_prop_of thm (SOME proof0);
       
   313         val _ = Thm.expose_proofs thy [thm];
       
   314       in
       
   315         (prop, deps, proof) |>
       
   316           let
       
   317             open XML.Encode Term_XML.Encode;
       
   318             val encode_proof = Proofterm.encode_standard_proof consts;
       
   319           in triple encode_prop (list string) encode_proof end
       
   320       end;
       
   321 
       
   322     fun export_thm (thm_id, thm_name) =
       
   323       let
       
   324         val markup = entity_markup_thm (#serial thm_id, thm_name);
       
   325         val body =
       
   326           if enabled then
       
   327             Global_Theory.get_thm_name thy (thm_name, Position.none)
       
   328             |> encode_thm thm_id
       
   329           else [];
       
   330       in XML.Elem (markup, body) end;
       
   331 
       
   332     val _ = export_body thy "thms" (map export_thm (Global_Theory.dest_thm_names thy));
       
   333 
       
   334 
       
   335     (* type classes *)
       
   336 
       
   337     val encode_class =
       
   338       let open XML.Encode Term_XML.Encode
       
   339       in pair (list (pair string typ)) (list (encode_axiom Name.context)) end;
       
   340 
       
   341     val _ =
       
   342       export_entities "classes" Sign.class_space
       
   343         (map (rpair ()) (Graph.keys (Sorts.classes_of (#2 (#classes rep_tsig)))))
       
   344         (fn name => fn () => SOME (fn () =>
       
   345           (case try (Axclass.get_info thy) name of
       
   346             NONE => ([], [])
       
   347           | SOME {params, axioms, ...} => (params, map (Thm.plain_prop_of o clean_thm) axioms))
       
   348           |> encode_class));
       
   349 
       
   350 
       
   351     (* sort algebra *)
       
   352 
       
   353     val _ =
       
   354       if enabled then
       
   355         let
       
   356           val prop = encode_axiom Name.context o Logic.varify_global;
       
   357 
       
   358           val encode_classrel =
       
   359             let open XML.Encode
       
   360             in list (pair prop (pair string string)) end;
       
   361 
       
   362           val encode_arities =
       
   363             let open XML.Encode Term_XML.Encode
       
   364             in list (pair prop (triple string (list sort) string)) end;
       
   365 
       
   366           val export_classrel =
       
   367             maps (fn (c, cs) => map (pair c) cs) #> map (`Logic.mk_classrel) #> encode_classrel;
       
   368 
       
   369           val export_arities = map (`Logic.mk_arity) #> encode_arities;
       
   370 
       
   371           val {classrel, arities} =
       
   372             Sorts.dest_algebra (map (#2 o #classes o Type.rep_tsig o Sign.tsig_of) parents)
       
   373               (#2 (#classes rep_tsig));
       
   374         in
       
   375           if null classrel then () else export_body thy "classrel" (export_classrel classrel);
       
   376           if null arities then () else export_body thy "arities" (export_arities arities)
       
   377         end
       
   378       else ();
       
   379 
       
   380 
       
   381     (* locales *)
       
   382 
       
   383     fun encode_locale used =
       
   384       let open XML.Encode Term_XML.Encode in
       
   385         triple (list (pair string sort)) (list (pair (pair string typ) encode_syntax))
       
   386           (list (encode_axiom used))
       
   387       end;
       
   388 
       
   389     val _ =
       
   390       export_entities "locales" Locale.locale_space (get_locales thy)
       
   391         (fn loc => fn () => SOME (fn () =>
       
   392           let
       
   393             val {typargs, args, axioms} = locale_content thy loc;
       
   394             val used = fold Name.declare (map #1 typargs @ map (#1 o #1) args) Name.context;
       
   395           in encode_locale used (typargs, args, axioms) end
       
   396           handle ERROR msg =>
       
   397             cat_error msg ("The error(s) above occurred in locale " ^
       
   398               quote (Locale.markup_name thy_ctxt loc))));
       
   399 
       
   400 
       
   401     (* locale dependencies *)
       
   402 
       
   403     fun encode_locale_dependency (dep: Locale.locale_dependency, subst) =
       
   404       (#source dep, (#target dep, (#prefix dep, subst))) |>
       
   405         let
       
   406           open XML.Encode Term_XML.Encode;
       
   407           val encode_subst =
       
   408             pair (list (pair (pair string sort) typ)) (list (pair (pair string typ) (term consts)));
       
   409         in pair string (pair string (pair (list (pair string bool)) encode_subst)) end;
       
   410 
       
   411     val _ =
       
   412       if enabled then
       
   413         get_dependencies parents thy |> map_index (fn (i, dep) =>
       
   414           let
       
   415             val xname = string_of_int (i + 1);
       
   416             val name = Long_Name.implode [Context.theory_base_name thy, xname];
       
   417             val markup = make_entity_markup name xname (#pos (#1 dep)) (#serial (#1 dep));
       
   418             val body = encode_locale_dependency dep;
       
   419           in XML.Elem (markup, body) end)
       
   420         |> export_body thy "locale_dependencies"
       
   421       else ();
       
   422 
       
   423 
       
   424     (* constdefs *)
       
   425 
       
   426     val _ =
       
   427       if enabled then
       
   428         let
       
   429           val constdefs =
       
   430             Defs.dest_constdefs (map Theory.defs_of (Theory.parents_of thy)) (Theory.defs_of thy)
       
   431             |> sort_by #1;
       
   432           val encode =
       
   433             let open XML.Encode
       
   434             in list (pair string string) end;
       
   435         in if null constdefs then () else export_body thy "constdefs" (encode constdefs) end
       
   436       else ();
       
   437 
       
   438 
       
   439     (* spec rules *)
       
   440 
       
   441     val encode_specs =
       
   442       let open XML.Encode Term_XML.Encode in
       
   443         list (fn {props, name, rough_classification, typargs, args, terms, rules} =>
       
   444           pair properties (pair string (pair Spec_Rules.encode_rough_classification
       
   445             (pair (list (pair string sort)) (pair (list (pair string typ))
       
   446               (pair (list (pair encode_term typ)) (list encode_term))))))
       
   447               (props, (name, (rough_classification, (typargs, (args, (terms, rules)))))))
       
   448       end;
       
   449 
       
   450     val _ =
       
   451       if enabled then
       
   452         (case Spec_Rules.dest_theory thy of
       
   453           [] => ()
       
   454         | spec_rules =>
       
   455             export_body thy "spec_rules" (encode_specs (map spec_rule_content spec_rules)))
       
   456       else ();
       
   457 
       
   458 
       
   459     (* other entities *)
       
   460 
       
   461     fun export_other get_space =
       
   462       let
       
   463         val space = get_space thy;
       
   464         val export_name = "other/" ^ Name_Space.kind_of space;
       
   465         val decls = Name_Space.get_names space |> map (rpair ());
       
   466       in export_entities export_name get_space decls (fn _ => fn () => SOME (K [])) end;
       
   467 
       
   468     val other_spaces = other_name_spaces thy;
       
   469     val other_kinds = map (fn get_space => Name_Space.kind_of (get_space thy)) other_spaces;
       
   470     val _ =
       
   471       if null other_kinds then ()
       
   472       else
       
   473         Export.export thy \<^path_binding>\<open>theory/other_kinds\<close>
       
   474           (XML.Encode.string (cat_lines other_kinds));
       
   475     val _ = List.app export_other other_spaces;
       
   476 
       
   477   in () end);
       
   478 
       
   479 end;