src/Pure/context.ML
author wenzelm
Wed, 03 Nov 2010 21:53:56 +0100
changeset 40335 3e4bb6e7c3ca
parent 39020 ac0f24f850c9
child 41711 3422ae5aff3a
permissions -rw-r--r--
feeder: treat header as escaped utf8 to allow initial ML text to refer to non-ASCII file/directory names (e.g. "Documents/" on Chinese Ubuntu);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
6185
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/context.ML
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
     2
    Author:     Markus Wenzel, TU Muenchen
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
     3
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
     4
Generic theory contexts with unique identity, arbitrarily typed data,
24141
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
     5
monotonic development graph and history support.  Generic proof
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
     6
contexts with arbitrarily typed data.
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
     7
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
     8
Firm naming conventions:
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
     9
   thy, thy', thy1, thy2: theory
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    10
   ctxt, ctxt', ctxt1, ctxt2: Proof.context
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    11
   context: Context.generic
6185
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
    12
*)
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
    13
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
    14
signature BASIC_CONTEXT =
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
    15
sig
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    16
  type theory
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    17
  type theory_ref
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    18
  exception THEORY of string * theory list
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    19
  structure Proof: sig type context end
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    20
  structure ProofContext:
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    21
  sig
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    22
    val theory_of: Proof.context -> theory
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 34259
diff changeset
    23
    val init_global: theory -> Proof.context
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    24
  end
6185
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
    25
end;
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
    26
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
    27
signature CONTEXT =
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
    28
sig
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
    29
  include BASIC_CONTEXT
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    30
  (*theory context*)
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    31
  val parents_of: theory -> theory list
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    32
  val ancestors_of: theory -> theory list
29069
c7ba485581ae unified ids for ancestors and checkpoints, removed obsolete history of checkpoints;
wenzelm
parents: 29001
diff changeset
    33
  val theory_name: theory -> string
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    34
  val is_stale: theory -> bool
26623
81547c8d51f8 export is_draft, not draftN;
wenzelm
parents: 26486
diff changeset
    35
  val is_draft: theory -> bool
28317
83c4fc383409 added reject_draft;
wenzelm
parents: 28122
diff changeset
    36
  val reject_draft: theory -> theory
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
    37
  val PureN: string
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
    38
  val display_names: theory -> string list
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    39
  val pretty_thy: theory -> Pretty.T
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    40
  val string_of_thy: theory -> string
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    41
  val pretty_abbrev_thy: theory -> Pretty.T
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    42
  val str_of_thy: theory -> string
37867
b9783e9e96e1 added Context.get_theory -- avoid referring to accidental theory loader state (cf. Thy_Info.get_theory);
wenzelm
parents: 37852
diff changeset
    43
  val get_theory: theory -> string -> theory
37871
c7ce7685e087 replaced Thy_Info.the_theory by Context.this_theory -- avoid referring to accidental theory loader state;
wenzelm
parents: 37867
diff changeset
    44
  val this_theory: theory -> string -> theory
24141
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
    45
  val deref: theory_ref -> theory
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
    46
  val check_thy: theory -> theory_ref
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    47
  val eq_thy: theory * theory -> bool
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    48
  val subthy: theory * theory -> bool
16594
5d73fbf4eb1e added joinable;
wenzelm
parents: 16533
diff changeset
    49
  val joinable: theory * theory -> bool
23355
d2c033fd4514 merge/merge_refs: plain error instead of exception TERM;
wenzelm
parents: 22847
diff changeset
    50
  val merge: theory * theory -> theory
d2c033fd4514 merge/merge_refs: plain error instead of exception TERM;
wenzelm
parents: 22847
diff changeset
    51
  val merge_refs: theory_ref * theory_ref -> theory_ref
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    52
  val copy_thy: theory -> theory
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    53
  val checkpoint_thy: theory -> theory
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
    54
  val finish_thy: theory -> theory
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
    55
  val begin_thy: (theory -> Pretty.pp) -> string -> theory list -> theory
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
    56
  (*proof context*)
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    57
  val raw_transfer: theory -> Proof.context -> Proof.context
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
    58
  (*generic context*)
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    59
  datatype generic = Theory of theory | Proof of Proof.context
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    60
  val cases: (theory -> 'a) -> (Proof.context -> 'a) -> generic -> 'a
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    61
  val mapping: (theory -> theory) -> (Proof.context -> Proof.context) -> generic -> generic
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    62
  val mapping_result: (theory -> 'a * theory) -> (Proof.context -> 'a * Proof.context) ->
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    63
    generic -> 'a * generic
18632
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
    64
  val the_theory: generic -> theory
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    65
  val the_proof: generic -> Proof.context
18731
3989c3c41983 rename map_theory/proof to theory/proof_map;
wenzelm
parents: 18711
diff changeset
    66
  val map_theory: (theory -> theory) -> generic -> generic
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    67
  val map_proof: (Proof.context -> Proof.context) -> generic -> generic
26486
b65a5272b360 added map_theory_result, map_proof_result;
wenzelm
parents: 26463
diff changeset
    68
  val map_theory_result: (theory -> 'a * theory) -> generic -> 'a * generic
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    69
  val map_proof_result: (Proof.context -> 'a * Proof.context) -> generic -> 'a * generic
18731
3989c3c41983 rename map_theory/proof to theory/proof_map;
wenzelm
parents: 18711
diff changeset
    70
  val theory_map: (generic -> generic) -> theory -> theory
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    71
  val proof_map: (generic -> generic) -> Proof.context -> Proof.context
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    72
  val theory_of: generic -> theory  (*total*)
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    73
  val proof_of: generic -> Proof.context  (*total*)
26413
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
    74
  (*thread data*)
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
    75
  val thread_data: unit -> generic option
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
    76
  val the_thread_data: unit -> generic
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
    77
  val set_thread_data: generic option -> unit
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
    78
  val setmp_thread_data: generic option -> ('a -> 'b) -> 'a -> 'b
26463
9283b4185fdf Context.>> : operate on Context.generic;
wenzelm
parents: 26435
diff changeset
    79
  val >> : (generic -> generic) -> unit
9283b4185fdf Context.>> : operate on Context.generic;
wenzelm
parents: 26435
diff changeset
    80
  val >>> : (generic -> 'a * generic) -> 'a
6185
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
    81
end;
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
    82
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    83
signature PRIVATE_CONTEXT =
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    84
sig
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    85
  include CONTEXT
33033
wenzelm
parents: 33031
diff changeset
    86
  structure Theory_Data:
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    87
  sig
34245
25bd3ed2ac9f dropped copy operation for legacy TheoryDataFun
haftmann
parents: 33606
diff changeset
    88
    val declare: Object.T -> (Object.T -> Object.T) ->
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
    89
      (Pretty.pp -> Object.T * Object.T -> Object.T) -> serial
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    90
    val get: serial -> (Object.T -> 'a) -> theory -> 'a
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    91
    val put: serial -> ('a -> Object.T) -> 'a -> theory -> theory
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
    92
  end
33033
wenzelm
parents: 33031
diff changeset
    93
  structure Proof_Data:
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
    94
  sig
22847
22da6c4bc422 simplified DataFun interfaces: removed name/print, use adhoc value for uninitialized data, init only required for impure data;
wenzelm
parents: 22827
diff changeset
    95
    val declare: (theory -> Object.T) -> serial
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    96
    val get: serial -> (Object.T -> 'a) -> Proof.context -> 'a
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    97
    val put: serial -> ('a -> Object.T) -> 'a -> Proof.context -> Proof.context
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
    98
  end
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    99
end;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   100
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   101
structure Context: PRIVATE_CONTEXT =
6185
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
   102
struct
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
   103
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   104
(*** theory context ***)
6185
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
   105
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   106
(** theory data **)
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   107
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   108
(* data kinds and access methods *)
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   109
19028
6c238953f66c structure Datatab: private copy avoids potential conflict of table exceptions;
wenzelm
parents: 18931
diff changeset
   110
(*private copy avoids potential conflict of table exceptions*)
31971
8c1b845ed105 renamed functor TableFun to Table, and GraphFun to Graph;
wenzelm
parents: 30628
diff changeset
   111
structure Datatab = Table(type key = int val ord = int_ord);
19028
6c238953f66c structure Datatab: private copy avoids potential conflict of table exceptions;
wenzelm
parents: 18931
diff changeset
   112
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   113
local
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   114
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   115
type kind =
22847
22da6c4bc422 simplified DataFun interfaces: removed name/print, use adhoc value for uninitialized data, init only required for impure data;
wenzelm
parents: 22827
diff changeset
   116
 {empty: Object.T,
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   117
  extend: Object.T -> Object.T,
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   118
  merge: Pretty.pp -> Object.T * Object.T -> Object.T};
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   119
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 31971
diff changeset
   120
val kinds = Unsynchronized.ref (Datatab.empty: kind Datatab.table);
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   121
22847
22da6c4bc422 simplified DataFun interfaces: removed name/print, use adhoc value for uninitialized data, init only required for impure data;
wenzelm
parents: 22827
diff changeset
   122
fun invoke f k =
19028
6c238953f66c structure Datatab: private copy avoids potential conflict of table exceptions;
wenzelm
parents: 18931
diff changeset
   123
  (case Datatab.lookup (! kinds) k of
22847
22da6c4bc422 simplified DataFun interfaces: removed name/print, use adhoc value for uninitialized data, init only required for impure data;
wenzelm
parents: 22827
diff changeset
   124
    SOME kind => f kind
37852
a902f158b4fc eliminated old-style sys_error/SYS_ERROR in favour of exception Fail -- after careful checking that there is no overlap with existing handling of that;
wenzelm
parents: 37216
diff changeset
   125
  | NONE => raise Fail "Invalid theory data identifier");
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   126
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   127
in
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   128
33033
wenzelm
parents: 33031
diff changeset
   129
fun invoke_empty k = invoke (K o #empty) k ();
wenzelm
parents: 33031
diff changeset
   130
val invoke_extend = invoke #extend;
wenzelm
parents: 33031
diff changeset
   131
fun invoke_merge pp = invoke (fn kind => #merge kind pp);
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   132
34245
25bd3ed2ac9f dropped copy operation for legacy TheoryDataFun
haftmann
parents: 33606
diff changeset
   133
fun declare_theory_data empty extend merge =
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   134
  let
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   135
    val k = serial ();
34245
25bd3ed2ac9f dropped copy operation for legacy TheoryDataFun
haftmann
parents: 33606
diff changeset
   136
    val kind = {empty = empty, extend = extend, merge = merge};
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 31971
diff changeset
   137
    val _ = CRITICAL (fn () => Unsynchronized.change kinds (Datatab.update (k, kind)));
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   138
  in k end;
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   139
39020
ac0f24f850c9 replaced Table.map' by Table.map
haftmann
parents: 37871
diff changeset
   140
val extend_data = Datatab.map invoke_extend;
29367
741373421318 parallelized merge_data;
wenzelm
parents: 29095
diff changeset
   141
741373421318 parallelized merge_data;
wenzelm
parents: 29095
diff changeset
   142
fun merge_data pp (data1, data2) =
741373421318 parallelized merge_data;
wenzelm
parents: 29095
diff changeset
   143
  Datatab.keys (Datatab.merge (K true) (data1, data2))
29368
503ce3f8f092 renamed structure ParList to Par_List;
wenzelm
parents: 29367
diff changeset
   144
  |> Par_List.map (fn k =>
29367
741373421318 parallelized merge_data;
wenzelm
parents: 29095
diff changeset
   145
    (case (Datatab.lookup data1 k, Datatab.lookup data2 k) of
741373421318 parallelized merge_data;
wenzelm
parents: 29095
diff changeset
   146
      (SOME x, NONE) => (k, invoke_extend k x)
741373421318 parallelized merge_data;
wenzelm
parents: 29095
diff changeset
   147
    | (NONE, SOME y) => (k, invoke_extend k y)
741373421318 parallelized merge_data;
wenzelm
parents: 29095
diff changeset
   148
    | (SOME x, SOME y) => (k, invoke_merge pp k (invoke_extend k x, invoke_extend k y))))
741373421318 parallelized merge_data;
wenzelm
parents: 29095
diff changeset
   149
  |> Datatab.make;
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   150
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   151
end;
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   152
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   153
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   154
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   155
(** datatype theory **)
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   156
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   157
datatype theory =
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   158
  Theory of
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   159
   (*identity*)
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 31971
diff changeset
   160
   {self: theory Unsynchronized.ref option,  (*dynamic self reference -- follows theory changes*)
29095
a75f3ed534a0 tuned comments;
wenzelm
parents: 29093
diff changeset
   161
    draft: bool,                  (*draft mode -- linear destructive changes*)
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   162
    id: serial,                   (*identifier*)
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   163
    ids: unit Inttab.table} *     (*cumulative identifiers of non-drafts -- symbolic body content*)
29095
a75f3ed534a0 tuned comments;
wenzelm
parents: 29093
diff changeset
   164
   (*data*)
a75f3ed534a0 tuned comments;
wenzelm
parents: 29093
diff changeset
   165
   Object.T Datatab.table *       (*body content*)
a75f3ed534a0 tuned comments;
wenzelm
parents: 29093
diff changeset
   166
   (*ancestry*)
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   167
   {parents: theory list,         (*immediate predecessors*)
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   168
    ancestors: theory list} *     (*all predecessors -- canonical reverse order*)
29095
a75f3ed534a0 tuned comments;
wenzelm
parents: 29093
diff changeset
   169
   (*history*)
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   170
   {name: string,                 (*official theory name*)
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   171
    stage: int};                  (*checkpoint counter*)
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   172
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   173
exception THEORY of string * theory list;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   174
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   175
fun rep_theory (Theory args) = args;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   176
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   177
val identity_of = #1 o rep_theory;
33033
wenzelm
parents: 33031
diff changeset
   178
val data_of = #2 o rep_theory;
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   179
val ancestry_of = #3 o rep_theory;
33033
wenzelm
parents: 33031
diff changeset
   180
val history_of = #4 o rep_theory;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   181
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   182
fun make_identity self draft id ids = {self = self, draft = draft, id = id, ids = ids};
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   183
fun make_ancestry parents ancestors = {parents = parents, ancestors = ancestors};
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   184
fun make_history name stage = {name = name, stage = stage};
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   185
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   186
val the_self = the o #self o identity_of;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   187
val parents_of = #parents o ancestry_of;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   188
val ancestors_of = #ancestors o ancestry_of;
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   189
val theory_name = #name o history_of;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   190
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   191
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   192
(* staleness *)
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   193
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   194
fun eq_id (i: int, j) = i = j;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   195
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   196
fun is_stale
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 31971
diff changeset
   197
    (Theory ({self =
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 31971
diff changeset
   198
        SOME (Unsynchronized.ref (Theory ({id = id', ...}, _, _, _))), id, ...}, _, _, _)) =
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   199
      not (eq_id (id, id'))
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   200
  | is_stale (Theory ({self = NONE, ...}, _, _, _)) = true;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   201
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   202
fun vitalize (thy as Theory ({self = SOME r, ...}, _, _, _)) = (r := thy; thy)
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   203
  | vitalize (thy as Theory ({self = NONE, draft, id, ids}, data, ancestry, history)) =
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   204
      let
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 31971
diff changeset
   205
        val r = Unsynchronized.ref thy;
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   206
        val thy' = Theory (make_identity (SOME r) draft id ids, data, ancestry, history);
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   207
      in r := thy'; thy' end;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   208
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   209
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   210
(* draft mode *)
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   211
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   212
val is_draft = #draft o identity_of;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   213
28317
83c4fc383409 added reject_draft;
wenzelm
parents: 28122
diff changeset
   214
fun reject_draft thy =
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   215
  if is_draft thy then
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   216
    raise THEORY ("Illegal draft theory -- stable checkpoint required", [thy])
28317
83c4fc383409 added reject_draft;
wenzelm
parents: 28122
diff changeset
   217
  else thy;
83c4fc383409 added reject_draft;
wenzelm
parents: 28122
diff changeset
   218
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   219
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   220
(* names *)
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   221
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   222
val PureN = "Pure";
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   223
val draftN = "#";
29095
a75f3ed534a0 tuned comments;
wenzelm
parents: 29093
diff changeset
   224
val finished = ~1;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   225
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   226
fun display_names thy =
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   227
  let
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   228
    val draft = if is_draft thy then [draftN] else [];
29095
a75f3ed534a0 tuned comments;
wenzelm
parents: 29093
diff changeset
   229
    val {stage, ...} = history_of thy;
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   230
    val name =
29095
a75f3ed534a0 tuned comments;
wenzelm
parents: 29093
diff changeset
   231
      if stage = finished then theory_name thy
a75f3ed534a0 tuned comments;
wenzelm
parents: 29093
diff changeset
   232
      else theory_name thy ^ ":" ^ string_of_int stage;
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   233
    val ancestor_names = map theory_name (ancestors_of thy);
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   234
    val stale = if is_stale thy then ["!"] else [];
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   235
  in rev (stale @ draft @ [name] @ ancestor_names) end;
29069
c7ba485581ae unified ids for ancestors and checkpoints, removed obsolete history of checkpoints;
wenzelm
parents: 29001
diff changeset
   236
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   237
val pretty_thy = Pretty.str_list "{" "}" o display_names;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   238
val string_of_thy = Pretty.string_of o pretty_thy;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   239
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   240
fun pretty_abbrev_thy thy =
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   241
  let
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   242
    val names = display_names thy;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   243
    val n = length names;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   244
    val abbrev = if n > 5 then "..." :: List.drop (names, n - 5) else names;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   245
  in Pretty.str_list "{" "}" abbrev end;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   246
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   247
val str_of_thy = Pretty.str_of o pretty_abbrev_thy;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   248
37867
b9783e9e96e1 added Context.get_theory -- avoid referring to accidental theory loader state (cf. Thy_Info.get_theory);
wenzelm
parents: 37852
diff changeset
   249
fun get_theory thy name =
b9783e9e96e1 added Context.get_theory -- avoid referring to accidental theory loader state (cf. Thy_Info.get_theory);
wenzelm
parents: 37852
diff changeset
   250
  if theory_name thy <> name then
b9783e9e96e1 added Context.get_theory -- avoid referring to accidental theory loader state (cf. Thy_Info.get_theory);
wenzelm
parents: 37852
diff changeset
   251
    (case find_first (fn thy' => theory_name thy' = name) (ancestors_of thy) of
b9783e9e96e1 added Context.get_theory -- avoid referring to accidental theory loader state (cf. Thy_Info.get_theory);
wenzelm
parents: 37852
diff changeset
   252
      SOME thy' => thy'
b9783e9e96e1 added Context.get_theory -- avoid referring to accidental theory loader state (cf. Thy_Info.get_theory);
wenzelm
parents: 37852
diff changeset
   253
    | NONE => error ("Unknown ancestor theory " ^ quote name))
b9783e9e96e1 added Context.get_theory -- avoid referring to accidental theory loader state (cf. Thy_Info.get_theory);
wenzelm
parents: 37852
diff changeset
   254
  else if #stage (history_of thy) = finished then thy
b9783e9e96e1 added Context.get_theory -- avoid referring to accidental theory loader state (cf. Thy_Info.get_theory);
wenzelm
parents: 37852
diff changeset
   255
  else error ("Unfinished theory " ^ quote name);
b9783e9e96e1 added Context.get_theory -- avoid referring to accidental theory loader state (cf. Thy_Info.get_theory);
wenzelm
parents: 37852
diff changeset
   256
37871
c7ce7685e087 replaced Thy_Info.the_theory by Context.this_theory -- avoid referring to accidental theory loader state;
wenzelm
parents: 37867
diff changeset
   257
fun this_theory thy name =
c7ce7685e087 replaced Thy_Info.the_theory by Context.this_theory -- avoid referring to accidental theory loader state;
wenzelm
parents: 37867
diff changeset
   258
  if theory_name thy = name then thy
c7ce7685e087 replaced Thy_Info.the_theory by Context.this_theory -- avoid referring to accidental theory loader state;
wenzelm
parents: 37867
diff changeset
   259
  else get_theory thy name;
c7ce7685e087 replaced Thy_Info.the_theory by Context.this_theory -- avoid referring to accidental theory loader state;
wenzelm
parents: 37867
diff changeset
   260
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   261
24141
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   262
(* theory references *)
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   263
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   264
(*theory_ref provides a safe way to store dynamic references to a
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   265
  theory in external data structures -- a plain theory value would
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   266
  become stale as the self reference moves on*)
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   267
33033
wenzelm
parents: 33031
diff changeset
   268
datatype theory_ref = Theory_Ref of theory Unsynchronized.ref;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   269
33033
wenzelm
parents: 33031
diff changeset
   270
fun deref (Theory_Ref (Unsynchronized.ref thy)) = thy;
24141
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   271
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   272
fun check_thy thy =  (*thread-safe version*)
33033
wenzelm
parents: 33031
diff changeset
   273
  let val thy_ref = Theory_Ref (the_self thy) in
24141
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   274
    if is_stale thy then error ("Stale theory encountered:\n" ^ string_of_thy thy)
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   275
    else thy_ref
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   276
  end;
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   277
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   278
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   279
(* build ids *)
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   280
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   281
fun insert_id draft id ids =
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   282
  if draft then ids
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   283
  else Inttab.update (id, ()) ids;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   284
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   285
fun merge_ids
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   286
    (Theory ({draft = draft1, id = id1, ids = ids1, ...}, _, _, _))
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   287
    (Theory ({draft = draft2, id = id2, ids = ids2, ...}, _, _, _)) =
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   288
  Inttab.merge (K true) (ids1, ids2)
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   289
  |> insert_id draft1 id1
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   290
  |> insert_id draft2 id2;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   291
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   292
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   293
(* equality and inclusion *)
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   294
24141
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   295
val eq_thy = eq_id o pairself (#id o identity_of);
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   296
29069
c7ba485581ae unified ids for ancestors and checkpoints, removed obsolete history of checkpoints;
wenzelm
parents: 29001
diff changeset
   297
fun proper_subthy (Theory ({id, ...}, _, _, _), Theory ({ids, ...}, _, _, _)) =
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   298
  Inttab.defined ids id;
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   299
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   300
fun subthy thys = eq_thy thys orelse proper_subthy thys;
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   301
16594
5d73fbf4eb1e added joinable;
wenzelm
parents: 16533
diff changeset
   302
fun joinable (thy1, thy2) = subthy (thy1, thy2) orelse subthy (thy2, thy1);
5d73fbf4eb1e added joinable;
wenzelm
parents: 16533
diff changeset
   303
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   304
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   305
(* consistent ancestors *)
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   306
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   307
fun extend_ancestors thy thys =
33033
wenzelm
parents: 33031
diff changeset
   308
  if member eq_thy thys thy then
wenzelm
parents: 33031
diff changeset
   309
    raise THEORY ("Duplicate theory node", thy :: thys)
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   310
  else thy :: thys;
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   311
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   312
fun extend_ancestors_of thy = extend_ancestors thy (ancestors_of thy);
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   313
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   314
val merge_ancestors = merge (fn (thy1, thy2) =>
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   315
  eq_thy (thy1, thy2) orelse
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   316
    theory_name thy1 = theory_name thy2 andalso
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   317
      raise THEORY ("Inconsistent theory versions", [thy1, thy2]));
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   318
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   319
23355
d2c033fd4514 merge/merge_refs: plain error instead of exception TERM;
wenzelm
parents: 22847
diff changeset
   320
(* trivial merge *)
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   321
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   322
fun merge (thy1, thy2) =
16719
5c5eb939f6eb check_thy: less invocations, less verbose;
wenzelm
parents: 16594
diff changeset
   323
  if eq_thy (thy1, thy2) then thy1
5c5eb939f6eb check_thy: less invocations, less verbose;
wenzelm
parents: 16594
diff changeset
   324
  else if proper_subthy (thy2, thy1) then thy1
5c5eb939f6eb check_thy: less invocations, less verbose;
wenzelm
parents: 16594
diff changeset
   325
  else if proper_subthy (thy1, thy2) then thy2
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   326
  else error (cat_lines ["Attempt to perform non-trivial merge of theories:",
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   327
    str_of_thy thy1, str_of_thy thy2]);
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   328
16719
5c5eb939f6eb check_thy: less invocations, less verbose;
wenzelm
parents: 16594
diff changeset
   329
fun merge_refs (ref1, ref2) =
5c5eb939f6eb check_thy: less invocations, less verbose;
wenzelm
parents: 16594
diff changeset
   330
  if ref1 = ref2 then ref1
24141
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   331
  else check_thy (merge (deref ref1, deref ref2));
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   332
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   333
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   334
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   335
(** build theories **)
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   336
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   337
(* primitives *)
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   338
33606
2b27020ffcb2 local mutex for theory content/identity operations;
wenzelm
parents: 33517
diff changeset
   339
local
2b27020ffcb2 local mutex for theory content/identity operations;
wenzelm
parents: 33517
diff changeset
   340
  val lock = Mutex.mutex ();
2b27020ffcb2 local mutex for theory content/identity operations;
wenzelm
parents: 33517
diff changeset
   341
in
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 36610
diff changeset
   342
  fun SYNCHRONIZED e = Simple_Thread.synchronized "theory" lock e;
33606
2b27020ffcb2 local mutex for theory content/identity operations;
wenzelm
parents: 33517
diff changeset
   343
end;
2b27020ffcb2 local mutex for theory content/identity operations;
wenzelm
parents: 33517
diff changeset
   344
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   345
fun create_thy self draft ids data ancestry history =
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   346
  let val identity = make_identity self draft (serial ()) ids;
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   347
  in vitalize (Theory (identity, data, ancestry, history)) end;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   348
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   349
fun change_thy draft' f thy =
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   350
  let
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   351
    val Theory ({self, draft, id, ids}, data, ancestry, history) = thy;
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   352
    val (self', data', ancestry') =
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   353
      if draft then (self, data, ancestry)    (*destructive change!*)
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   354
      else if #stage history > 0
34245
25bd3ed2ac9f dropped copy operation for legacy TheoryDataFun
haftmann
parents: 33606
diff changeset
   355
      then (NONE, data, ancestry)
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   356
      else (NONE, extend_data data, make_ancestry [thy] (extend_ancestors_of thy));
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   357
    val ids' = insert_id draft id ids;
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   358
    val data'' = f data';
33606
2b27020ffcb2 local mutex for theory content/identity operations;
wenzelm
parents: 33517
diff changeset
   359
    val thy' = SYNCHRONIZED (fn () =>
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   360
      (check_thy thy; create_thy self' draft' ids' data'' ancestry' history));
24369
0cb1f4d76452 tuned CRITICAL sections;
wenzelm
parents: 24184
diff changeset
   361
  in thy' end;
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   362
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   363
val name_thy = change_thy false I;
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   364
val extend_thy = change_thy true I;
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   365
val modify_thy = change_thy true;
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   366
24369
0cb1f4d76452 tuned CRITICAL sections;
wenzelm
parents: 24184
diff changeset
   367
fun copy_thy thy =
24141
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   368
  let
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   369
    val Theory ({draft, id, ids, ...}, data, ancestry, history) = thy;
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   370
    val ids' = insert_id draft id ids;
33606
2b27020ffcb2 local mutex for theory content/identity operations;
wenzelm
parents: 33517
diff changeset
   371
    val thy' = SYNCHRONIZED (fn () =>
34245
25bd3ed2ac9f dropped copy operation for legacy TheoryDataFun
haftmann
parents: 33606
diff changeset
   372
      (check_thy thy; create_thy NONE true ids' data ancestry history));
24369
0cb1f4d76452 tuned CRITICAL sections;
wenzelm
parents: 24184
diff changeset
   373
  in thy' end;
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   374
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   375
val pre_pure_thy = create_thy NONE true Inttab.empty
29069
c7ba485581ae unified ids for ancestors and checkpoints, removed obsolete history of checkpoints;
wenzelm
parents: 29001
diff changeset
   376
  Datatab.empty (make_ancestry [] []) (make_history PureN 0);
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   377
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   378
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   379
(* named theory nodes *)
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   380
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   381
fun merge_thys pp (thy1, thy2) =
26957
e3f04fdd994d eliminated theory CPure;
wenzelm
parents: 26889
diff changeset
   382
  let
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   383
    val ids = merge_ids thy1 thy2;
26957
e3f04fdd994d eliminated theory CPure;
wenzelm
parents: 26889
diff changeset
   384
    val data = merge_data (pp thy1) (data_of thy1, data_of thy2);
e3f04fdd994d eliminated theory CPure;
wenzelm
parents: 26889
diff changeset
   385
    val ancestry = make_ancestry [] [];
29069
c7ba485581ae unified ids for ancestors and checkpoints, removed obsolete history of checkpoints;
wenzelm
parents: 29001
diff changeset
   386
    val history = make_history "" 0;
33606
2b27020ffcb2 local mutex for theory content/identity operations;
wenzelm
parents: 33517
diff changeset
   387
    val thy' = SYNCHRONIZED (fn () =>
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   388
     (check_thy thy1; check_thy thy2; create_thy NONE true ids data ancestry history));
26957
e3f04fdd994d eliminated theory CPure;
wenzelm
parents: 26889
diff changeset
   389
  in thy' end;
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   390
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   391
fun maximal_thys thys =
28617
wenzelm
parents: 28375
diff changeset
   392
  thys |> filter_out (fn thy => exists (fn thy' => proper_subthy (thy, thy')) thys);
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   393
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   394
fun begin_thy pp name imports =
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   395
  if name = "" orelse name = draftN then error ("Bad theory name: " ^ quote name)
24369
0cb1f4d76452 tuned CRITICAL sections;
wenzelm
parents: 24184
diff changeset
   396
  else
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   397
    let
24141
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   398
      val parents = maximal_thys (distinct eq_thy imports);
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   399
      val ancestors =
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   400
        Library.foldl merge_ancestors ([], map ancestors_of parents)
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   401
        |> fold extend_ancestors parents;
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   402
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   403
      val Theory ({ids, ...}, data, _, _) =
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   404
        (case parents of
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   405
          [] => error "No parent theories"
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   406
        | [thy] => extend_thy thy
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   407
        | thy :: thys => Library.foldl (merge_thys pp) (thy, thys));
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   408
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   409
      val ancestry = make_ancestry parents ancestors;
29069
c7ba485581ae unified ids for ancestors and checkpoints, removed obsolete history of checkpoints;
wenzelm
parents: 29001
diff changeset
   410
      val history = make_history name 0;
33606
2b27020ffcb2 local mutex for theory content/identity operations;
wenzelm
parents: 33517
diff changeset
   411
      val thy' = SYNCHRONIZED (fn () =>
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   412
        (map check_thy imports; create_thy NONE true ids data ancestry history));
24369
0cb1f4d76452 tuned CRITICAL sections;
wenzelm
parents: 24184
diff changeset
   413
    in thy' end;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   414
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   415
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   416
(* history stages *)
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   417
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   418
fun history_stage f thy =
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   419
  let
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   420
    val {name, stage} = history_of thy;
29095
a75f3ed534a0 tuned comments;
wenzelm
parents: 29093
diff changeset
   421
    val _ = stage = finished andalso raise THEORY ("Theory already finished", [thy]);
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   422
    val history' = make_history name (f stage);
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   423
    val thy' as Theory (identity', data', ancestry', _) = name_thy thy;
33606
2b27020ffcb2 local mutex for theory content/identity operations;
wenzelm
parents: 33517
diff changeset
   424
    val thy'' = SYNCHRONIZED (fn () =>
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   425
      (check_thy thy'; vitalize (Theory (identity', data', ancestry', history'))));
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   426
  in thy'' end;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   427
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   428
fun checkpoint_thy thy =
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   429
  if is_draft thy then history_stage (fn stage => stage + 1) thy
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   430
  else thy;
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   431
29095
a75f3ed534a0 tuned comments;
wenzelm
parents: 29093
diff changeset
   432
val finish_thy = history_stage (fn _ => finished);
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   433
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   434
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   435
(* theory data *)
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   436
33033
wenzelm
parents: 33031
diff changeset
   437
structure Theory_Data =
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   438
struct
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   439
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   440
val declare = declare_theory_data;
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   441
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   442
fun get k dest thy =
34253
5930c6391126 removed further remains of mutable theory data (cf. 25bd3ed2ac9f);
wenzelm
parents: 34245
diff changeset
   443
  (case Datatab.lookup (data_of thy) k of
22847
22da6c4bc422 simplified DataFun interfaces: removed name/print, use adhoc value for uninitialized data, init only required for impure data;
wenzelm
parents: 22827
diff changeset
   444
    SOME x => x
34253
5930c6391126 removed further remains of mutable theory data (cf. 25bd3ed2ac9f);
wenzelm
parents: 34245
diff changeset
   445
  | NONE => invoke_empty k) |> dest;
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   446
22847
22da6c4bc422 simplified DataFun interfaces: removed name/print, use adhoc value for uninitialized data, init only required for impure data;
wenzelm
parents: 22827
diff changeset
   447
fun put k mk x = modify_thy (Datatab.update (k, mk x));
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   448
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   449
end;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   450
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   451
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   452
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   453
(*** proof context ***)
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   454
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   455
(* datatype Proof.context *)
17060
cca2f3938443 type proof: theory_ref instead of theory (make proof contexts independent entities);
wenzelm
parents: 16894
diff changeset
   456
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   457
structure Proof =
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   458
struct
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   459
  datatype context = Context of Object.T Datatab.table * theory_ref;
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   460
end;
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   461
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   462
fun theory_of_proof (Proof.Context (_, thy_ref)) = deref thy_ref;
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   463
fun data_of_proof (Proof.Context (data, _)) = data;
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   464
fun map_prf f (Proof.Context (data, thy_ref)) = Proof.Context (f data, thy_ref);
17060
cca2f3938443 type proof: theory_ref instead of theory (make proof contexts independent entities);
wenzelm
parents: 16894
diff changeset
   465
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   466
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   467
(* proof data kinds *)
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   468
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   469
local
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   470
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 31971
diff changeset
   471
val kinds = Unsynchronized.ref (Datatab.empty: (theory -> Object.T) Datatab.table);
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   472
22847
22da6c4bc422 simplified DataFun interfaces: removed name/print, use adhoc value for uninitialized data, init only required for impure data;
wenzelm
parents: 22827
diff changeset
   473
fun invoke_init k =
19028
6c238953f66c structure Datatab: private copy avoids potential conflict of table exceptions;
wenzelm
parents: 18931
diff changeset
   474
  (case Datatab.lookup (! kinds) k of
22847
22da6c4bc422 simplified DataFun interfaces: removed name/print, use adhoc value for uninitialized data, init only required for impure data;
wenzelm
parents: 22827
diff changeset
   475
    SOME init => init
37852
a902f158b4fc eliminated old-style sys_error/SYS_ERROR in favour of exception Fail -- after careful checking that there is no overlap with existing handling of that;
wenzelm
parents: 37216
diff changeset
   476
  | NONE => raise Fail "Invalid proof data identifier");
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   477
22847
22da6c4bc422 simplified DataFun interfaces: removed name/print, use adhoc value for uninitialized data, init only required for impure data;
wenzelm
parents: 22827
diff changeset
   478
fun init_data thy =
39020
ac0f24f850c9 replaced Table.map' by Table.map
haftmann
parents: 37871
diff changeset
   479
  Datatab.map (fn k => fn _ => invoke_init k thy) (! kinds);
22847
22da6c4bc422 simplified DataFun interfaces: removed name/print, use adhoc value for uninitialized data, init only required for impure data;
wenzelm
parents: 22827
diff changeset
   480
22da6c4bc422 simplified DataFun interfaces: removed name/print, use adhoc value for uninitialized data, init only required for impure data;
wenzelm
parents: 22827
diff changeset
   481
fun init_new_data data thy =
22da6c4bc422 simplified DataFun interfaces: removed name/print, use adhoc value for uninitialized data, init only required for impure data;
wenzelm
parents: 22827
diff changeset
   482
  Datatab.merge (K true) (data, init_data thy);
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   483
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   484
in
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   485
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   486
fun raw_transfer thy' (Proof.Context (data, thy_ref)) =
24141
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   487
  let
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   488
    val thy = deref thy_ref;
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   489
    val _ = subthy (thy, thy') orelse error "transfer proof context: not a super theory";
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   490
    val _ = check_thy thy;
24184
19cb051154fd thread-safeness: when creating certified items, perform Theory.check_thy *last*;
wenzelm
parents: 24141
diff changeset
   491
    val data' = init_new_data data thy';
24141
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   492
    val thy_ref' = check_thy thy';
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   493
  in Proof.Context (data', thy_ref') end;
22847
22da6c4bc422 simplified DataFun interfaces: removed name/print, use adhoc value for uninitialized data, init only required for impure data;
wenzelm
parents: 22827
diff changeset
   494
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   495
structure ProofContext =
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   496
struct
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   497
  val theory_of = theory_of_proof;
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 34259
diff changeset
   498
  fun init_global thy = Proof.Context (init_data thy, check_thy thy);
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   499
end;
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   500
33033
wenzelm
parents: 33031
diff changeset
   501
structure Proof_Data =
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   502
struct
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   503
22847
22da6c4bc422 simplified DataFun interfaces: removed name/print, use adhoc value for uninitialized data, init only required for impure data;
wenzelm
parents: 22827
diff changeset
   504
fun declare init =
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   505
  let
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   506
    val k = serial ();
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 31971
diff changeset
   507
    val _ = CRITICAL (fn () => Unsynchronized.change kinds (Datatab.update (k, init)));
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   508
  in k end;
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   509
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   510
fun get k dest prf =
22847
22da6c4bc422 simplified DataFun interfaces: removed name/print, use adhoc value for uninitialized data, init only required for impure data;
wenzelm
parents: 22827
diff changeset
   511
  dest (case Datatab.lookup (data_of_proof prf) k of
22da6c4bc422 simplified DataFun interfaces: removed name/print, use adhoc value for uninitialized data, init only required for impure data;
wenzelm
parents: 22827
diff changeset
   512
    SOME x => x
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   513
  | NONE => invoke_init k (ProofContext.theory_of prf));   (*adhoc value*)
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   514
19028
6c238953f66c structure Datatab: private copy avoids potential conflict of table exceptions;
wenzelm
parents: 18931
diff changeset
   515
fun put k mk x = map_prf (Datatab.update (k, mk x));
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   516
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   517
end;
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   518
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   519
end;
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   520
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   521
18632
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   522
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   523
(*** generic context ***)
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   524
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   525
datatype generic = Theory of theory | Proof of Proof.context;
18632
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   526
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   527
fun cases f _ (Theory thy) = f thy
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   528
  | cases _ g (Proof prf) = g prf;
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   529
19678
d1a15431de34 added mapping;
wenzelm
parents: 19482
diff changeset
   530
fun mapping f g = cases (Theory o f) (Proof o g);
21660
c86b761d6c06 added mapping_result;
wenzelm
parents: 21518
diff changeset
   531
fun mapping_result f g = cases (apsnd Theory o f) (apsnd Proof o g);
19678
d1a15431de34 added mapping;
wenzelm
parents: 19482
diff changeset
   532
23595
7ca68a2c8575 the_theory/proof: error instead of exception Fail;
wenzelm
parents: 23355
diff changeset
   533
val the_theory = cases I (fn _ => error "Ill-typed context: theory expected");
7ca68a2c8575 the_theory/proof: error instead of exception Fail;
wenzelm
parents: 23355
diff changeset
   534
val the_proof = cases (fn _ => error "Ill-typed context: proof expected") I;
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   535
18731
3989c3c41983 rename map_theory/proof to theory/proof_map;
wenzelm
parents: 18711
diff changeset
   536
fun map_theory f = Theory o f o the_theory;
3989c3c41983 rename map_theory/proof to theory/proof_map;
wenzelm
parents: 18711
diff changeset
   537
fun map_proof f = Proof o f o the_proof;
3989c3c41983 rename map_theory/proof to theory/proof_map;
wenzelm
parents: 18711
diff changeset
   538
26486
b65a5272b360 added map_theory_result, map_proof_result;
wenzelm
parents: 26463
diff changeset
   539
fun map_theory_result f = apsnd Theory o f o the_theory;
b65a5272b360 added map_theory_result, map_proof_result;
wenzelm
parents: 26463
diff changeset
   540
fun map_proof_result f = apsnd Proof o f o the_proof;
b65a5272b360 added map_theory_result, map_proof_result;
wenzelm
parents: 26463
diff changeset
   541
18731
3989c3c41983 rename map_theory/proof to theory/proof_map;
wenzelm
parents: 18711
diff changeset
   542
fun theory_map f = the_theory o f o Theory;
3989c3c41983 rename map_theory/proof to theory/proof_map;
wenzelm
parents: 18711
diff changeset
   543
fun proof_map f = the_proof o f o Proof;
18665
5198a2c4c00e added map_theory, map_proof;
wenzelm
parents: 18632
diff changeset
   544
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   545
val theory_of = cases I ProofContext.theory_of;
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 34259
diff changeset
   546
val proof_of = cases ProofContext.init_global I;
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   547
22085
c138cfd500f7 ML context: full generic context, tuned signature;
wenzelm
parents: 21962
diff changeset
   548
c138cfd500f7 ML context: full generic context, tuned signature;
wenzelm
parents: 21962
diff changeset
   549
26413
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   550
(** thread data **)
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   551
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   552
local val tag = Universal.tag () : generic option Universal.tag in
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   553
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   554
fun thread_data () =
28122
3d099ce624e7 Thread.getLocal/setLocal;
wenzelm
parents: 27341
diff changeset
   555
  (case Thread.getLocal tag of
26413
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   556
    SOME (SOME context) => SOME context
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   557
  | _ => NONE);
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   558
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   559
fun the_thread_data () =
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   560
  (case thread_data () of
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   561
    SOME context => context
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   562
  | _ => error "Unknown context");
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   563
28122
3d099ce624e7 Thread.getLocal/setLocal;
wenzelm
parents: 27341
diff changeset
   564
fun set_thread_data context = Thread.setLocal (tag, context);
26413
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   565
fun setmp_thread_data context = Library.setmp_thread_data tag (thread_data ()) context;
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   566
26428
5b2beca2087d eliminated theory ProtoPure;
wenzelm
parents: 26421
diff changeset
   567
end;
5b2beca2087d eliminated theory ProtoPure;
wenzelm
parents: 26421
diff changeset
   568
5b2beca2087d eliminated theory ProtoPure;
wenzelm
parents: 26421
diff changeset
   569
fun >>> f =
5b2beca2087d eliminated theory ProtoPure;
wenzelm
parents: 26421
diff changeset
   570
  let
26463
9283b4185fdf Context.>> : operate on Context.generic;
wenzelm
parents: 26435
diff changeset
   571
    val (res, context') = f (the_thread_data ());
9283b4185fdf Context.>> : operate on Context.generic;
wenzelm
parents: 26435
diff changeset
   572
    val _ = set_thread_data (SOME context');
26428
5b2beca2087d eliminated theory ProtoPure;
wenzelm
parents: 26421
diff changeset
   573
  in res end;
5b2beca2087d eliminated theory ProtoPure;
wenzelm
parents: 26421
diff changeset
   574
26421
3dfb36923a56 nonfix >>;
wenzelm
parents: 26413
diff changeset
   575
nonfix >>;
26463
9283b4185fdf Context.>> : operate on Context.generic;
wenzelm
parents: 26435
diff changeset
   576
fun >> f = >>> (fn context => ((), f context));
26413
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   577
26428
5b2beca2087d eliminated theory ProtoPure;
wenzelm
parents: 26421
diff changeset
   578
val _ = set_thread_data (SOME (Theory pre_pure_thy));
26413
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   579
6185
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
   580
end;
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
   581
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   582
structure Basic_Context: BASIC_CONTEXT = Context;
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   583
open Basic_Context;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   584
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   585
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   586
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   587
(*** type-safe interfaces for data declarations ***)
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   588
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   589
(** theory data **)
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   590
34259
2ba492b8b6e8 discontinued old TheoryDataFun, but retain Theory_Data_PP with is Pretty.pp argument to merge (still required in exotic situations -- hard to get rid of);
wenzelm
parents: 34253
diff changeset
   591
signature THEORY_DATA_PP_ARGS =
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   592
sig
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   593
  type T
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   594
  val empty: T
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   595
  val extend: T -> T
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   596
  val merge: Pretty.pp -> T * T -> T
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   597
end;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   598
34259
2ba492b8b6e8 discontinued old TheoryDataFun, but retain Theory_Data_PP with is Pretty.pp argument to merge (still required in exotic situations -- hard to get rid of);
wenzelm
parents: 34253
diff changeset
   599
signature THEORY_DATA_ARGS =
2ba492b8b6e8 discontinued old TheoryDataFun, but retain Theory_Data_PP with is Pretty.pp argument to merge (still required in exotic situations -- hard to get rid of);
wenzelm
parents: 34253
diff changeset
   600
sig
2ba492b8b6e8 discontinued old TheoryDataFun, but retain Theory_Data_PP with is Pretty.pp argument to merge (still required in exotic situations -- hard to get rid of);
wenzelm
parents: 34253
diff changeset
   601
  type T
2ba492b8b6e8 discontinued old TheoryDataFun, but retain Theory_Data_PP with is Pretty.pp argument to merge (still required in exotic situations -- hard to get rid of);
wenzelm
parents: 34253
diff changeset
   602
  val empty: T
2ba492b8b6e8 discontinued old TheoryDataFun, but retain Theory_Data_PP with is Pretty.pp argument to merge (still required in exotic situations -- hard to get rid of);
wenzelm
parents: 34253
diff changeset
   603
  val extend: T -> T
2ba492b8b6e8 discontinued old TheoryDataFun, but retain Theory_Data_PP with is Pretty.pp argument to merge (still required in exotic situations -- hard to get rid of);
wenzelm
parents: 34253
diff changeset
   604
  val merge: T * T -> T
2ba492b8b6e8 discontinued old TheoryDataFun, but retain Theory_Data_PP with is Pretty.pp argument to merge (still required in exotic situations -- hard to get rid of);
wenzelm
parents: 34253
diff changeset
   605
end;
2ba492b8b6e8 discontinued old TheoryDataFun, but retain Theory_Data_PP with is Pretty.pp argument to merge (still required in exotic situations -- hard to get rid of);
wenzelm
parents: 34253
diff changeset
   606
34253
5930c6391126 removed further remains of mutable theory data (cf. 25bd3ed2ac9f);
wenzelm
parents: 34245
diff changeset
   607
signature THEORY_DATA =
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   608
sig
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   609
  type T
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   610
  val get: theory -> T
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   611
  val put: T -> theory -> theory
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   612
  val map: (T -> T) -> theory -> theory
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   613
end;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   614
34259
2ba492b8b6e8 discontinued old TheoryDataFun, but retain Theory_Data_PP with is Pretty.pp argument to merge (still required in exotic situations -- hard to get rid of);
wenzelm
parents: 34253
diff changeset
   615
functor Theory_Data_PP(Data: THEORY_DATA_PP_ARGS): THEORY_DATA =
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   616
struct
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   617
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   618
type T = Data.T;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   619
exception Data of T;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   620
33033
wenzelm
parents: 33031
diff changeset
   621
val kind = Context.Theory_Data.declare
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   622
  (Data Data.empty)
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   623
  (fn Data x => Data (Data.extend x))
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   624
  (fn pp => fn (Data x1, Data x2) => Data (Data.merge pp (x1, x2)));
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   625
33033
wenzelm
parents: 33031
diff changeset
   626
val get = Context.Theory_Data.get kind (fn Data x => x);
wenzelm
parents: 33031
diff changeset
   627
val put = Context.Theory_Data.put kind Data;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   628
fun map f thy = put (f (get thy)) thy;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   629
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   630
end;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   631
33517
d064fa48f305 modernized/simplified functor Theory_Data, Proof_Data, Generic_Data: eliminated Pretty.pp, discontinued mutable data;
wenzelm
parents: 33033
diff changeset
   632
functor Theory_Data(Data: THEORY_DATA_ARGS): THEORY_DATA =
34259
2ba492b8b6e8 discontinued old TheoryDataFun, but retain Theory_Data_PP with is Pretty.pp argument to merge (still required in exotic situations -- hard to get rid of);
wenzelm
parents: 34253
diff changeset
   633
  Theory_Data_PP
2ba492b8b6e8 discontinued old TheoryDataFun, but retain Theory_Data_PP with is Pretty.pp argument to merge (still required in exotic situations -- hard to get rid of);
wenzelm
parents: 34253
diff changeset
   634
  (
2ba492b8b6e8 discontinued old TheoryDataFun, but retain Theory_Data_PP with is Pretty.pp argument to merge (still required in exotic situations -- hard to get rid of);
wenzelm
parents: 34253
diff changeset
   635
    type T = Data.T;
2ba492b8b6e8 discontinued old TheoryDataFun, but retain Theory_Data_PP with is Pretty.pp argument to merge (still required in exotic situations -- hard to get rid of);
wenzelm
parents: 34253
diff changeset
   636
    val empty = Data.empty;
2ba492b8b6e8 discontinued old TheoryDataFun, but retain Theory_Data_PP with is Pretty.pp argument to merge (still required in exotic situations -- hard to get rid of);
wenzelm
parents: 34253
diff changeset
   637
    val extend = Data.extend;
2ba492b8b6e8 discontinued old TheoryDataFun, but retain Theory_Data_PP with is Pretty.pp argument to merge (still required in exotic situations -- hard to get rid of);
wenzelm
parents: 34253
diff changeset
   638
    fun merge _ = Data.merge;
2ba492b8b6e8 discontinued old TheoryDataFun, but retain Theory_Data_PP with is Pretty.pp argument to merge (still required in exotic situations -- hard to get rid of);
wenzelm
parents: 34253
diff changeset
   639
  );
33517
d064fa48f305 modernized/simplified functor Theory_Data, Proof_Data, Generic_Data: eliminated Pretty.pp, discontinued mutable data;
wenzelm
parents: 33033
diff changeset
   640
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   641
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   642
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   643
(** proof data **)
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   644
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   645
signature PROOF_DATA_ARGS =
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   646
sig
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   647
  type T
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   648
  val init: theory -> T
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   649
end;
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   650
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   651
signature PROOF_DATA =
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   652
sig
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   653
  type T
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   654
  val get: Proof.context -> T
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   655
  val put: T -> Proof.context -> Proof.context
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   656
  val map: (T -> T) -> Proof.context -> Proof.context
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   657
end;
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   658
33517
d064fa48f305 modernized/simplified functor Theory_Data, Proof_Data, Generic_Data: eliminated Pretty.pp, discontinued mutable data;
wenzelm
parents: 33033
diff changeset
   659
functor Proof_Data(Data: PROOF_DATA_ARGS): PROOF_DATA =
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   660
struct
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   661
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   662
type T = Data.T;
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   663
exception Data of T;
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   664
33033
wenzelm
parents: 33031
diff changeset
   665
val kind = Context.Proof_Data.declare (Data o Data.init);
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   666
33033
wenzelm
parents: 33031
diff changeset
   667
val get = Context.Proof_Data.get kind (fn Data x => x);
wenzelm
parents: 33031
diff changeset
   668
val put = Context.Proof_Data.put kind Data;
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   669
fun map f prf = put (f (get prf)) prf;
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   670
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   671
end;
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   672
18632
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   673
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   674
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   675
(** generic data **)
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   676
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   677
signature GENERIC_DATA_ARGS =
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   678
sig
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   679
  type T
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   680
  val empty: T
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   681
  val extend: T -> T
33517
d064fa48f305 modernized/simplified functor Theory_Data, Proof_Data, Generic_Data: eliminated Pretty.pp, discontinued mutable data;
wenzelm
parents: 33033
diff changeset
   682
  val merge: T * T -> T
18632
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   683
end;
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   684
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   685
signature GENERIC_DATA =
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   686
sig
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   687
  type T
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   688
  val get: Context.generic -> T
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   689
  val put: T -> Context.generic -> Context.generic
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   690
  val map: (T -> T) -> Context.generic -> Context.generic
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   691
end;
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   692
33517
d064fa48f305 modernized/simplified functor Theory_Data, Proof_Data, Generic_Data: eliminated Pretty.pp, discontinued mutable data;
wenzelm
parents: 33033
diff changeset
   693
functor Generic_Data(Data: GENERIC_DATA_ARGS): GENERIC_DATA =
18632
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   694
struct
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   695
33517
d064fa48f305 modernized/simplified functor Theory_Data, Proof_Data, Generic_Data: eliminated Pretty.pp, discontinued mutable data;
wenzelm
parents: 33033
diff changeset
   696
structure Thy_Data = Theory_Data(Data);
d064fa48f305 modernized/simplified functor Theory_Data, Proof_Data, Generic_Data: eliminated Pretty.pp, discontinued mutable data;
wenzelm
parents: 33033
diff changeset
   697
structure Prf_Data = Proof_Data(type T = Data.T val init = Thy_Data.get);
18632
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   698
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   699
type T = Data.T;
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   700
33033
wenzelm
parents: 33031
diff changeset
   701
fun get (Context.Theory thy) = Thy_Data.get thy
wenzelm
parents: 33031
diff changeset
   702
  | get (Context.Proof prf) = Prf_Data.get prf;
18632
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   703
33033
wenzelm
parents: 33031
diff changeset
   704
fun put x (Context.Theory thy) = Context.Theory (Thy_Data.put x thy)
wenzelm
parents: 33031
diff changeset
   705
  | put x (Context.Proof prf) = Context.Proof (Prf_Data.put x prf);
18632
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   706
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   707
fun map f ctxt = put (f (get ctxt)) ctxt;
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   708
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   709
end;
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   710
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   711
(*hide private interface*)
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   712
structure Context: CONTEXT = Context;
20297
a9a917b356af fake predeclaration of type Proof.context;
wenzelm
parents: 19815
diff changeset
   713