src/Pure/Isar/theory_target.ML
author wenzelm
Mon, 09 Oct 2006 02:20:08 +0200
changeset 20915 dcb0a3e2f1bd
parent 20894 784eefc906aa
child 20962 e404275bff33
permissions -rw-r--r--
added exit; notes: simplified locale target;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
20894
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/theory_target.ML
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
     3
    Author:     Makarius
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
     4
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
     5
Common theory targets.
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
     6
*)
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
     7
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
     8
signature THEORY_TARGET =
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
     9
sig
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    10
  val init: xstring option -> theory -> local_theory
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    11
  val init_i: string option -> theory -> local_theory
20915
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
    12
  val exit: local_theory -> Proof.context * theory
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
    13
  val mapping: xstring option -> (local_theory -> local_theory) -> theory -> Proof.context * theory
20894
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    14
end;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    15
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    16
structure TheoryTarget: THEORY_TARGET =
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    17
struct
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    18
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    19
(** locale targets **)
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    20
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    21
(* pretty *)
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    22
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    23
fun pretty loc ctxt =
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    24
  let
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    25
    val thy = ProofContext.theory_of ctxt;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    26
    val fixes = map (fn (x, T) => (x, SOME T, NoSyn)) (#1 (ProofContext.inferred_fixes ctxt));
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    27
    val assumes = map (fn A => (("", []), [(A, [])])) (Assumption.assms_of ctxt);
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    28
    val elems =
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    29
      (if null fixes then [] else [Element.Fixes fixes]) @
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    30
      (if null assumes then [] else [Element.Assumes assumes]);
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    31
  in
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    32
    ([Pretty.block
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    33
      [Pretty.str "theory", Pretty.brk 1, Pretty.str (Context.theory_name thy),
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    34
        Pretty.str " =", Pretty.brk 1, ThyInfo.pretty_theory thy]] @
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    35
      (if loc = "" then []
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    36
       else if null elems then [Pretty.str ("locale " ^ Locale.extern thy loc)]
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    37
       else
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    38
         [Pretty.big_list ("locale " ^ Locale.extern thy loc ^ " =")
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    39
           (map (Pretty.chunks o Element.pretty_ctxt ctxt) elems)]))
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    40
    |> Pretty.chunks
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    41
  end;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    42
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    43
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    44
(* consts *)
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    45
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    46
fun consts loc depends decls lthy =
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    47
  let
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    48
    val xs = filter depends (#1 (ProofContext.inferred_fixes lthy));
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    49
    val ys = filter (Variable.newly_fixed lthy (LocalTheory.target_of lthy) o #1) xs;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    50
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    51
    fun const ((c, T), mx) thy =
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    52
      let
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    53
        val U = map #2 xs ---> T;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    54
        val mx' = if null ys then mx else NoSyn;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    55
        val mx'' = Syntax.unlocalize_mixfix (loc <> "") mx';
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    56
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    57
        val t = Term.list_comb (Const (Sign.full_name thy c, U), map Free xs);
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    58
        val defn = ((c, if loc <> "" then mx else NoSyn (* FIXME !? *)), (("", []), t));
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    59
        val abbr = ((c, mx'), fold_rev (lambda o Free) ys t);
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    60
        val thy' = Sign.add_consts_authentic [(c, U, mx'')] thy;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    61
      in ((defn, abbr), thy') end;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    62
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    63
    val ((defns, abbrs), lthy') = lthy
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    64
      |> LocalTheory.theory_result (fold_map const decls) |>> split_list;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    65
  in
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    66
    lthy'
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    67
    |> K (loc <> "") ? LocalTheory.abbrevs Syntax.default_mode abbrs
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    68
    |> LocalDefs.add_defs defns |>> map (apsnd snd)
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    69
  end;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    70
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    71
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    72
(* axioms *)
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    73
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    74
local
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    75
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    76
fun add_axiom hyps (name, prop) thy =
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    77
  let
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    78
    val name' = if name = "" then "axiom_" ^ serial_string () else name;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    79
    val prop' = Logic.list_implies (hyps, prop);
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    80
    val thy' = thy |> Theory.add_axioms_i [(name', prop')];
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    81
    val axm = Drule.unvarify (Thm.get_axiom_i thy' (Sign.full_name thy' name'));
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    82
    val prems = map (Thm.assume o Thm.cterm_of thy') hyps;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    83
  in (Drule.implies_elim_list axm prems, thy') end;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    84
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    85
in
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    86
20915
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
    87
fun axioms specs lthy =
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
    88
  let
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
    89
    val hyps = Assumption.assms_of lthy;
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
    90
    fun axiom ((name, atts), props) thy = thy
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
    91
      |> fold_map (add_axiom hyps) (PureThy.name_multi name props)
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
    92
      |-> (fn ths => pair ((name, atts), [(ths, [])]));
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
    93
  in
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
    94
    lthy
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
    95
    |> LocalTheory.theory_result (fold_map axiom specs)
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
    96
    |-> LocalTheory.notes
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
    97
  end;
20894
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    98
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
    99
end;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   100
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   101
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   102
(* defs *)
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   103
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   104
local
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   105
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   106
fun add_def (name, prop) =
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   107
  Theory.add_defs_i false false [(name, prop)] #>
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   108
  (fn thy => (Drule.unvarify (Thm.get_axiom_i thy (Sign.full_name thy name)), thy));
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   109
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   110
fun expand_defs lthy =
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   111
  Drule.term_rule (ProofContext.theory_of lthy)
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   112
    (Assumption.export false lthy (LocalTheory.target_of lthy));
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   113
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   114
in
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   115
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   116
fun defs args lthy =
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   117
  let
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   118
    fun def ((x, mx), ((name, atts), rhs)) lthy1 =
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   119
      let
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   120
        val name' = Thm.def_name_optional x name;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   121
        val T = Term.fastype_of rhs;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   122
        val rhs' = expand_defs lthy1 rhs;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   123
        val depends = member (op =) (Term.add_frees rhs' []);
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   124
        val defined = filter_out depends (Term.add_frees rhs []);
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   125
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   126
        val rhs_conv = rhs
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   127
          |> Thm.cterm_of (ProofContext.theory_of lthy1)
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   128
          |> Tactic.rewrite true (map_filter (LocalDefs.find_def lthy1 o #1) defined);
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   129
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   130
        val ([(lhs, local_def)], lthy2) = lthy1
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   131
          |> LocalTheory.consts depends [((x, T), mx)];
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   132
        val lhs' = #2 (Logic.dest_equals (Thm.prop_of local_def));
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   133
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   134
        val (global_def, lthy3) = lthy2
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   135
          |> LocalTheory.theory_result (add_def (name', Logic.mk_equals (lhs', rhs')));
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   136
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   137
        val eq = Thm.transitive (Thm.transitive local_def global_def) (Thm.symmetric rhs_conv);
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   138
      in ((lhs, ((name', atts), [([eq], [])])), lthy3) end;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   139
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   140
    val ((lhss, facts), lthy') = lthy |> fold_map def args |>> split_list;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   141
    val (res, lthy'') = lthy' |> LocalTheory.notes facts;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   142
  in (lhss ~~ map (apsnd the_single) res, lthy'') end;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   143
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   144
end;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   145
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   146
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   147
(* notes *)
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   148
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   149
fun context_notes facts lthy =
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   150
  let
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   151
    val facts' = facts
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   152
      |> Element.export_standard_facts lthy lthy
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   153
      |> Attrib.map_facts (Attrib.attribute_i (ProofContext.theory_of lthy));
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   154
  in
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   155
    lthy
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   156
    |> ProofContext.qualified_names
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   157
    |> ProofContext.note_thmss_i facts'
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   158
    ||> ProofContext.restore_naming lthy
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   159
  end;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   160
20915
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
   161
fun theory_notes keep_atts facts lthy = lthy |> LocalTheory.theory (fn thy =>
20894
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   162
  let
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   163
    val facts' = facts
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   164
      |> Element.export_standard_facts lthy (ProofContext.init thy)
20915
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
   165
      |> Attrib.map_facts (if keep_atts then Attrib.attribute_i thy else K I);
20894
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   166
  in
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   167
    thy
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   168
    |> Sign.qualified_names
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   169
    |> PureThy.note_thmss_i "" facts' |> #2
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   170
    |> Sign.restore_naming thy
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   171
  end);
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   172
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   173
fun locale_notes loc facts lthy = lthy |> LocalTheory.target (fn ctxt =>
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   174
  #2 (Locale.add_thmss "" loc (Element.export_standard_facts lthy ctxt facts) ctxt));
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   175
20915
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
   176
fun notes "" facts = theory_notes true facts #> context_notes facts
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
   177
  | notes loc facts = theory_notes false facts #> locale_notes loc facts #> context_notes facts;
20894
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   178
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   179
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   180
(* declarations *)
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   181
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   182
fun term_syntax "" f = LocalTheory.theory (Context.theory_map f)
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   183
  | term_syntax loc f = LocalTheory.target (Locale.add_term_syntax loc (Context.proof_map f));
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   184
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   185
fun declaration "" f = LocalTheory.theory (Context.theory_map f)
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   186
  | declaration loc f = I;    (* FIXME *)
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   187
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   188
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   189
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   190
(* init and exit *)
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   191
20915
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
   192
fun target_operations loc exit : LocalTheory.operations =
20894
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   193
 {pretty = pretty loc,
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   194
  consts = consts loc,
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   195
  axioms = axioms,
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   196
  defs = defs,
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   197
  notes = notes loc,
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   198
  term_syntax = term_syntax loc,
20915
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
   199
  declaration = declaration loc,
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
   200
  exit = exit};
20894
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   201
20915
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
   202
fun init_i NONE thy = LocalTheory.init (SOME "") (target_operations "" I) (ProofContext.init thy)
20894
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   203
  | init_i (SOME loc) thy =
20915
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
   204
      LocalTheory.init (SOME (NameSpace.base loc)) (target_operations loc I) (Locale.init loc thy);
20894
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   205
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   206
fun init loc thy = init_i (Option.map (Locale.intern thy) loc) thy;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   207
20915
dcb0a3e2f1bd added exit;
wenzelm
parents: 20894
diff changeset
   208
val exit = LocalTheory.exit #> (fn lthy => (lthy, ProofContext.theory_of lthy));
20894
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   209
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   210
fun mapping loc f = init loc #> f #> exit;
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   211
784eefc906aa Common theory targets.
wenzelm
parents:
diff changeset
   212
end;