src/Pure/Isar/local_theory.ML
author wenzelm
Sat, 14 Oct 2006 23:25:53 +0200
changeset 21034 99697a1597b2
parent 20981 c4d3fc6e1e64
child 21206 2af4c7b3f7ef
permissions -rw-r--r--
added assert;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/local_theory.ML
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
     3
    Author:     Makarius
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
     4
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
     5
Local theory operations, with abstract target context.
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
     6
*)
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
     7
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
     8
type local_theory = Proof.context;
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
     9
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    10
signature LOCAL_THEORY =
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    11
sig
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    12
  type operations
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    13
  val target_of: local_theory -> Proof.context
20960
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
    14
  val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
    15
  val raw_theory: (theory -> theory) -> local_theory -> local_theory
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
    16
  val theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    17
  val theory: (theory -> theory) -> local_theory -> local_theory
20960
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
    18
  val target_result: (Proof.context -> 'a * Proof.context) -> local_theory -> 'a * local_theory
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    19
  val target: (Proof.context -> Proof.context) -> local_theory -> local_theory
21034
99697a1597b2 added assert;
wenzelm
parents: 20981
diff changeset
    20
  val assert: local_theory -> local_theory
20960
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
    21
  val pretty: local_theory -> Pretty.T list
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    22
  val consts: (string * typ -> bool) ->
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    23
    ((bstring * typ) * mixfix) list -> local_theory -> (term * thm) list * local_theory
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    24
  val axioms: ((bstring * Attrib.src list) * term list) list -> local_theory ->
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    25
    (bstring * thm list) list * local_theory
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    26
  val defs: ((bstring * mixfix) * ((bstring * Attrib.src list) * term)) list ->
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    27
    local_theory -> (term * (bstring * thm)) list * local_theory
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    28
  val notes: ((bstring * Attrib.src list) * (thm list * Attrib.src list) list) list ->
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    29
    local_theory -> (bstring * thm list) list * local_theory
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    30
  val term_syntax: (Context.generic -> Context.generic) -> local_theory -> local_theory
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    31
  val declaration: (Context.generic -> Context.generic) -> local_theory -> local_theory
18823
916c493b7f0c added print_consts;
wenzelm
parents: 18805
diff changeset
    32
  val def: (bstring * mixfix) * ((bstring * Attrib.src list) * term) ->
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    33
    local_theory -> (term * (bstring * thm)) * local_theory
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    34
  val note: (bstring * Attrib.src list) * thm list ->
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    35
    local_theory -> (bstring * thm list) * Proof.context
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    36
  val const_syntax: Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    37
  val abbrevs: Syntax.mode -> ((bstring * mixfix) * term) list -> local_theory -> local_theory
20910
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
    38
  val init: string option -> operations -> Proof.context -> local_theory
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
    39
  val reinit: local_theory -> local_theory
20981
c4d3fc6e1e64 exit: pass interactive flag, toplevel result convention;
wenzelm
parents: 20960
diff changeset
    40
  val exit: bool -> local_theory -> Proof.context * theory
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    41
end;
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    42
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    43
structure LocalTheory: LOCAL_THEORY =
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    44
struct
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    45
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    46
(** local theory data **)
19059
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
    47
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    48
(* type lthy *)
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    49
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    50
type operations =
20960
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
    51
 {pretty: local_theory -> Pretty.T list,
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    52
  consts: (string * typ -> bool) -> ((bstring * typ) * mixfix) list -> local_theory ->
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    53
    (term * thm) list * local_theory,
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    54
  axioms: ((bstring * Attrib.src list) * term list) list -> local_theory ->
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    55
    (bstring * thm list) list * local_theory,
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    56
  defs: ((bstring * mixfix) * ((bstring * Attrib.src list) * term)) list -> local_theory ->
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    57
    (term * (bstring * thm)) list * local_theory,
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    58
  notes: ((bstring * Attrib.src list) * (thm list * Attrib.src list) list) list -> local_theory ->
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    59
    (bstring * thm list) list * local_theory,
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    60
  term_syntax: (Context.generic -> Context.generic) -> local_theory -> local_theory,
20910
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
    61
  declaration: (Context.generic -> Context.generic) -> local_theory -> local_theory,
20981
c4d3fc6e1e64 exit: pass interactive flag, toplevel result convention;
wenzelm
parents: 20960
diff changeset
    62
  exit: bool -> local_theory -> local_theory};
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    63
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    64
datatype lthy = LThy of
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    65
 {theory_prefix: string option,
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    66
  operations: operations,
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    67
  target: Proof.context};
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    68
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    69
fun make_lthy (theory_prefix, operations, target) =
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    70
  LThy {theory_prefix = theory_prefix, operations = operations, target = target};
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    71
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    72
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    73
(* context data *)
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    74
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    75
structure Data = ProofDataFun
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    76
(
18876
ddb6803da197 added consts_retricted;
wenzelm
parents: 18823
diff changeset
    77
  val name = "Pure/local_theory";
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    78
  type T = lthy option;
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    79
  fun init _ = NONE;
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    80
  fun print _ _ = ();
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    81
);
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    82
val _ = Context.add_setup Data.init;
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    83
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    84
fun get_lthy lthy =
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    85
  (case Data.get lthy of
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    86
    NONE => error "No local theory context"
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    87
  | SOME (LThy data) => data);
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    88
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    89
fun map_lthy f lthy =
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    90
  let val {theory_prefix, operations, target} = get_lthy lthy
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    91
  in Data.put (SOME (make_lthy (f (theory_prefix, operations, target)))) lthy end;
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    92
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    93
val target_of = #target o get_lthy;
21034
99697a1597b2 added assert;
wenzelm
parents: 20981
diff changeset
    94
val assert = tap target_of;
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    95
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    96
fun map_target f = map_lthy (fn (theory_prefix, operations, target) =>
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    97
  (theory_prefix, operations, f target));
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    98
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    99
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   100
(* substructure mappings *)
19661
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   101
20960
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   102
fun raw_theory_result f lthy =
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   103
  let
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   104
    val (res, thy') = f (ProofContext.theory_of lthy);
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   105
    val lthy' = lthy
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   106
      |> map_target (ProofContext.transfer thy')
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   107
      |> ProofContext.transfer thy';
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   108
  in (res, lthy') end;
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   109
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   110
fun raw_theory f = #2 o raw_theory_result (f #> pair ());
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   111
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   112
fun theory_result f lthy =
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   113
  (case #theory_prefix (get_lthy lthy) of
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   114
    NONE => error "Cannot change background theory"
20960
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   115
  | SOME prefix => lthy |> raw_theory_result (fn thy => thy
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   116
      |> Sign.sticky_prefix prefix
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   117
      |> Sign.qualified_names
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   118
      |> f
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   119
      ||> Sign.restore_naming thy));
19661
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   120
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   121
fun theory f = #2 o theory_result (f #> pair ());
19661
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   122
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   123
fun target_result f lthy =
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   124
  let
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   125
    val (res, ctxt') = f (#target (get_lthy lthy));
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   126
    val thy' = ProofContext.theory_of ctxt';
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   127
    val lthy' = lthy
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   128
      |> map_target (K ctxt')
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   129
      |> ProofContext.transfer thy';
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   130
  in (res, lthy') end;
18876
ddb6803da197 added consts_retricted;
wenzelm
parents: 18823
diff changeset
   131
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   132
fun target f = #2 o target_result (f #> pair ());
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   133
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   134
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   135
(* primitive operations *)
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   136
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   137
fun operation f lthy = f (#operations (get_lthy lthy)) lthy;
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   138
fun operation1 f x = operation (fn ops => f ops x);
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   139
fun operation2 f x y = operation (fn ops => f ops x y);
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   140
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   141
val pretty = operation #pretty;
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   142
val consts = operation2 #consts;
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   143
val axioms = operation1 #axioms;
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   144
val defs = operation1 #defs;
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   145
val notes = operation1 #notes;
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   146
val term_syntax = operation1 #term_syntax;
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   147
val declaration = operation1 #declaration;
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   148
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   149
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   150
(* derived operations *)
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   151
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   152
fun def arg lthy = lthy |> defs [arg] |>> hd;
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   153
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   154
fun note (a, ths) lthy = lthy |> notes [(a, [(ths, [])])] |>> hd;
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   155
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   156
fun const_syntax mode args =
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   157
  term_syntax
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   158
    (Context.mapping (Sign.add_const_syntax mode args) (ProofContext.add_const_syntax mode args));
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   159
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   160
fun abbrevs mode args =
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   161
  term_syntax (Context.mapping (Sign.add_abbrevs mode args) (ProofContext.add_abbrevs mode args));
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   162
20910
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   163
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   164
(* init -- exit *)
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   165
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   166
fun init theory_prefix operations target = target |> Data.map
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   167
  (fn NONE => SOME (make_lthy (theory_prefix, operations, target))
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   168
    | SOME _ => error "Local theory already initialized");
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   169
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   170
fun reinit lthy =
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   171
  let val {theory_prefix, operations, target} = get_lthy lthy
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   172
  in init theory_prefix operations target end;
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   173
20981
c4d3fc6e1e64 exit: pass interactive flag, toplevel result convention;
wenzelm
parents: 20960
diff changeset
   174
fun exit int lthy = lthy
c4d3fc6e1e64 exit: pass interactive flag, toplevel result convention;
wenzelm
parents: 20960
diff changeset
   175
  |> operation1 #exit int |> target_of
c4d3fc6e1e64 exit: pass interactive flag, toplevel result convention;
wenzelm
parents: 20960
diff changeset
   176
  |> (fn ctxt => (ctxt, ProofContext.theory_of ctxt));
20910
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   177
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   178
end;