src/Pure/Isar/local_theory.ML
author haftmann
Fri, 15 Feb 2008 16:09:12 +0100
changeset 26072 f65a7fa2da6c
parent 25984 da0399c9ffcb
child 26131 91024979b9eb
permissions -rw-r--r--
<= and < on nat no longer depend on wellfounded relations
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
22240
36cc1875619f added full_naming;
wenzelm
parents: 22203
diff changeset
    16
  val full_naming: local_theory -> NameSpace.naming
21614
89105c15b436 added full_name;
wenzelm
parents: 21529
diff changeset
    17
  val full_name: local_theory -> bstring -> string
20960
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
    18
  val theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    19
  val theory: (theory -> theory) -> local_theory -> local_theory
20960
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
    20
  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
    21
  val target: (Proof.context -> Proof.context) -> local_theory -> local_theory
21860
c4492c6bf450 renamed LocalTheory.assert to affirm;
wenzelm
parents: 21802
diff changeset
    22
  val affirm: local_theory -> local_theory
20960
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
    23
  val pretty: local_theory -> Pretty.T list
24984
952045a8dcf2 local_theory: incorporated consts into axioms;
wenzelm
parents: 24949
diff changeset
    24
  val axioms: string ->
952045a8dcf2 local_theory: incorporated consts into axioms;
wenzelm
parents: 24949
diff changeset
    25
    ((bstring * typ) * mixfix) list * ((bstring * Attrib.src list) * term list) list -> local_theory
952045a8dcf2 local_theory: incorporated consts into axioms;
wenzelm
parents: 24949
diff changeset
    26
    -> (term list * (bstring * thm list) list) * local_theory
25984
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
    27
  val axioms_grouped: string -> string ->
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
    28
    ((bstring * typ) * mixfix) list * ((bstring * Attrib.src list) * term list) list -> local_theory
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
    29
    -> (term list * (bstring * thm list) list) * local_theory
21802
e024aa65f4ce abbrev: tuned signature;
wenzelm
parents: 21743
diff changeset
    30
  val abbrev: Syntax.mode -> (bstring * mixfix) * term -> local_theory ->
25120
23fbc38f6432 tuned abbrev interface;
wenzelm
parents: 25104
diff changeset
    31
    (term * term) * local_theory
25021
8f00edb993bd renamed def to define;
wenzelm
parents: 24984
diff changeset
    32
  val define: string -> (bstring * mixfix) * ((bstring * Attrib.src list) * term) -> local_theory ->
8f00edb993bd renamed def to define;
wenzelm
parents: 24984
diff changeset
    33
    (term * (bstring * thm)) * local_theory
25984
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
    34
  val define_grouped: string -> string ->
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
    35
    (bstring * mixfix) * ((bstring * Attrib.src list) * term) -> local_theory ->
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
    36
    (term * (bstring * thm)) * local_theory
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
    37
  val note: string -> (bstring * Attrib.src list) * thm list ->
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
    38
    local_theory -> (bstring * thm list) * local_theory
21433
89104dca628e LocalTheory.axioms/notes/defs: proper kind;
wenzelm
parents: 21292
diff changeset
    39
  val notes: string -> ((bstring * Attrib.src list) * (thm list * Attrib.src list) list) list ->
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    40
    local_theory -> (bstring * thm list) list * local_theory
25984
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
    41
  val note_grouped: string -> string ->
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
    42
    (bstring * Attrib.src list) * thm list ->
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
    43
    local_theory -> (bstring * thm list) * local_theory
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
    44
  val notes_grouped: string -> string ->
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
    45
    ((bstring * Attrib.src list) * (thm list * Attrib.src list) list) list ->
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
    46
    local_theory -> (bstring * thm list) list * local_theory
24020
ed4d7abffee7 tuned signature;
wenzelm
parents: 22846
diff changeset
    47
  val type_syntax: declaration -> local_theory -> local_theory
ed4d7abffee7 tuned signature;
wenzelm
parents: 22846
diff changeset
    48
  val term_syntax: declaration -> local_theory -> local_theory
ed4d7abffee7 tuned signature;
wenzelm
parents: 22846
diff changeset
    49
  val declaration: declaration -> local_theory -> local_theory
24949
5f00e3532418 generalized notation interface (add or del);
wenzelm
parents: 24933
diff changeset
    50
  val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory
21529
bfe99f603933 abbrevs: no result;
wenzelm
parents: 21498
diff changeset
    51
  val target_morphism: local_theory -> morphism
bfe99f603933 abbrevs: no result;
wenzelm
parents: 21498
diff changeset
    52
  val init: string -> operations -> Proof.context -> local_theory
21273
3b01db9504a8 abbrevs: return result;
wenzelm
parents: 21206
diff changeset
    53
  val restore: local_theory -> local_theory
25289
3d332d8a827c simplified LocalTheory.reinit;
wenzelm
parents: 25120
diff changeset
    54
  val reinit: local_theory -> local_theory
21292
11143b6ad87f simplified exit;
wenzelm
parents: 21273
diff changeset
    55
  val exit: local_theory -> Proof.context
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    56
end;
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    57
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    58
structure LocalTheory: LOCAL_THEORY =
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    59
struct
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    60
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    61
(** local theory data **)
19059
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
    62
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    63
(* type lthy *)
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    64
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    65
type operations =
20960
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
    66
 {pretty: local_theory -> Pretty.T list,
25984
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
    67
  axioms: string -> string ->
24984
952045a8dcf2 local_theory: incorporated consts into axioms;
wenzelm
parents: 24949
diff changeset
    68
    ((bstring * typ) * mixfix) list * ((bstring * Attrib.src list) * term list) list -> local_theory
952045a8dcf2 local_theory: incorporated consts into axioms;
wenzelm
parents: 24949
diff changeset
    69
    -> (term list * (bstring * thm list) list) * local_theory,
25120
23fbc38f6432 tuned abbrev interface;
wenzelm
parents: 25104
diff changeset
    70
  abbrev: Syntax.mode -> (bstring * mixfix) * term -> local_theory -> (term * term) * local_theory,
25984
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
    71
  define: string -> string ->
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
    72
    (bstring * mixfix) * ((bstring * Attrib.src list) * term) -> local_theory ->
25021
8f00edb993bd renamed def to define;
wenzelm
parents: 24984
diff changeset
    73
    (term * (bstring * thm)) * local_theory,
25984
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
    74
  notes: string -> string ->
21433
89104dca628e LocalTheory.axioms/notes/defs: proper kind;
wenzelm
parents: 21292
diff changeset
    75
    ((bstring * Attrib.src list) * (thm list * Attrib.src list) list) list ->
89104dca628e LocalTheory.axioms/notes/defs: proper kind;
wenzelm
parents: 21292
diff changeset
    76
    local_theory -> (bstring * thm list) list * local_theory,
24020
ed4d7abffee7 tuned signature;
wenzelm
parents: 22846
diff changeset
    77
  type_syntax: declaration -> local_theory -> local_theory,
ed4d7abffee7 tuned signature;
wenzelm
parents: 22846
diff changeset
    78
  term_syntax: declaration -> local_theory -> local_theory,
ed4d7abffee7 tuned signature;
wenzelm
parents: 22846
diff changeset
    79
  declaration: declaration -> local_theory -> local_theory,
25289
3d332d8a827c simplified LocalTheory.reinit;
wenzelm
parents: 25120
diff changeset
    80
  reinit: local_theory -> local_theory,
21292
11143b6ad87f simplified exit;
wenzelm
parents: 21273
diff changeset
    81
  exit: local_theory -> Proof.context};
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    82
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    83
datatype lthy = LThy of
21529
bfe99f603933 abbrevs: no result;
wenzelm
parents: 21498
diff changeset
    84
 {theory_prefix: string,
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    85
  operations: operations,
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    86
  target: Proof.context};
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    87
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    88
fun make_lthy (theory_prefix, operations, target) =
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    89
  LThy {theory_prefix = theory_prefix, operations = operations, target = target};
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    90
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    91
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    92
(* context data *)
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    93
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    94
structure Data = ProofDataFun
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    95
(
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    96
  type T = lthy option;
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    97
  fun init _ = NONE;
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    98
);
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    99
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   100
fun get_lthy lthy =
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   101
  (case Data.get lthy of
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   102
    NONE => error "No local theory context"
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   103
  | SOME (LThy data) => data);
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
   104
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   105
fun map_lthy f lthy =
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   106
  let val {theory_prefix, operations, target} = get_lthy lthy
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   107
  in Data.put (SOME (make_lthy (f (theory_prefix, operations, target)))) lthy end;
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   108
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   109
val target_of = #target o get_lthy;
21860
c4492c6bf450 renamed LocalTheory.assert to affirm;
wenzelm
parents: 21802
diff changeset
   110
val affirm = tap target_of;
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   111
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   112
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
   113
  (theory_prefix, operations, f target));
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   114
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   115
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   116
(* substructure mappings *)
19661
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   117
20960
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   118
fun raw_theory_result f lthy =
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   119
  let
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   120
    val (res, thy') = f (ProofContext.theory_of lthy);
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   121
    val lthy' = lthy
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   122
      |> map_target (ProofContext.transfer thy')
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   123
      |> ProofContext.transfer thy';
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   124
  in (res, lthy') end;
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   125
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   126
fun raw_theory f = #2 o raw_theory_result (f #> pair ());
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   127
22240
36cc1875619f added full_naming;
wenzelm
parents: 22203
diff changeset
   128
fun full_naming lthy =
36cc1875619f added full_naming;
wenzelm
parents: 22203
diff changeset
   129
  Sign.naming_of (ProofContext.theory_of lthy)
36cc1875619f added full_naming;
wenzelm
parents: 22203
diff changeset
   130
  |> NameSpace.sticky_prefix (#theory_prefix (get_lthy lthy))
36cc1875619f added full_naming;
wenzelm
parents: 22203
diff changeset
   131
  |> NameSpace.qualified_names;
36cc1875619f added full_naming;
wenzelm
parents: 22203
diff changeset
   132
36cc1875619f added full_naming;
wenzelm
parents: 22203
diff changeset
   133
val full_name = NameSpace.full o full_naming;
21614
89105c15b436 added full_name;
wenzelm
parents: 21529
diff changeset
   134
21529
bfe99f603933 abbrevs: no result;
wenzelm
parents: 21498
diff changeset
   135
fun theory_result f lthy = lthy |> raw_theory_result (fn thy => thy
bfe99f603933 abbrevs: no result;
wenzelm
parents: 21498
diff changeset
   136
  |> Sign.sticky_prefix (#theory_prefix (get_lthy lthy))
bfe99f603933 abbrevs: no result;
wenzelm
parents: 21498
diff changeset
   137
  |> Sign.qualified_names
bfe99f603933 abbrevs: no result;
wenzelm
parents: 21498
diff changeset
   138
  |> f
bfe99f603933 abbrevs: no result;
wenzelm
parents: 21498
diff changeset
   139
  ||> Sign.restore_naming thy);
19661
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   140
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   141
fun theory f = #2 o theory_result (f #> pair ());
19661
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   142
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   143
fun target_result f lthy =
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   144
  let
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   145
    val (res, ctxt') = f (#target (get_lthy lthy));
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   146
    val thy' = ProofContext.theory_of ctxt';
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   147
    val lthy' = lthy
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   148
      |> map_target (K ctxt')
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   149
      |> ProofContext.transfer thy';
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   150
  in (res, lthy') end;
18876
ddb6803da197 added consts_retricted;
wenzelm
parents: 18823
diff changeset
   151
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   152
fun target f = #2 o target_result (f #> pair ());
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   153
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   154
21664
dd4e89123359 notation/abbreviation: more serious handling of morphisms;
wenzelm
parents: 21614
diff changeset
   155
(* basic operations *)
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   156
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   157
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
   158
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
   159
fun operation2 f x y = operation (fn ops => f ops x y);
25984
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
   160
fun operation3 f x y z = operation (fn ops => f ops x y z);
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   161
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   162
val pretty = operation #pretty;
25984
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
   163
val axioms_grouped = operation3 #axioms;
21743
fe94c6797c09 added notation/abbrev (from term_syntax.ML);
wenzelm
parents: 21700
diff changeset
   164
val abbrev = operation2 #abbrev;
25984
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
   165
val define_grouped = operation3 #define;
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
   166
val notes_grouped = operation3 #notes;
21498
6382c3a1e7cf uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
wenzelm
parents: 21433
diff changeset
   167
val type_syntax = operation1 #type_syntax;
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   168
val term_syntax = operation1 #term_syntax;
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   169
val declaration = operation1 #declaration;
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   170
25984
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
   171
fun note_grouped kind group (a, ths) lthy = lthy
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
   172
  |> notes_grouped kind group [(a, [(ths, [])])] |>> the_single;
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
   173
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
   174
fun axioms kind = axioms_grouped kind "";
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
   175
fun define kind = define_grouped kind "";
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
   176
fun notes kind = notes_grouped kind "";
da0399c9ffcb grouped versions of axioms/define/notes;
wenzelm
parents: 25381
diff changeset
   177
fun note kind = note_grouped kind "";
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   178
24933
5471b164813b removed LocalTheory.defs/target_morphism operations;
wenzelm
parents: 24554
diff changeset
   179
fun target_morphism lthy =
5471b164813b removed LocalTheory.defs/target_morphism operations;
wenzelm
parents: 24554
diff changeset
   180
  ProofContext.export_morphism lthy (target_of lthy) $>
5471b164813b removed LocalTheory.defs/target_morphism operations;
wenzelm
parents: 24554
diff changeset
   181
  Morphism.thm_morphism Goal.norm_result;
5471b164813b removed LocalTheory.defs/target_morphism operations;
wenzelm
parents: 24554
diff changeset
   182
24949
5f00e3532418 generalized notation interface (add or del);
wenzelm
parents: 24933
diff changeset
   183
fun notation add mode raw_args lthy =
21743
fe94c6797c09 added notation/abbrev (from term_syntax.ML);
wenzelm
parents: 21700
diff changeset
   184
  let val args = map (apfst (Morphism.term (target_morphism lthy))) raw_args
24949
5f00e3532418 generalized notation interface (add or del);
wenzelm
parents: 24933
diff changeset
   185
  in term_syntax (ProofContext.target_notation add mode args) lthy end;
21743
fe94c6797c09 added notation/abbrev (from term_syntax.ML);
wenzelm
parents: 21700
diff changeset
   186
21664
dd4e89123359 notation/abbreviation: more serious handling of morphisms;
wenzelm
parents: 21614
diff changeset
   187
20910
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   188
(* init -- exit *)
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   189
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   190
fun init theory_prefix operations target = target |> Data.map
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   191
  (fn NONE => SOME (make_lthy (theory_prefix, operations, target))
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   192
    | SOME _ => error "Local theory already initialized");
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   193
21273
3b01db9504a8 abbrevs: return result;
wenzelm
parents: 21206
diff changeset
   194
fun restore lthy =
20910
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   195
  let val {theory_prefix, operations, target} = get_lthy lthy
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   196
  in init theory_prefix operations target end;
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   197
21292
11143b6ad87f simplified exit;
wenzelm
parents: 21273
diff changeset
   198
val reinit = operation #reinit;
11143b6ad87f simplified exit;
wenzelm
parents: 21273
diff changeset
   199
val exit = operation #exit;
20910
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   200
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   201
end;