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