src/Pure/context.ML
author wenzelm
Thu, 07 Oct 2021 13:12:08 +0200
changeset 74482 bd5998580edb
parent 74461 8e9f38240c05
child 74561 8e6c973003c8
permissions -rw-r--r--
build minisat, using recent fork from original sources;
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
  exception THEORY of string * theory list
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    18
  structure Proof: sig type context end
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 41711
diff changeset
    19
  structure Proof_Context:
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    20
  sig
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    21
    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
    22
    val init_global: theory -> Proof.context
51686
532e0ac5a66d added ML antiquotation @{theory_context};
wenzelm
parents: 51368
diff changeset
    23
    val get_global: theory -> string -> 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*)
60947
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
    31
  type theory_id
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
    32
  val theory_id: theory -> theory_id
42818
128cc195ced3 timing of Theory_Data operations, with implicit thread positions when functor is applied;
wenzelm
parents: 42383
diff changeset
    33
  val timing: bool Unsynchronized.ref
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    34
  val parents_of: theory -> theory list
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    35
  val ancestors_of: theory -> theory list
70586
57df8a85317a clarified signature;
wenzelm
parents: 70459
diff changeset
    36
  val theory_id_ord: theory_id ord
65459
da59e8e0a663 clarified theory_long_name (for qualified access to Thy_Info) vs. short theory_name (which is unique within any given theory context);
wenzelm
parents: 62889
diff changeset
    37
  val theory_id_long_name: theory_id -> string
60948
b710a5087116 prefer theory_id operations;
wenzelm
parents: 60947
diff changeset
    38
  val theory_id_name: theory_id -> string
65459
da59e8e0a663 clarified theory_long_name (for qualified access to Thy_Info) vs. short theory_name (which is unique within any given theory context);
wenzelm
parents: 62889
diff changeset
    39
  val theory_long_name: theory -> string
29069
c7ba485581ae unified ids for ancestors and checkpoints, removed obsolete history of checkpoints;
wenzelm
parents: 29001
diff changeset
    40
  val theory_name: theory -> string
68482
cb84beb84ca9 clarified signature;
wenzelm
parents: 68193
diff changeset
    41
  val theory_name': {long: bool} -> theory -> string
74461
8e9f38240c05 more exports, notably for Isabelle/Naproche;
wenzelm
parents: 74234
diff changeset
    42
  val theory_identifier: theory -> serial
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
    43
  val PureN: string
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    44
  val pretty_thy: theory -> Pretty.T
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    45
  val pretty_abbrev_thy: theory -> Pretty.T
68482
cb84beb84ca9 clarified signature;
wenzelm
parents: 68193
diff changeset
    46
  val get_theory: {long: bool} -> theory -> string -> theory
60947
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
    47
  val eq_thy_id: theory_id * theory_id -> bool
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    48
  val eq_thy: theory * theory -> bool
60947
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
    49
  val proper_subthy_id: theory_id * theory_id -> bool
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
    50
  val proper_subthy: theory * theory -> bool
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
    51
  val subthy_id: theory_id * theory_id -> bool
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    52
  val subthy: theory * theory -> bool
67623
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
    53
  val trace_theories: bool Unsynchronized.ref
67640
c89270d67169 more informative theories_trace;
wenzelm
parents: 67623
diff changeset
    54
  val theories_trace: unit -> {active_positions: Position.T list, active: int, total: int}
70361
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
    55
  val join_thys: theory * theory -> theory
67623
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
    56
  val begin_thy: string -> theory list -> theory
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
    57
  val finish_thy: theory -> theory
67621
8f93d878f855 auxiliary operation for space profiling;
wenzelm
parents: 66452
diff changeset
    58
  val theory_data_size: theory -> (Position.T * int) list
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
    59
  (*proof context*)
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    60
  val raw_transfer: theory -> Proof.context -> Proof.context
61044
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
    61
  (*certificate*)
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
    62
  datatype certificate = Certificate of theory | Certificate_Id of theory_id
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
    63
  val certificate_theory: certificate -> theory
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
    64
  val certificate_theory_id: certificate -> theory_id
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
    65
  val eq_certificate: certificate * certificate -> bool
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
    66
  val join_certificate: certificate * certificate -> certificate
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
    67
  (*generic context*)
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    68
  datatype generic = Theory of theory | Proof of Proof.context
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    69
  val cases: (theory -> 'a) -> (Proof.context -> 'a) -> generic -> 'a
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    70
  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
    71
  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
    72
    generic -> 'a * generic
18632
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
    73
  val the_theory: generic -> theory
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    74
  val the_proof: generic -> Proof.context
18731
3989c3c41983 rename map_theory/proof to theory/proof_map;
wenzelm
parents: 18711
diff changeset
    75
  val map_theory: (theory -> theory) -> generic -> generic
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    76
  val map_proof: (Proof.context -> Proof.context) -> generic -> generic
26486
b65a5272b360 added map_theory_result, map_proof_result;
wenzelm
parents: 26463
diff changeset
    77
  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
    78
  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
    79
  val theory_map: (generic -> generic) -> theory -> theory
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    80
  val proof_map: (generic -> generic) -> Proof.context -> Proof.context
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    81
  val theory_of: generic -> theory  (*total*)
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
    82
  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
    83
  (*thread data*)
62889
99c7f31615c2 clarified modules;
wenzelm
parents: 62876
diff changeset
    84
  val get_generic_context: unit -> generic option
99c7f31615c2 clarified modules;
wenzelm
parents: 62876
diff changeset
    85
  val put_generic_context: generic option -> unit
99c7f31615c2 clarified modules;
wenzelm
parents: 62876
diff changeset
    86
  val setmp_generic_context: generic option -> ('a -> 'b) -> 'a -> 'b
62876
507c90523113 clarified modules -- simplified bootstrap;
wenzelm
parents: 62819
diff changeset
    87
  val the_generic_context: unit -> generic
507c90523113 clarified modules -- simplified bootstrap;
wenzelm
parents: 62819
diff changeset
    88
  val the_global_context: unit -> theory
507c90523113 clarified modules -- simplified bootstrap;
wenzelm
parents: 62819
diff changeset
    89
  val the_local_context: unit -> Proof.context
26463
9283b4185fdf Context.>> : operate on Context.generic;
wenzelm
parents: 26435
diff changeset
    90
  val >> : (generic -> generic) -> unit
9283b4185fdf Context.>> : operate on Context.generic;
wenzelm
parents: 26435
diff changeset
    91
  val >>> : (generic -> 'a * generic) -> 'a
6185
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
    92
end;
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
    93
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    94
signature PRIVATE_CONTEXT =
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    95
sig
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    96
  include CONTEXT
33033
wenzelm
parents: 33031
diff changeset
    97
  structure Theory_Data:
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
    98
  sig
51368
2ea5c7c2d825 tuned signature -- prefer terminology of Scala and Axiom;
wenzelm
parents: 50431
diff changeset
    99
    val declare: Position.T -> Any.T -> (Any.T -> Any.T) ->
61262
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   100
      (theory * theory -> Any.T * Any.T -> Any.T) -> serial
51368
2ea5c7c2d825 tuned signature -- prefer terminology of Scala and Axiom;
wenzelm
parents: 50431
diff changeset
   101
    val get: serial -> (Any.T -> 'a) -> theory -> 'a
2ea5c7c2d825 tuned signature -- prefer terminology of Scala and Axiom;
wenzelm
parents: 50431
diff changeset
   102
    val put: serial -> ('a -> Any.T) -> 'a -> theory -> theory
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   103
  end
33033
wenzelm
parents: 33031
diff changeset
   104
  structure Proof_Data:
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   105
  sig
51368
2ea5c7c2d825 tuned signature -- prefer terminology of Scala and Axiom;
wenzelm
parents: 50431
diff changeset
   106
    val declare: (theory -> Any.T) -> serial
2ea5c7c2d825 tuned signature -- prefer terminology of Scala and Axiom;
wenzelm
parents: 50431
diff changeset
   107
    val get: serial -> (Any.T -> 'a) -> Proof.context -> 'a
2ea5c7c2d825 tuned signature -- prefer terminology of Scala and Axiom;
wenzelm
parents: 50431
diff changeset
   108
    val put: serial -> ('a -> Any.T) -> 'a -> Proof.context -> Proof.context
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   109
  end
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   110
end;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   111
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   112
structure Context: PRIVATE_CONTEXT =
6185
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
   113
struct
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
   114
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   115
(*** theory context ***)
6185
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
   116
19028
6c238953f66c structure Datatab: private copy avoids potential conflict of table exceptions;
wenzelm
parents: 18931
diff changeset
   117
(*private copy avoids potential conflict of table exceptions*)
31971
8c1b845ed105 renamed functor TableFun to Table, and GraphFun to Graph;
wenzelm
parents: 30628
diff changeset
   118
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
   119
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   120
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   121
(** datatype theory **)
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   122
70360
03430649a7d2 clarified history stage: allow independent updates that are merged later;
wenzelm
parents: 68482
diff changeset
   123
type stage = int * int Synchronized.var;
03430649a7d2 clarified history stage: allow independent updates that are merged later;
wenzelm
parents: 68482
diff changeset
   124
fun init_stage () : stage = (0, Synchronized.var "Context.stage" 1);
03430649a7d2 clarified history stage: allow independent updates that are merged later;
wenzelm
parents: 68482
diff changeset
   125
fun next_stage ((_, state): stage) = (Synchronized.change_result state (fn n => (n, n + 1)), state);
03430649a7d2 clarified history stage: allow independent updates that are merged later;
wenzelm
parents: 68482
diff changeset
   126
70452
70019ab5e57f abstract type theory_id -- ensure non-equality type independently of implementation;
wenzelm
parents: 70361
diff changeset
   127
abstype theory_id =
60947
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   128
  Theory_Id of
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   129
   (*identity*)
52696
38466f4f3483 immutable theory values with full stamp record of every update (increase of stamp size for HOL: 20000 -> 100000, JinjaThreads: 65000 -> 300000) -- minimal measurable impact on inference kernel performance;
wenzelm
parents: 52682
diff changeset
   130
   {id: serial,                   (*identifier*)
38466f4f3483 immutable theory values with full stamp record of every update (increase of stamp size for HOL: 20000 -> 100000, JinjaThreads: 65000 -> 300000) -- minimal measurable impact on inference kernel performance;
wenzelm
parents: 52682
diff changeset
   131
    ids: Inttab.set} *            (*cumulative identifiers -- symbolic body content*)
29095
a75f3ed534a0 tuned comments;
wenzelm
parents: 29093
diff changeset
   132
   (*history*)
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   133
   {name: string,                 (*official theory name*)
70452
70019ab5e57f abstract type theory_id -- ensure non-equality type independently of implementation;
wenzelm
parents: 70361
diff changeset
   134
    stage: stage option}          (*index and counter for anonymous updates*)
70019ab5e57f abstract type theory_id -- ensure non-equality type independently of implementation;
wenzelm
parents: 70361
diff changeset
   135
with
70019ab5e57f abstract type theory_id -- ensure non-equality type independently of implementation;
wenzelm
parents: 70361
diff changeset
   136
70019ab5e57f abstract type theory_id -- ensure non-equality type independently of implementation;
wenzelm
parents: 70361
diff changeset
   137
fun rep_theory_id (Theory_Id args) = args;
70019ab5e57f abstract type theory_id -- ensure non-equality type independently of implementation;
wenzelm
parents: 70361
diff changeset
   138
val make_theory_id = Theory_Id;
70019ab5e57f abstract type theory_id -- ensure non-equality type independently of implementation;
wenzelm
parents: 70361
diff changeset
   139
70019ab5e57f abstract type theory_id -- ensure non-equality type independently of implementation;
wenzelm
parents: 70361
diff changeset
   140
end;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   141
60947
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   142
datatype theory =
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   143
  Theory of
67623
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   144
   (*identity*)
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   145
   {theory_id: theory_id,
67640
c89270d67169 more informative theories_trace;
wenzelm
parents: 67623
diff changeset
   146
    token: Position.T Unsynchronized.ref} *
60947
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   147
   (*ancestry*)
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   148
   {parents: theory list,         (*immediate predecessors*)
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   149
    ancestors: theory list} *     (*all predecessors -- canonical reverse order*)
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   150
   (*data*)
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   151
   Any.T Datatab.table;           (*body content*)
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   152
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   153
exception THEORY of string * theory list;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   154
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   155
fun rep_theory (Theory args) = args;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   156
67623
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   157
val theory_identity = #1 o rep_theory;
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   158
val theory_id = #theory_id o theory_identity;
60947
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   159
val identity_of_id = #1 o rep_theory_id;
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   160
val identity_of = identity_of_id o theory_id;
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   161
val history_of_id = #2 o rep_theory_id;
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   162
val history_of = history_of_id o theory_id;
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   163
val ancestry_of = #2 o rep_theory;
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   164
val data_of = #3 o rep_theory;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   165
52696
38466f4f3483 immutable theory values with full stamp record of every update (increase of stamp size for HOL: 20000 -> 100000, JinjaThreads: 65000 -> 300000) -- minimal measurable impact on inference kernel performance;
wenzelm
parents: 52682
diff changeset
   166
fun make_identity id ids = {id = id, ids = ids};
70360
03430649a7d2 clarified history stage: allow independent updates that are merged later;
wenzelm
parents: 68482
diff changeset
   167
fun make_history name = {name = name, stage = SOME (init_stage ())};
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   168
fun make_ancestry parents ancestors = {parents = parents, ancestors = ancestors};
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   169
70459
f0a445c5a82c maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
wenzelm
parents: 70452
diff changeset
   170
val theory_id_ord = int_ord o apply2 (#id o identity_of_id);
65459
da59e8e0a663 clarified theory_long_name (for qualified access to Thy_Info) vs. short theory_name (which is unique within any given theory context);
wenzelm
parents: 62889
diff changeset
   171
val theory_id_long_name = #name o history_of_id;
da59e8e0a663 clarified theory_long_name (for qualified access to Thy_Info) vs. short theory_name (which is unique within any given theory context);
wenzelm
parents: 62889
diff changeset
   172
val theory_id_name = Long_Name.base_name o theory_id_long_name;
da59e8e0a663 clarified theory_long_name (for qualified access to Thy_Info) vs. short theory_name (which is unique within any given theory context);
wenzelm
parents: 62889
diff changeset
   173
val theory_long_name = #name o history_of;
da59e8e0a663 clarified theory_long_name (for qualified access to Thy_Info) vs. short theory_name (which is unique within any given theory context);
wenzelm
parents: 62889
diff changeset
   174
val theory_name = Long_Name.base_name o theory_long_name;
68482
cb84beb84ca9 clarified signature;
wenzelm
parents: 68193
diff changeset
   175
fun theory_name' {long} = if long then theory_long_name else theory_name;
74461
8e9f38240c05 more exports, notably for Isabelle/Naproche;
wenzelm
parents: 74234
diff changeset
   176
val theory_identifier = #id o identity_of_id o theory_id;
65459
da59e8e0a663 clarified theory_long_name (for qualified access to Thy_Info) vs. short theory_name (which is unique within any given theory context);
wenzelm
parents: 62889
diff changeset
   177
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   178
val parents_of = #parents o ancestry_of;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   179
val ancestors_of = #ancestors o ancestry_of;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   180
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
(* names *)
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   183
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   184
val PureN = "Pure";
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   185
61044
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   186
fun display_name thy_id =
70360
03430649a7d2 clarified history stage: allow independent updates that are merged later;
wenzelm
parents: 68482
diff changeset
   187
  (case history_of_id thy_id of
03430649a7d2 clarified history stage: allow independent updates that are merged later;
wenzelm
parents: 68482
diff changeset
   188
    {name, stage = NONE} => name
03430649a7d2 clarified history stage: allow independent updates that are merged later;
wenzelm
parents: 68482
diff changeset
   189
  | {name, stage = SOME (i, _)} => name ^ ":" ^ string_of_int i);
61044
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   190
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   191
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
   192
  let
61044
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   193
    val name = display_name (theory_id thy);
68192
73a1b393d6f9 more uniform output (cf. 450cefec7c11);
wenzelm
parents: 67640
diff changeset
   194
    val ancestor_names = map theory_long_name (ancestors_of thy);
61044
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   195
  in rev (name :: ancestor_names) end;
29069
c7ba485581ae unified ids for ancestors and checkpoints, removed obsolete history of checkpoints;
wenzelm
parents: 29001
diff changeset
   196
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   197
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
   198
62819
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62663
diff changeset
   199
val _ = ML_system_pp (fn _ => fn _ => Pretty.to_polyml o pretty_thy);
62663
bea354f6ff21 clarified modules;
wenzelm
parents: 61878
diff changeset
   200
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   201
fun pretty_abbrev_thy thy =
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   202
  let
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   203
    val names = display_names thy;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   204
    val n = length names;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   205
    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
   206
  in Pretty.str_list "{" "}" abbrev end;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   207
68482
cb84beb84ca9 clarified signature;
wenzelm
parents: 68193
diff changeset
   208
fun get_theory long thy name =
cb84beb84ca9 clarified signature;
wenzelm
parents: 68193
diff changeset
   209
  if theory_name' long thy <> name then
cb84beb84ca9 clarified signature;
wenzelm
parents: 68193
diff changeset
   210
    (case find_first (fn thy' => theory_name' long thy' = name) (ancestors_of thy) of
37867
b9783e9e96e1 added Context.get_theory -- avoid referring to accidental theory loader state (cf. Thy_Info.get_theory);
wenzelm
parents: 37852
diff changeset
   211
      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
   212
    | NONE => error ("Unknown ancestor theory " ^ quote name))
70360
03430649a7d2 clarified history stage: allow independent updates that are merged later;
wenzelm
parents: 68482
diff changeset
   213
  else if is_none (#stage (history_of thy)) then thy
37867
b9783e9e96e1 added Context.get_theory -- avoid referring to accidental theory loader state (cf. Thy_Info.get_theory);
wenzelm
parents: 37852
diff changeset
   214
  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
   215
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   216
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   217
(* build ids *)
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   218
52696
38466f4f3483 immutable theory values with full stamp record of every update (increase of stamp size for HOL: 20000 -> 100000, JinjaThreads: 65000 -> 300000) -- minimal measurable impact on inference kernel performance;
wenzelm
parents: 52682
diff changeset
   219
fun insert_id id ids = Inttab.update (id, ()) ids;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   220
67623
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   221
val merge_ids =
70452
70019ab5e57f abstract type theory_id -- ensure non-equality type independently of implementation;
wenzelm
parents: 70361
diff changeset
   222
  apply2 (theory_id #> rep_theory_id #> #1) #>
70019ab5e57f abstract type theory_id -- ensure non-equality type independently of implementation;
wenzelm
parents: 70361
diff changeset
   223
  (fn ({id = id1, ids = ids1, ...}, {id = id2, ids = ids2, ...}) =>
67623
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   224
    Inttab.merge (K true) (ids1, ids2)
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   225
    |> insert_id id1
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   226
    |> insert_id id2);
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   227
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   228
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   229
(* equality and inclusion *)
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   230
60947
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   231
val eq_thy_id = op = o apply2 (#id o identity_of_id);
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 55547
diff changeset
   232
val eq_thy = op = o apply2 (#id o identity_of);
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   233
70452
70019ab5e57f abstract type theory_id -- ensure non-equality type independently of implementation;
wenzelm
parents: 70361
diff changeset
   234
val proper_subthy_id =
70019ab5e57f abstract type theory_id -- ensure non-equality type independently of implementation;
wenzelm
parents: 70361
diff changeset
   235
  apply2 (rep_theory_id #> #1) #> (fn ({id, ...}, {ids, ...}) => Inttab.defined ids id);
60947
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   236
val proper_subthy = proper_subthy_id o apply2 theory_id;
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   237
60947
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   238
fun subthy_id p = eq_thy_id p orelse proper_subthy_id p;
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   239
val subthy = subthy_id o apply2 theory_id;
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   240
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   241
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   242
(* consistent ancestors *)
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   243
50431
955c4aa44f60 check consistent theory names for direct imports as well -- as claimed in the comments (see also 1cc36c0ec9eb);
wenzelm
parents: 48992
diff changeset
   244
fun eq_thy_consistent (thy1, thy2) =
955c4aa44f60 check consistent theory names for direct imports as well -- as claimed in the comments (see also 1cc36c0ec9eb);
wenzelm
parents: 48992
diff changeset
   245
  eq_thy (thy1, thy2) orelse
955c4aa44f60 check consistent theory names for direct imports as well -- as claimed in the comments (see also 1cc36c0ec9eb);
wenzelm
parents: 48992
diff changeset
   246
    (theory_name thy1 = theory_name thy2 andalso
955c4aa44f60 check consistent theory names for direct imports as well -- as claimed in the comments (see also 1cc36c0ec9eb);
wenzelm
parents: 48992
diff changeset
   247
      raise THEORY ("Duplicate theory name", [thy1, thy2]));
955c4aa44f60 check consistent theory names for direct imports as well -- as claimed in the comments (see also 1cc36c0ec9eb);
wenzelm
parents: 48992
diff changeset
   248
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   249
fun extend_ancestors thy thys =
50431
955c4aa44f60 check consistent theory names for direct imports as well -- as claimed in the comments (see also 1cc36c0ec9eb);
wenzelm
parents: 48992
diff changeset
   250
  if member eq_thy_consistent thys thy then
33033
wenzelm
parents: 33031
diff changeset
   251
    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
   252
  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
   253
50431
955c4aa44f60 check consistent theory names for direct imports as well -- as claimed in the comments (see also 1cc36c0ec9eb);
wenzelm
parents: 48992
diff changeset
   254
val merge_ancestors = merge eq_thy_consistent;
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   255
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   256
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   257
61262
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   258
(** theory data **)
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   259
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   260
(* data kinds and access methods *)
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   261
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   262
val timing = Unsynchronized.ref false;
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   263
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   264
local
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   265
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   266
type kind =
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   267
 {pos: Position.T,
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   268
  empty: Any.T,
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   269
  extend: Any.T -> Any.T,
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   270
  merge: theory * theory -> Any.T * Any.T -> Any.T};
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   271
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   272
val kinds = Synchronized.var "Theory_Data" (Datatab.empty: kind Datatab.table);
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   273
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   274
fun invoke name f k x =
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   275
  (case Datatab.lookup (Synchronized.value kinds) k of
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   276
    SOME kind =>
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   277
      if ! timing andalso name <> "" then
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   278
        Timing.cond_timeit true ("Theory_Data." ^ name ^ Position.here (#pos kind))
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   279
          (fn () => f kind x)
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   280
      else f kind x
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   281
  | NONE => raise Fail "Invalid theory data identifier");
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   282
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   283
in
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   284
67621
8f93d878f855 auxiliary operation for space profiling;
wenzelm
parents: 66452
diff changeset
   285
fun invoke_pos k = invoke "" (K o #pos) k ();
61262
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   286
fun invoke_empty k = invoke "" (K o #empty) k ();
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   287
val invoke_extend = invoke "extend" #extend;
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   288
fun invoke_merge thys = invoke "merge" (fn kind => #merge kind thys);
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   289
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   290
fun declare_theory_data pos empty extend merge =
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   291
  let
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   292
    val k = serial ();
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   293
    val kind = {pos = pos, empty = empty, extend = extend, merge = merge};
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   294
    val _ = Synchronized.change kinds (Datatab.update (k, kind));
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   295
  in k end;
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   296
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   297
val extend_data = Datatab.map invoke_extend;
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   298
fun merge_data thys = Datatab.join (invoke_merge thys) o apply2 extend_data;
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   299
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   300
end;
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   301
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   302
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   303
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   304
(** build theories **)
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   305
67623
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   306
(* create theory *)
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   307
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   308
val trace_theories = Unsynchronized.ref false;
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   309
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   310
local
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   311
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   312
val theories =
67640
c89270d67169 more informative theories_trace;
wenzelm
parents: 67623
diff changeset
   313
  Synchronized.var "theory_tokens"
c89270d67169 more informative theories_trace;
wenzelm
parents: 67623
diff changeset
   314
    ([]: Position.T Unsynchronized.ref option Unsynchronized.ref list);
67623
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   315
67640
c89270d67169 more informative theories_trace;
wenzelm
parents: 67623
diff changeset
   316
val dummy_token = Unsynchronized.ref Position.none;
67623
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   317
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   318
fun make_token () =
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   319
  if ! trace_theories then
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   320
    let
67640
c89270d67169 more informative theories_trace;
wenzelm
parents: 67623
diff changeset
   321
      val token = Unsynchronized.ref (Position.thread_data ());
67623
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   322
      val _ = Synchronized.change theories (cons (Weak.weak (SOME token)));
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   323
    in token end
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   324
  else dummy_token;
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   325
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   326
in
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   327
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   328
fun theories_trace () =
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   329
  let
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   330
    val trace = Synchronized.value theories;
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   331
    val _ = ML_Heap.full_gc ();
67640
c89270d67169 more informative theories_trace;
wenzelm
parents: 67623
diff changeset
   332
    val active_positions =
c89270d67169 more informative theories_trace;
wenzelm
parents: 67623
diff changeset
   333
      fold (fn Unsynchronized.ref (SOME pos) => cons (! pos) | _ => I) trace [];
c89270d67169 more informative theories_trace;
wenzelm
parents: 67623
diff changeset
   334
  in
c89270d67169 more informative theories_trace;
wenzelm
parents: 67623
diff changeset
   335
    {active_positions = active_positions,
c89270d67169 more informative theories_trace;
wenzelm
parents: 67623
diff changeset
   336
     active = length active_positions,
c89270d67169 more informative theories_trace;
wenzelm
parents: 67623
diff changeset
   337
     total = length trace}
c89270d67169 more informative theories_trace;
wenzelm
parents: 67623
diff changeset
   338
  end;
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   339
60947
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   340
fun create_thy ids history ancestry data =
67623
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   341
  let
70452
70019ab5e57f abstract type theory_id -- ensure non-equality type independently of implementation;
wenzelm
parents: 70361
diff changeset
   342
    val theory_id = make_theory_id (make_identity (serial ()) ids, history);
67623
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   343
    val token = make_token ();
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   344
  in Theory ({theory_id = theory_id, token = token}, ancestry, data) end;
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   345
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   346
end;
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   347
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   348
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   349
(* primitives *)
33606
2b27020ffcb2 local mutex for theory content/identity operations;
wenzelm
parents: 33517
diff changeset
   350
52696
38466f4f3483 immutable theory values with full stamp record of every update (increase of stamp size for HOL: 20000 -> 100000, JinjaThreads: 65000 -> 300000) -- minimal measurable impact on inference kernel performance;
wenzelm
parents: 52682
diff changeset
   351
val pre_pure_thy =
70360
03430649a7d2 clarified history stage: allow independent updates that are merged later;
wenzelm
parents: 68482
diff changeset
   352
  create_thy Inttab.empty (make_history PureN) (make_ancestry [] []) Datatab.empty;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   353
52696
38466f4f3483 immutable theory values with full stamp record of every update (increase of stamp size for HOL: 20000 -> 100000, JinjaThreads: 65000 -> 300000) -- minimal measurable impact on inference kernel performance;
wenzelm
parents: 52682
diff changeset
   354
local
38466f4f3483 immutable theory values with full stamp record of every update (increase of stamp size for HOL: 20000 -> 100000, JinjaThreads: 65000 -> 300000) -- minimal measurable impact on inference kernel performance;
wenzelm
parents: 52682
diff changeset
   355
38466f4f3483 immutable theory values with full stamp record of every update (increase of stamp size for HOL: 20000 -> 100000, JinjaThreads: 65000 -> 300000) -- minimal measurable impact on inference kernel performance;
wenzelm
parents: 52682
diff changeset
   356
fun change_thy finish f thy =
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   357
  let
70452
70019ab5e57f abstract type theory_id -- ensure non-equality type independently of implementation;
wenzelm
parents: 70361
diff changeset
   358
    val ({id, ids}, {name, stage}) = rep_theory_id (theory_id thy);
67623
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   359
    val Theory (_, ancestry, data) = thy;
60947
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   360
    val (ancestry', data') =
70360
03430649a7d2 clarified history stage: allow independent updates that are merged later;
wenzelm
parents: 68482
diff changeset
   361
      if is_none stage then
60947
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   362
        (make_ancestry [thy] (extend_ancestors thy (ancestors_of thy)), extend_data data)
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   363
      else (ancestry, data);
70360
03430649a7d2 clarified history stage: allow independent updates that are merged later;
wenzelm
parents: 68482
diff changeset
   364
    val history' = {name = name, stage = if finish then NONE else Option.map next_stage stage};
52696
38466f4f3483 immutable theory values with full stamp record of every update (increase of stamp size for HOL: 20000 -> 100000, JinjaThreads: 65000 -> 300000) -- minimal measurable impact on inference kernel performance;
wenzelm
parents: 52682
diff changeset
   365
    val ids' = insert_id id ids;
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   366
    val data'' = f data';
60947
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   367
  in create_thy ids' history' ancestry' data'' end;
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   368
52696
38466f4f3483 immutable theory values with full stamp record of every update (increase of stamp size for HOL: 20000 -> 100000, JinjaThreads: 65000 -> 300000) -- minimal measurable impact on inference kernel performance;
wenzelm
parents: 52682
diff changeset
   369
in
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   370
52696
38466f4f3483 immutable theory values with full stamp record of every update (increase of stamp size for HOL: 20000 -> 100000, JinjaThreads: 65000 -> 300000) -- minimal measurable impact on inference kernel performance;
wenzelm
parents: 52682
diff changeset
   371
val update_thy = change_thy false;
38466f4f3483 immutable theory values with full stamp record of every update (increase of stamp size for HOL: 20000 -> 100000, JinjaThreads: 65000 -> 300000) -- minimal measurable impact on inference kernel performance;
wenzelm
parents: 52682
diff changeset
   372
val extend_thy = update_thy I;
68193
14dd78f036ba more thorough checks for theory name consistency (for extend, not just merge);
wenzelm
parents: 68192
diff changeset
   373
val finish_thy = change_thy true I #> tap extend_thy;
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   374
52696
38466f4f3483 immutable theory values with full stamp record of every update (increase of stamp size for HOL: 20000 -> 100000, JinjaThreads: 65000 -> 300000) -- minimal measurable impact on inference kernel performance;
wenzelm
parents: 52682
diff changeset
   375
end;
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   376
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   377
70361
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   378
(* join: anonymous theory nodes *)
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   379
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   380
local
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   381
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   382
fun bad_join (thy1, thy2) = raise THEORY ("Cannot join theories", [thy1, thy2]);
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   383
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   384
fun join_history thys =
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   385
  apply2 history_of thys |>
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   386
  (fn ({name, stage}, {name = name', stage = stage'}) =>
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   387
    if name = name' andalso is_some stage andalso is_some stage' then
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   388
      {name = name, stage = Option.map next_stage stage}
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   389
    else bad_join thys);
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   390
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   391
fun join_ancestry thys =
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   392
  apply2 ancestry_of thys |>
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   393
  (fn (ancestry as {parents, ancestors}, {parents = parents', ancestors = ancestors'}) =>
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   394
    if eq_list eq_thy (parents, parents') andalso eq_list eq_thy (ancestors, ancestors')
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   395
    then ancestry else bad_join thys);
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   396
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   397
in
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   398
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   399
fun join_thys thys =
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   400
  let
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   401
    val ids = merge_ids thys;
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   402
    val history = join_history thys;
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   403
    val ancestry = join_ancestry thys;
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   404
    val data = merge_data thys (apply2 data_of thys);
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   405
  in create_thy ids history ancestry data end;
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   406
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   407
end;
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   408
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   409
34b271c4f400 support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
wenzelm
parents: 70360
diff changeset
   410
(* merge: named theory nodes *)
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   411
61044
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   412
local
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   413
67623
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   414
fun merge_thys thys =
26957
e3f04fdd994d eliminated theory CPure;
wenzelm
parents: 26889
diff changeset
   415
  let
67623
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   416
    val ids = merge_ids thys;
70360
03430649a7d2 clarified history stage: allow independent updates that are merged later;
wenzelm
parents: 68482
diff changeset
   417
    val history = make_history "";
26957
e3f04fdd994d eliminated theory CPure;
wenzelm
parents: 26889
diff changeset
   418
    val ancestry = make_ancestry [] [];
67623
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   419
    val data = merge_data thys (apply2 data_of thys);
60947
d5f7b424ba47 separate type theory_id;
wenzelm
parents: 59163
diff changeset
   420
  in create_thy ids history ancestry data end;
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   421
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   422
fun maximal_thys thys =
28617
wenzelm
parents: 28375
diff changeset
   423
  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
   424
61044
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   425
in
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   426
61262
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   427
fun begin_thy name imports =
52696
38466f4f3483 immutable theory values with full stamp record of every update (increase of stamp size for HOL: 20000 -> 100000, JinjaThreads: 65000 -> 300000) -- minimal measurable impact on inference kernel performance;
wenzelm
parents: 52682
diff changeset
   428
  if name = "" then error ("Bad theory name: " ^ quote name)
24369
0cb1f4d76452 tuned CRITICAL sections;
wenzelm
parents: 24184
diff changeset
   429
  else
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   430
    let
24141
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   431
      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
   432
      val ancestors =
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   433
        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
   434
        |> 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
   435
67623
dee9d2924617 optional trace of created theory values;
wenzelm
parents: 67621
diff changeset
   436
      val thy0 =
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   437
        (case parents of
48638
22d65e375c01 more standard bootstrapping of Pure.thy;
wenzelm
parents: 47005
diff changeset
   438
          [] => error "Missing theory imports"
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   439
        | [thy] => extend_thy thy
61262
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   440
        | thy :: thys => Library.foldl merge_thys (thy, thys));
70452
70019ab5e57f abstract type theory_id -- ensure non-equality type independently of implementation;
wenzelm
parents: 70361
diff changeset
   441
      val ({ids, ...}, _) = rep_theory_id (theory_id thy0);
29093
1cc36c0ec9eb refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
wenzelm
parents: 29069
diff changeset
   442
70360
03430649a7d2 clarified history stage: allow independent updates that are merged later;
wenzelm
parents: 68482
diff changeset
   443
      val history = make_history name;
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   444
      val ancestry = make_ancestry parents ancestors;
68193
14dd78f036ba more thorough checks for theory name consistency (for extend, not just merge);
wenzelm
parents: 68192
diff changeset
   445
    in create_thy ids history ancestry (data_of thy0) |> tap finish_thy end;
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   446
61044
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   447
end;
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   448
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   449
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   450
(* theory data *)
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   451
33033
wenzelm
parents: 33031
diff changeset
   452
structure Theory_Data =
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   453
struct
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   454
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   455
val declare = declare_theory_data;
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   456
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   457
fun get k dest thy =
34253
5930c6391126 removed further remains of mutable theory data (cf. 25bd3ed2ac9f);
wenzelm
parents: 34245
diff changeset
   458
  (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
   459
    SOME x => x
34253
5930c6391126 removed further remains of mutable theory data (cf. 25bd3ed2ac9f);
wenzelm
parents: 34245
diff changeset
   460
  | NONE => invoke_empty k) |> dest;
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   461
52696
38466f4f3483 immutable theory values with full stamp record of every update (increase of stamp size for HOL: 20000 -> 100000, JinjaThreads: 65000 -> 300000) -- minimal measurable impact on inference kernel performance;
wenzelm
parents: 52682
diff changeset
   462
fun put k mk x = update_thy (Datatab.update (k, mk x));
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   463
67621
8f93d878f855 auxiliary operation for space profiling;
wenzelm
parents: 66452
diff changeset
   464
fun obj_size k thy =
8f93d878f855 auxiliary operation for space profiling;
wenzelm
parents: 66452
diff changeset
   465
  Datatab.lookup (data_of thy) k |> Option.map ML_Heap.obj_size;
8f93d878f855 auxiliary operation for space profiling;
wenzelm
parents: 66452
diff changeset
   466
16489
f66ab8a4e98f improved treatment of intermediate checkpoints: actual copy
wenzelm
parents: 16436
diff changeset
   467
end;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   468
67621
8f93d878f855 auxiliary operation for space profiling;
wenzelm
parents: 66452
diff changeset
   469
fun theory_data_size thy =
74234
4f2bd13edce3 clarified signature;
wenzelm
parents: 72060
diff changeset
   470
  build (data_of thy |> Datatab.fold_rev (fn (k, _) =>
67621
8f93d878f855 auxiliary operation for space profiling;
wenzelm
parents: 66452
diff changeset
   471
    (case Theory_Data.obj_size k thy of
8f93d878f855 auxiliary operation for space profiling;
wenzelm
parents: 66452
diff changeset
   472
      NONE => I
74234
4f2bd13edce3 clarified signature;
wenzelm
parents: 72060
diff changeset
   473
    | SOME n => (cons (invoke_pos k, n)))));
67621
8f93d878f855 auxiliary operation for space profiling;
wenzelm
parents: 66452
diff changeset
   474
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   475
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   476
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   477
(*** proof context ***)
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   478
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   479
(* datatype Proof.context *)
17060
cca2f3938443 type proof: theory_ref instead of theory (make proof contexts independent entities);
wenzelm
parents: 16894
diff changeset
   480
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   481
structure Proof =
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   482
struct
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 52696
diff changeset
   483
  datatype context = Context of Any.T Datatab.table * theory;
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   484
end;
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   485
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   486
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   487
(* proof data kinds *)
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   488
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   489
local
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   490
51368
2ea5c7c2d825 tuned signature -- prefer terminology of Scala and Axiom;
wenzelm
parents: 50431
diff changeset
   491
val kinds = Synchronized.var "Proof_Data" (Datatab.empty: (theory -> Any.T) Datatab.table);
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   492
59163
wenzelm
parents: 59146
diff changeset
   493
fun init_data thy =
wenzelm
parents: 59146
diff changeset
   494
  Synchronized.value kinds |> Datatab.map (fn _ => fn init => init thy);
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   495
59163
wenzelm
parents: 59146
diff changeset
   496
fun init_new_data thy =
wenzelm
parents: 59146
diff changeset
   497
  Synchronized.value kinds |> Datatab.fold (fn (k, init) => fn data =>
wenzelm
parents: 59146
diff changeset
   498
    if Datatab.defined data k then data
wenzelm
parents: 59146
diff changeset
   499
    else Datatab.update (k, init thy) data);
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
   500
59163
wenzelm
parents: 59146
diff changeset
   501
fun init_fallback k thy =
wenzelm
parents: 59146
diff changeset
   502
  (case Datatab.lookup (Synchronized.value kinds) k of
wenzelm
parents: 59146
diff changeset
   503
    SOME init => init thy
wenzelm
parents: 59146
diff changeset
   504
  | NONE => raise Fail "Invalid proof data identifier");
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   505
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   506
in
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   507
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 52696
diff changeset
   508
fun raw_transfer thy' (Proof.Context (data, thy)) =
24141
73baca986087 improved check_thy: produce a checked theory_ref (thread-safe version);
wenzelm
parents: 23944
diff changeset
   509
  let
52682
wenzelm
parents: 51686
diff changeset
   510
    val _ = subthy (thy, thy') orelse error "Cannot transfer proof context: not a super theory";
59163
wenzelm
parents: 59146
diff changeset
   511
    val data' = init_new_data thy' data;
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 52696
diff changeset
   512
  in Proof.Context (data', thy') 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
   513
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 41711
diff changeset
   514
structure Proof_Context =
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   515
struct
59163
wenzelm
parents: 59146
diff changeset
   516
  fun theory_of (Proof.Context (_, thy)) = thy;
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 52696
diff changeset
   517
  fun init_global thy = Proof.Context (init_data thy, thy);
68482
cb84beb84ca9 clarified signature;
wenzelm
parents: 68193
diff changeset
   518
  fun get_global thy name = init_global (get_theory {long = false} thy name);
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   519
end;
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   520
33033
wenzelm
parents: 33031
diff changeset
   521
structure Proof_Data =
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   522
struct
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   523
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
   524
fun declare init =
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   525
  let
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   526
    val k = serial ();
43684
85388f5570c4 prefer Synchronized.var;
wenzelm
parents: 43610
diff changeset
   527
    val _ = Synchronized.change kinds (Datatab.update (k, init));
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   528
  in k end;
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   529
59163
wenzelm
parents: 59146
diff changeset
   530
fun get k dest (Proof.Context (data, thy)) =
wenzelm
parents: 59146
diff changeset
   531
  (case Datatab.lookup data 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
   532
    SOME x => x
59163
wenzelm
parents: 59146
diff changeset
   533
  | NONE => init_fallback k thy) |> dest;
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   534
59163
wenzelm
parents: 59146
diff changeset
   535
fun put k mk x (Proof.Context (data, thy)) =
wenzelm
parents: 59146
diff changeset
   536
  Proof.Context (Datatab.update (k, mk x) data, thy);
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   537
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   538
end;
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   539
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   540
end;
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   541
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   542
18632
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   543
61044
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   544
(*** theory certificate ***)
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   545
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   546
datatype certificate = Certificate of theory | Certificate_Id of theory_id;
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   547
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   548
fun certificate_theory (Certificate thy) = thy
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   549
  | certificate_theory (Certificate_Id thy_id) =
61062
52f3256a6d85 tuned message;
wenzelm
parents: 61050
diff changeset
   550
      error ("No content for theory certificate " ^ display_name thy_id);
61044
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   551
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   552
fun certificate_theory_id (Certificate thy) = theory_id thy
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   553
  | certificate_theory_id (Certificate_Id thy_id) = thy_id;
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   554
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   555
fun eq_certificate (Certificate thy1, Certificate thy2) = eq_thy (thy1, thy2)
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   556
  | eq_certificate (Certificate_Id thy_id1, Certificate_Id thy_id2) = eq_thy_id (thy_id1, thy_id2)
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   557
  | eq_certificate _ = false;
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   558
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   559
fun join_certificate (cert1, cert2) =
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   560
  let val (thy_id1, thy_id2) = apply2 certificate_theory_id (cert1, cert2) in
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   561
    if eq_thy_id (thy_id1, thy_id2) then (case cert1 of Certificate _ => cert1 | _ => cert2)
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   562
    else if proper_subthy_id (thy_id2, thy_id1) then cert1
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   563
    else if proper_subthy_id (thy_id1, thy_id2) then cert2
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   564
    else
61045
c7a7f063704a clarified exceptions: avoid interference of formal context failure with regular rule application failure (which is routinely handled in user-space);
wenzelm
parents: 61044
diff changeset
   565
      error ("Cannot join unrelated theory certificates " ^
61062
52f3256a6d85 tuned message;
wenzelm
parents: 61050
diff changeset
   566
        display_name thy_id1 ^ " and " ^ display_name thy_id2)
61044
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   567
  end;
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   568
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   569
b7af255dd200 more abstract theory certificate, which is not necessarily the full theory;
wenzelm
parents: 60948
diff changeset
   570
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   571
(*** generic context ***)
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   572
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   573
datatype generic = Theory of theory | Proof of Proof.context;
18632
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   574
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   575
fun cases f _ (Theory thy) = f thy
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   576
  | cases _ g (Proof prf) = g prf;
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   577
19678
d1a15431de34 added mapping;
wenzelm
parents: 19482
diff changeset
   578
fun mapping f g = cases (Theory o f) (Proof o g);
21660
c86b761d6c06 added mapping_result;
wenzelm
parents: 21518
diff changeset
   579
fun mapping_result f g = cases (apsnd Theory o f) (apsnd Proof o g);
19678
d1a15431de34 added mapping;
wenzelm
parents: 19482
diff changeset
   580
23595
7ca68a2c8575 the_theory/proof: error instead of exception Fail;
wenzelm
parents: 23355
diff changeset
   581
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
   582
val the_proof = cases (fn _ => error "Ill-typed context: proof expected") I;
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   583
18731
3989c3c41983 rename map_theory/proof to theory/proof_map;
wenzelm
parents: 18711
diff changeset
   584
fun map_theory f = Theory o f o the_theory;
3989c3c41983 rename map_theory/proof to theory/proof_map;
wenzelm
parents: 18711
diff changeset
   585
fun map_proof f = Proof o f o the_proof;
3989c3c41983 rename map_theory/proof to theory/proof_map;
wenzelm
parents: 18711
diff changeset
   586
26486
b65a5272b360 added map_theory_result, map_proof_result;
wenzelm
parents: 26463
diff changeset
   587
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
   588
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
   589
18731
3989c3c41983 rename map_theory/proof to theory/proof_map;
wenzelm
parents: 18711
diff changeset
   590
fun theory_map f = the_theory o f o Theory;
3989c3c41983 rename map_theory/proof to theory/proof_map;
wenzelm
parents: 18711
diff changeset
   591
fun proof_map f = the_proof o f o Proof;
18665
5198a2c4c00e added map_theory, map_proof;
wenzelm
parents: 18632
diff changeset
   592
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 41711
diff changeset
   593
val theory_of = cases I Proof_Context.theory_of;
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 41711
diff changeset
   594
val proof_of = cases Proof_Context.init_global I;
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   595
22085
c138cfd500f7 ML context: full generic context, tuned signature;
wenzelm
parents: 21962
diff changeset
   596
c138cfd500f7 ML context: full generic context, tuned signature;
wenzelm
parents: 21962
diff changeset
   597
26413
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   598
(** thread data **)
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   599
62889
99c7f31615c2 clarified modules;
wenzelm
parents: 62876
diff changeset
   600
local val generic_context_var = Thread_Data.var () : generic Thread_Data.var in
26413
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   601
62889
99c7f31615c2 clarified modules;
wenzelm
parents: 62876
diff changeset
   602
fun get_generic_context () = Thread_Data.get generic_context_var;
99c7f31615c2 clarified modules;
wenzelm
parents: 62876
diff changeset
   603
val put_generic_context = Thread_Data.put generic_context_var;
99c7f31615c2 clarified modules;
wenzelm
parents: 62876
diff changeset
   604
fun setmp_generic_context opt_context = Thread_Data.setmp generic_context_var opt_context;
62876
507c90523113 clarified modules -- simplified bootstrap;
wenzelm
parents: 62819
diff changeset
   605
507c90523113 clarified modules -- simplified bootstrap;
wenzelm
parents: 62819
diff changeset
   606
fun the_generic_context () =
62889
99c7f31615c2 clarified modules;
wenzelm
parents: 62876
diff changeset
   607
  (case get_generic_context () of
26413
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   608
    SOME context => context
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   609
  | _ => error "Unknown context");
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   610
62876
507c90523113 clarified modules -- simplified bootstrap;
wenzelm
parents: 62819
diff changeset
   611
val the_global_context = theory_of o the_generic_context;
507c90523113 clarified modules -- simplified bootstrap;
wenzelm
parents: 62819
diff changeset
   612
val the_local_context = proof_of o the_generic_context;
26413
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   613
26428
5b2beca2087d eliminated theory ProtoPure;
wenzelm
parents: 26421
diff changeset
   614
end;
5b2beca2087d eliminated theory ProtoPure;
wenzelm
parents: 26421
diff changeset
   615
5b2beca2087d eliminated theory ProtoPure;
wenzelm
parents: 26421
diff changeset
   616
fun >>> f =
5b2beca2087d eliminated theory ProtoPure;
wenzelm
parents: 26421
diff changeset
   617
  let
62876
507c90523113 clarified modules -- simplified bootstrap;
wenzelm
parents: 62819
diff changeset
   618
    val (res, context') = f (the_generic_context ());
62889
99c7f31615c2 clarified modules;
wenzelm
parents: 62876
diff changeset
   619
    val _ = put_generic_context (SOME context');
26428
5b2beca2087d eliminated theory ProtoPure;
wenzelm
parents: 26421
diff changeset
   620
  in res end;
5b2beca2087d eliminated theory ProtoPure;
wenzelm
parents: 26421
diff changeset
   621
26421
3dfb36923a56 nonfix >>;
wenzelm
parents: 26413
diff changeset
   622
nonfix >>;
26463
9283b4185fdf Context.>> : operate on Context.generic;
wenzelm
parents: 26435
diff changeset
   623
fun >> f = >>> (fn context => ((), f context));
26413
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   624
62889
99c7f31615c2 clarified modules;
wenzelm
parents: 62876
diff changeset
   625
val _ = put_generic_context (SOME (Theory pre_pure_thy));
26413
003dd6155870 added thread data (formerly global ref in ML/ml_context.ML);
wenzelm
parents: 24559
diff changeset
   626
6185
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
   627
end;
11bf7a8b6a02 Global theory context (used to be in Thy/context.ML);
wenzelm
parents:
diff changeset
   628
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   629
structure Basic_Context: BASIC_CONTEXT = Context;
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   630
open Basic_Context;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   631
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   632
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   633
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   634
(*** type-safe interfaces for data declarations ***)
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   635
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   636
(** theory data **)
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   637
61262
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   638
signature THEORY_DATA'_ARGS =
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   639
sig
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   640
  type T
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   641
  val empty: T
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   642
  val extend: T -> T
61262
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   643
  val merge: theory * theory -> T * T -> T
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   644
end;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   645
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
   646
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
   647
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
   648
  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
   649
  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
   650
  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
   651
  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
   652
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
   653
34253
5930c6391126 removed further remains of mutable theory data (cf. 25bd3ed2ac9f);
wenzelm
parents: 34245
diff changeset
   654
signature THEORY_DATA =
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   655
sig
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   656
  type T
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   657
  val get: theory -> T
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   658
  val put: T -> theory -> theory
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   659
  val map: (T -> T) -> theory -> theory
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   660
end;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   661
61262
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   662
functor Theory_Data'(Data: THEORY_DATA'_ARGS): THEORY_DATA =
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   663
struct
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   664
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   665
type T = Data.T;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   666
exception Data of T;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   667
42818
128cc195ced3 timing of Theory_Data operations, with implicit thread positions when functor is applied;
wenzelm
parents: 42383
diff changeset
   668
val kind =
72060
efb7fd4a6d1f subtle change of Theory_Data extend/merge semantics due to Theory.join_theory;
wenzelm
parents: 70586
diff changeset
   669
  let val pos = Position.thread_data () in
efb7fd4a6d1f subtle change of Theory_Data extend/merge semantics due to Theory.join_theory;
wenzelm
parents: 70586
diff changeset
   670
    Context.Theory_Data.declare
efb7fd4a6d1f subtle change of Theory_Data extend/merge semantics due to Theory.join_theory;
wenzelm
parents: 70586
diff changeset
   671
      pos
efb7fd4a6d1f subtle change of Theory_Data extend/merge semantics due to Theory.join_theory;
wenzelm
parents: 70586
diff changeset
   672
      (Data Data.empty)
efb7fd4a6d1f subtle change of Theory_Data extend/merge semantics due to Theory.join_theory;
wenzelm
parents: 70586
diff changeset
   673
      (fn Data x =>
efb7fd4a6d1f subtle change of Theory_Data extend/merge semantics due to Theory.join_theory;
wenzelm
parents: 70586
diff changeset
   674
        let
efb7fd4a6d1f subtle change of Theory_Data extend/merge semantics due to Theory.join_theory;
wenzelm
parents: 70586
diff changeset
   675
          val y = Data.extend x;
efb7fd4a6d1f subtle change of Theory_Data extend/merge semantics due to Theory.join_theory;
wenzelm
parents: 70586
diff changeset
   676
          val _ =
efb7fd4a6d1f subtle change of Theory_Data extend/merge semantics due to Theory.join_theory;
wenzelm
parents: 70586
diff changeset
   677
            if pointer_eq (x, y) then ()
efb7fd4a6d1f subtle change of Theory_Data extend/merge semantics due to Theory.join_theory;
wenzelm
parents: 70586
diff changeset
   678
            else raise Fail ("Theory_Data extend needs to be identity" ^ Position.here pos)
efb7fd4a6d1f subtle change of Theory_Data extend/merge semantics due to Theory.join_theory;
wenzelm
parents: 70586
diff changeset
   679
        in Data y end)
efb7fd4a6d1f subtle change of Theory_Data extend/merge semantics due to Theory.join_theory;
wenzelm
parents: 70586
diff changeset
   680
      (fn thys => fn (Data x1, Data x2) => Data (Data.merge thys (x1, x2)))
efb7fd4a6d1f subtle change of Theory_Data extend/merge semantics due to Theory.join_theory;
wenzelm
parents: 70586
diff changeset
   681
  end;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   682
33033
wenzelm
parents: 33031
diff changeset
   683
val get = Context.Theory_Data.get kind (fn Data x => x);
wenzelm
parents: 33031
diff changeset
   684
val put = Context.Theory_Data.put kind Data;
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   685
fun map f thy = put (f (get thy)) thy;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   686
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   687
end;
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   688
33517
d064fa48f305 modernized/simplified functor Theory_Data, Proof_Data, Generic_Data: eliminated Pretty.pp, discontinued mutable data;
wenzelm
parents: 33033
diff changeset
   689
functor Theory_Data(Data: THEORY_DATA_ARGS): THEORY_DATA =
61262
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61062
diff changeset
   690
  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
   691
  (
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
   692
    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
   693
    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
   694
    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
   695
    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
   696
  );
33517
d064fa48f305 modernized/simplified functor Theory_Data, Proof_Data, Generic_Data: eliminated Pretty.pp, discontinued mutable data;
wenzelm
parents: 33033
diff changeset
   697
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   698
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   699
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   700
(** proof data **)
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   701
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   702
signature PROOF_DATA_ARGS =
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   703
sig
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   704
  type T
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   705
  val init: theory -> T
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   706
end;
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   707
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   708
signature PROOF_DATA =
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   709
sig
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   710
  type T
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   711
  val get: Proof.context -> T
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   712
  val put: T -> Proof.context -> Proof.context
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32784
diff changeset
   713
  val map: (T -> T) -> Proof.context -> Proof.context
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   714
end;
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   715
33517
d064fa48f305 modernized/simplified functor Theory_Data, Proof_Data, Generic_Data: eliminated Pretty.pp, discontinued mutable data;
wenzelm
parents: 33033
diff changeset
   716
functor Proof_Data(Data: PROOF_DATA_ARGS): PROOF_DATA =
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   717
struct
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   718
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   719
type T = Data.T;
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   720
exception Data of T;
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   721
33033
wenzelm
parents: 33031
diff changeset
   722
val kind = Context.Proof_Data.declare (Data o Data.init);
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   723
33033
wenzelm
parents: 33031
diff changeset
   724
val get = Context.Proof_Data.get kind (fn Data x => x);
wenzelm
parents: 33031
diff changeset
   725
val put = Context.Proof_Data.put kind Data;
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   726
fun map f prf = put (f (get prf)) prf;
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   727
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   728
end;
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   729
18632
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   730
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   731
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   732
(** generic data **)
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   733
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   734
signature GENERIC_DATA_ARGS =
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   735
sig
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   736
  type T
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   737
  val empty: T
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   738
  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
   739
  val merge: T * T -> T
18632
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   740
end;
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   741
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   742
signature GENERIC_DATA =
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   743
sig
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   744
  type T
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   745
  val get: Context.generic -> T
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   746
  val put: T -> Context.generic -> Context.generic
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   747
  val map: (T -> T) -> Context.generic -> Context.generic
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   748
end;
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   749
33517
d064fa48f305 modernized/simplified functor Theory_Data, Proof_Data, Generic_Data: eliminated Pretty.pp, discontinued mutable data;
wenzelm
parents: 33033
diff changeset
   750
functor Generic_Data(Data: GENERIC_DATA_ARGS): GENERIC_DATA =
18632
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   751
struct
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   752
33517
d064fa48f305 modernized/simplified functor Theory_Data, Proof_Data, Generic_Data: eliminated Pretty.pp, discontinued mutable data;
wenzelm
parents: 33033
diff changeset
   753
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
   754
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
   755
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   756
type T = Data.T;
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   757
33033
wenzelm
parents: 33031
diff changeset
   758
fun get (Context.Theory thy) = Thy_Data.get thy
wenzelm
parents: 33031
diff changeset
   759
  | get (Context.Proof prf) = Prf_Data.get prf;
18632
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   760
33033
wenzelm
parents: 33031
diff changeset
   761
fun put x (Context.Theory thy) = Context.Theory (Thy_Data.put x thy)
wenzelm
parents: 33031
diff changeset
   762
  | 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
   763
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   764
fun map f ctxt = put (f (get ctxt)) ctxt;
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   765
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   766
end;
3149c6abe876 support for generic contexts with data;
wenzelm
parents: 17756
diff changeset
   767
16533
f1152f75f6fc begin_thy: merge maximal imports;
wenzelm
parents: 16489
diff changeset
   768
(*hide private interface*)
16436
7eb6b6cbd166 added type theory: generic theory contexts with unique identity,
wenzelm
parents: 15801
diff changeset
   769
structure Context: CONTEXT = Context;