| author | wenzelm | 
| Mon, 04 Nov 2024 20:55:01 +0100 | |
| changeset 81340 | 30f7eb65d679 | 
| parent 80809 | 4a64fc4d1cde | 
| child 83014 | fa6c4ad21a24 | 
| permissions | -rw-r--r-- | 
| 6185 | 1 | (* Title: Pure/context.ML | 
| 2 | Author: Markus Wenzel, TU Muenchen | |
| 3 | ||
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
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: 
23944diff
changeset | 5 | monotonic development graph and history support. Generic proof | 
| 
73baca986087
improved check_thy: produce a checked theory_ref (thread-safe version);
 wenzelm parents: 
23944diff
changeset | 6 | contexts with arbitrarily typed data. | 
| 33031 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 7 | |
| 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 8 | Firm naming conventions: | 
| 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 9 | thy, thy', thy1, thy2: theory | 
| 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 10 | ctxt, ctxt', ctxt1, ctxt2: Proof.context | 
| 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 11 | context: Context.generic | 
| 6185 | 12 | *) | 
| 13 | ||
| 14 | signature BASIC_CONTEXT = | |
| 15 | sig | |
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 16 | type theory | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 17 | exception THEORY of string * theory list | 
| 33031 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 18 | structure Proof: sig type context end | 
| 42360 | 19 | structure Proof_Context: | 
| 33031 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 20 | sig | 
| 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
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: 
34259diff
changeset | 22 | val init_global: theory -> Proof.context | 
| 77889 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
77886diff
changeset | 23 |     val get_global: {long: bool} -> theory -> string -> Proof.context
 | 
| 33031 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 24 | end | 
| 6185 | 25 | end; | 
| 26 | ||
| 27 | signature CONTEXT = | |
| 28 | sig | |
| 29 | include BASIC_CONTEXT | |
| 77897 | 30 | (*theory data*) | 
| 31 | type data_kind = int | |
| 32 | val data_kinds: unit -> (data_kind * Position.T) list | |
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 33 | (*theory context*) | 
| 77897 | 34 | type id = int | 
| 60947 | 35 | type theory_id | 
| 36 | val theory_id: theory -> theory_id | |
| 78018 
dfa44d85d751
proper system options to control context tracing/timing;
 wenzelm parents: 
78004diff
changeset | 37 | val data_timing: bool Unsynchronized.ref | 
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 38 | val parents_of: theory -> theory list | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 39 | val ancestors_of: theory -> theory list | 
| 70586 | 40 | val theory_id_ord: theory_id ord | 
| 77889 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
77886diff
changeset | 41 |   val theory_id_name: {long: bool} -> 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: 
62889diff
changeset | 42 | val theory_long_name: theory -> string | 
| 77889 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
77886diff
changeset | 43 | val theory_base_name: theory -> string | 
| 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
77886diff
changeset | 44 |   val theory_name: {long: bool} -> theory -> string
 | 
| 77897 | 45 | val theory_identifier: theory -> id | 
| 29093 
1cc36c0ec9eb
refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
 wenzelm parents: 
29069diff
changeset | 46 | val PureN: string | 
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 47 | val pretty_thy: theory -> Pretty.T | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 48 | val pretty_abbrev_thy: theory -> Pretty.T | 
| 68482 | 49 |   val get_theory: {long: bool} -> theory -> string -> theory
 | 
| 60947 | 50 | val eq_thy_id: theory_id * theory_id -> bool | 
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 51 | val eq_thy: theory * theory -> bool | 
| 60947 | 52 | val proper_subthy_id: theory_id * theory_id -> bool | 
| 53 | val proper_subthy: theory * theory -> bool | |
| 54 | val subthy_id: theory_id * theory_id -> bool | |
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 55 | val subthy: theory * theory -> bool | 
| 77895 | 56 | val join_thys: theory list -> theory | 
| 67623 | 57 | val begin_thy: string -> theory list -> theory | 
| 16489 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 58 | val finish_thy: theory -> theory | 
| 77766 | 59 | val theory_data_sizeof1: theory -> (Position.T * int) list | 
| 16533 | 60 | (*proof context*) | 
| 33031 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 61 | val raw_transfer: theory -> Proof.context -> Proof.context | 
| 61044 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 62 | (*certificate*) | 
| 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 63 | datatype certificate = Certificate of theory | Certificate_Id of theory_id | 
| 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 64 | val certificate_theory: certificate -> theory | 
| 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 65 | val certificate_theory_id: certificate -> theory_id | 
| 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 66 | val eq_certificate: certificate * certificate -> bool | 
| 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 67 | val join_certificate: certificate * certificate -> certificate | 
| 78062 
edb195122938
support for context within morphism (for background theory);
 wenzelm parents: 
78035diff
changeset | 68 | val join_certificate_theory: theory * theory -> theory | 
| 16533 | 69 | (*generic context*) | 
| 33031 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 70 | datatype generic = Theory of theory | Proof of Proof.context | 
| 78018 
dfa44d85d751
proper system options to control context tracing/timing;
 wenzelm parents: 
78004diff
changeset | 71 | val theory_tracing: bool Unsynchronized.ref | 
| 
dfa44d85d751
proper system options to control context tracing/timing;
 wenzelm parents: 
78004diff
changeset | 72 | val proof_tracing: bool Unsynchronized.ref | 
| 78035 | 73 | val enabled_tracing: unit -> bool | 
| 78020 | 74 | val finish_tracing: unit -> | 
| 78004 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 75 |    {contexts: (generic * Position.T) list,
 | 
| 77996 | 76 | active_contexts: int, | 
| 77 | active_theories: int, | |
| 78 | active_proofs: int, | |
| 79 | total_contexts: int, | |
| 80 | total_theories: int, | |
| 81 | total_proofs: int} | |
| 33031 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 82 | val cases: (theory -> 'a) -> (Proof.context -> 'a) -> generic -> 'a | 
| 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 83 | val mapping: (theory -> theory) -> (Proof.context -> Proof.context) -> generic -> generic | 
| 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 84 | val mapping_result: (theory -> 'a * theory) -> (Proof.context -> 'a * Proof.context) -> | 
| 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 85 | generic -> 'a * generic | 
| 18632 | 86 | val the_theory: generic -> theory | 
| 33031 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 87 | val the_proof: generic -> Proof.context | 
| 18731 | 88 | val map_theory: (theory -> theory) -> generic -> generic | 
| 33031 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 89 | val map_proof: (Proof.context -> Proof.context) -> generic -> generic | 
| 26486 | 90 | val map_theory_result: (theory -> 'a * theory) -> generic -> 'a * generic | 
| 33031 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 91 | val map_proof_result: (Proof.context -> 'a * Proof.context) -> generic -> 'a * generic | 
| 18731 | 92 | val theory_map: (generic -> generic) -> theory -> theory | 
| 33031 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 93 | val proof_map: (generic -> generic) -> Proof.context -> Proof.context | 
| 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 94 | val theory_of: generic -> theory (*total*) | 
| 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 95 | val proof_of: generic -> Proof.context (*total*) | 
| 26413 
003dd6155870
added thread data (formerly global ref in ML/ml_context.ML);
 wenzelm parents: 
24559diff
changeset | 96 | (*thread data*) | 
| 62889 | 97 | val get_generic_context: unit -> generic option | 
| 98 | val put_generic_context: generic option -> unit | |
| 99 |   val setmp_generic_context: generic option -> ('a -> 'b) -> 'a -> 'b
 | |
| 62876 | 100 | val the_generic_context: unit -> generic | 
| 101 | val the_global_context: unit -> theory | |
| 102 | val the_local_context: unit -> Proof.context | |
| 26463 | 103 | val >> : (generic -> generic) -> unit | 
| 104 | val >>> : (generic -> 'a * generic) -> 'a | |
| 6185 | 105 | end; | 
| 106 | ||
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 107 | signature PRIVATE_CONTEXT = | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 108 | sig | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 109 | include CONTEXT | 
| 33033 | 110 | structure Theory_Data: | 
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 111 | sig | 
| 77897 | 112 | val declare: Position.T -> Any.T -> ((theory * Any.T) list -> Any.T) -> data_kind | 
| 113 | val get: data_kind -> (Any.T -> 'a) -> theory -> 'a | |
| 114 |     val put: data_kind -> ('a -> Any.T) -> 'a -> theory -> theory
 | |
| 16489 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 115 | end | 
| 33033 | 116 | structure Proof_Data: | 
| 16533 | 117 | sig | 
| 77897 | 118 | val declare: (theory -> Any.T) -> data_kind | 
| 119 | val get: data_kind -> (Any.T -> 'a) -> Proof.context -> 'a | |
| 120 |     val put: data_kind -> ('a -> Any.T) -> 'a -> Proof.context -> Proof.context
 | |
| 16533 | 121 | end | 
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 122 | end; | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 123 | |
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 124 | structure Context: PRIVATE_CONTEXT = | 
| 6185 | 125 | struct | 
| 126 | ||
| 77994 | 127 | (*** type definitions ***) | 
| 6185 | 128 | |
| 77994 | 129 | (* context data *) | 
| 77897 | 130 | |
| 19028 
6c238953f66c
structure Datatab: private copy avoids potential conflict of table exceptions;
 wenzelm parents: 
18931diff
changeset | 131 | (*private copy avoids potential conflict of table exceptions*) | 
| 31971 
8c1b845ed105
renamed functor TableFun to Table, and GraphFun to Graph;
 wenzelm parents: 
30628diff
changeset | 132 | structure Datatab = Table(type key = int val ord = int_ord); | 
| 19028 
6c238953f66c
structure Datatab: private copy avoids potential conflict of table exceptions;
 wenzelm parents: 
18931diff
changeset | 133 | |
| 77897 | 134 | type data_kind = int; | 
| 135 | val data_kind = Counter.make (); | |
| 136 | ||
| 16489 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 137 | |
| 77994 | 138 | (* theory identity *) | 
| 139 | ||
| 140 | type id = int; | |
| 79081 
9d6359b71264
more compact representation of theory_id -- via consecutive thread-local ids;
 wenzelm parents: 
79076diff
changeset | 141 | |
| 
9d6359b71264
more compact representation of theory_id -- via consecutive thread-local ids;
 wenzelm parents: 
79076diff
changeset | 142 | local | 
| 
9d6359b71264
more compact representation of theory_id -- via consecutive thread-local ids;
 wenzelm parents: 
79076diff
changeset | 143 | val new_block = Counter.make (); | 
| 
9d6359b71264
more compact representation of theory_id -- via consecutive thread-local ids;
 wenzelm parents: 
79076diff
changeset | 144 | fun new_elem () = Bitset.make_elem (new_block (), 0); | 
| 
9d6359b71264
more compact representation of theory_id -- via consecutive thread-local ids;
 wenzelm parents: 
79076diff
changeset | 145 | val var = Thread_Data.var () : id Thread_Data.var; | 
| 
9d6359b71264
more compact representation of theory_id -- via consecutive thread-local ids;
 wenzelm parents: 
79076diff
changeset | 146 | in | 
| 
9d6359b71264
more compact representation of theory_id -- via consecutive thread-local ids;
 wenzelm parents: 
79076diff
changeset | 147 | |
| 
9d6359b71264
more compact representation of theory_id -- via consecutive thread-local ids;
 wenzelm parents: 
79076diff
changeset | 148 | fun new_id () = | 
| 
9d6359b71264
more compact representation of theory_id -- via consecutive thread-local ids;
 wenzelm parents: 
79076diff
changeset | 149 | let | 
| 
9d6359b71264
more compact representation of theory_id -- via consecutive thread-local ids;
 wenzelm parents: 
79076diff
changeset | 150 | val id = | 
| 
9d6359b71264
more compact representation of theory_id -- via consecutive thread-local ids;
 wenzelm parents: 
79076diff
changeset | 151 | (case Option.map Bitset.dest_elem (Thread_Data.get var) of | 
| 
9d6359b71264
more compact representation of theory_id -- via consecutive thread-local ids;
 wenzelm parents: 
79076diff
changeset | 152 | NONE => new_elem () | 
| 
9d6359b71264
more compact representation of theory_id -- via consecutive thread-local ids;
 wenzelm parents: 
79076diff
changeset | 153 | | SOME (m, n) => | 
| 
9d6359b71264
more compact representation of theory_id -- via consecutive thread-local ids;
 wenzelm parents: 
79076diff
changeset | 154 | (case try Bitset.make_elem (m, n + 1) of | 
| 
9d6359b71264
more compact representation of theory_id -- via consecutive thread-local ids;
 wenzelm parents: 
79076diff
changeset | 155 | NONE => new_elem () | 
| 
9d6359b71264
more compact representation of theory_id -- via consecutive thread-local ids;
 wenzelm parents: 
79076diff
changeset | 156 | | SOME elem => elem)); | 
| 
9d6359b71264
more compact representation of theory_id -- via consecutive thread-local ids;
 wenzelm parents: 
79076diff
changeset | 157 | val _ = Thread_Data.put var (SOME id); | 
| 
9d6359b71264
more compact representation of theory_id -- via consecutive thread-local ids;
 wenzelm parents: 
79076diff
changeset | 158 | in id end; | 
| 
9d6359b71264
more compact representation of theory_id -- via consecutive thread-local ids;
 wenzelm parents: 
79076diff
changeset | 159 | |
| 
9d6359b71264
more compact representation of theory_id -- via consecutive thread-local ids;
 wenzelm parents: 
79076diff
changeset | 160 | end; | 
| 16489 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 161 | |
| 77994 | 162 | abstype theory_id = | 
| 163 | Thy_Id of | |
| 164 |    {id: id,         (*identifier*)
 | |
| 79076 | 165 | ids: Bitset.T, (*cumulative identifiers -- symbolic body content*) | 
| 77994 | 166 | name: string, (*official theory name*) | 
| 167 | stage: int} (*index for anonymous updates*) | |
| 168 | with | |
| 169 | fun rep_theory_id (Thy_Id args) = args; | |
| 170 | val make_theory_id = Thy_Id; | |
| 171 | end; | |
| 172 | ||
| 173 | ||
| 174 | (* theory allocation state *) | |
| 77895 | 175 | |
| 176 | type state = {stage: int} Synchronized.var;
 | |
| 177 | ||
| 178 | fun make_state () : state = | |
| 179 |   Synchronized.var "Context.state" {stage = 0};
 | |
| 180 | ||
| 181 | fun next_stage (state: state) = | |
| 182 |   Synchronized.change_result state (fn {stage} => (stage + 1, {stage = stage + 1}));
 | |
| 183 | ||
| 184 | ||
| 77994 | 185 | (* theory and proof context *) | 
| 77895 | 186 | |
| 60947 | 187 | datatype theory = | 
| 77996 | 188 | Thy_Undef | 
| 189 | | Thy of | |
| 190 | (*allocation state*) | |
| 191 | state * | |
| 192 | (*identity*) | |
| 193 |     {theory_id: theory_id,
 | |
| 78004 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 194 | theory_token: theory Unsynchronized.ref, | 
| 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 195 | theory_token_pos: Position.T} * | 
| 77996 | 196 | (*ancestry*) | 
| 197 |     {parents: theory list,         (*immediate predecessors*)
 | |
| 198 | ancestors: theory list} * (*all predecessors -- canonical reverse order*) | |
| 199 | (*data*) | |
| 200 | Any.T Datatab.table; (*body content*) | |
| 60947 | 201 | |
| 77996 | 202 | datatype proof = | 
| 203 | Prf_Undef | |
| 204 | | Prf of | |
| 205 | (*identity*) | |
| 78004 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 206 | proof Unsynchronized.ref * (*token*) | 
| 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 207 | Position.T * (*token_pos*) | 
| 77996 | 208 | theory * | 
| 209 | (*data*) | |
| 210 | Any.T Datatab.table; | |
| 211 | ||
| 77994 | 212 | structure Proof = struct type context = proof end; | 
| 213 | ||
| 214 | datatype generic = Theory of theory | Proof of Proof.context; | |
| 215 | ||
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 216 | |
| 77996 | 217 | (* heap allocations *) | 
| 218 | ||
| 78018 
dfa44d85d751
proper system options to control context tracing/timing;
 wenzelm parents: 
78004diff
changeset | 219 | val theory_tracing = Unsynchronized.ref false; | 
| 
dfa44d85d751
proper system options to control context tracing/timing;
 wenzelm parents: 
78004diff
changeset | 220 | val proof_tracing = Unsynchronized.ref false; | 
| 77996 | 221 | |
| 78035 | 222 | fun enabled_tracing () = ! theory_tracing orelse ! proof_tracing; | 
| 223 | ||
| 77996 | 224 | local | 
| 225 | ||
| 78464 
12af46f2cd94
avoid excessive accumulation of garbage, for profiling of huge sessions;
 wenzelm parents: 
78062diff
changeset | 226 | val m = Integer.pow 18 2; | 
| 
12af46f2cd94
avoid excessive accumulation of garbage, for profiling of huge sessions;
 wenzelm parents: 
78062diff
changeset | 227 | |
| 78020 | 228 | fun cons_tokens var token = | 
| 78464 
12af46f2cd94
avoid excessive accumulation of garbage, for profiling of huge sessions;
 wenzelm parents: 
78062diff
changeset | 229 | Synchronized.change var (fn (n, tokens) => | 
| 
12af46f2cd94
avoid excessive accumulation of garbage, for profiling of huge sessions;
 wenzelm parents: 
78062diff
changeset | 230 | let val tokens' = if n mod m = 0 then filter Unsynchronized.weak_active tokens else tokens | 
| 
12af46f2cd94
avoid excessive accumulation of garbage, for profiling of huge sessions;
 wenzelm parents: 
78062diff
changeset | 231 | in (n + 1, Weak.weak (SOME token) :: tokens') end); | 
| 78020 | 232 | |
| 233 | fun finish_tokens var = | |
| 234 | Synchronized.change_result var (fn (n, tokens) => | |
| 235 | let | |
| 236 | val tokens' = filter Unsynchronized.weak_active tokens; | |
| 237 | val results = map_filter Unsynchronized.weak_peek tokens'; | |
| 238 | in ((n, results), (n, tokens')) end); | |
| 239 | ||
| 77996 | 240 | fun make_token guard var token0 = | 
| 241 | if ! guard then | |
| 242 | let | |
| 243 | val token = Unsynchronized.ref (! token0); | |
| 78004 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 244 | val pos = Position.thread_data (); | 
| 78020 | 245 | fun assign res = (token := res; cons_tokens var token; res); | 
| 78004 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 246 | in (token, pos, assign) end | 
| 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 247 | else (token0, Position.none, I); | 
| 77996 | 248 | |
| 78020 | 249 | val theory_tokens = Synchronized.var "theory_tokens" (0, []: theory Unsynchronized.weak_ref list); | 
| 250 | val proof_tokens = Synchronized.var "proof_tokens" (0, []: proof Unsynchronized.weak_ref list); | |
| 77996 | 251 | |
| 252 | val theory_token0 = Unsynchronized.ref Thy_Undef; | |
| 253 | val proof_token0 = Unsynchronized.ref Prf_Undef; | |
| 254 | ||
| 255 | in | |
| 256 | ||
| 78018 
dfa44d85d751
proper system options to control context tracing/timing;
 wenzelm parents: 
78004diff
changeset | 257 | fun theory_token () = make_token theory_tracing theory_tokens theory_token0; | 
| 
dfa44d85d751
proper system options to control context tracing/timing;
 wenzelm parents: 
78004diff
changeset | 258 | fun proof_token () = make_token proof_tracing proof_tokens proof_token0; | 
| 77996 | 259 | |
| 78020 | 260 | fun finish_tracing () = | 
| 77996 | 261 | let | 
| 262 | val _ = ML_Heap.full_gc (); | |
| 78020 | 263 | val (total_theories, token_theories) = finish_tokens theory_tokens; | 
| 264 | val (total_proofs, token_proofs) = finish_tokens proof_tokens; | |
| 77996 | 265 | |
| 78020 | 266 |     fun cons1 (thy as Thy (_, {theory_token_pos, ...}, _, _)) = cons (Theory thy, theory_token_pos)
 | 
| 267 | | cons1 _ = I; | |
| 268 | fun cons2 (ctxt as Prf (_, proof_token_pos, _, _)) = cons (Proof ctxt, proof_token_pos) | |
| 269 | | cons2 _ = I; | |
| 77996 | 270 | |
| 78020 | 271 | val contexts = build (fold cons1 token_theories #> fold cons2 token_proofs); | 
| 78004 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 272 | val active_theories = fold (fn (Theory _, _) => Integer.add 1 | _ => I) contexts 0; | 
| 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 273 | val active_proofs = fold (fn (Proof _, _) => Integer.add 1 | _ => I) contexts 0; | 
| 77996 | 274 | in | 
| 275 |     {contexts = contexts,
 | |
| 276 | active_contexts = active_theories + active_proofs, | |
| 277 | active_theories = active_theories, | |
| 278 | active_proofs = active_proofs, | |
| 279 | total_contexts = total_theories + total_proofs, | |
| 280 | total_theories = total_theories, | |
| 281 | total_proofs = total_proofs} | |
| 282 | end; | |
| 283 | ||
| 284 | end; | |
| 285 | ||
| 286 | ||
| 77994 | 287 | |
| 288 | (*** theory operations ***) | |
| 289 | ||
| 290 | fun rep_theory (Thy args) = args; | |
| 291 | ||
| 292 | exception THEORY of string * theory list; | |
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 293 | |
| 77895 | 294 | val state_of = #1 o rep_theory; | 
| 295 | val theory_identity = #2 o rep_theory; | |
| 67623 | 296 | val theory_id = #theory_id o theory_identity; | 
| 77886 | 297 | val identity_of = rep_theory_id o theory_id; | 
| 298 | val ancestry_of = #3 o rep_theory; | |
| 299 | val data_of = #4 o rep_theory; | |
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 300 | |
| 16489 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 301 | fun make_ancestry parents ancestors = {parents = parents, ancestors = ancestors};
 | 
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 302 | |
| 77886 | 303 | fun stage_final stage = stage = 0; | 
| 304 | ||
| 305 | val theory_id_stage = #stage o rep_theory_id; | |
| 306 | val theory_id_final = stage_final o theory_id_stage; | |
| 307 | val theory_id_ord = int_ord o apply2 (#id o rep_theory_id); | |
| 77889 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
77886diff
changeset | 308 | fun theory_id_name {long} thy_id =
 | 
| 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
77886diff
changeset | 309 | let val name = #name (rep_theory_id thy_id) | 
| 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
77886diff
changeset | 310 | in if long then name else Long_Name.base_name name end; | 
| 77886 | 311 | |
| 312 | val theory_long_name = #name o identity_of; | |
| 77889 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
77886diff
changeset | 313 | val theory_base_name = Long_Name.base_name o theory_long_name; | 
| 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
77886diff
changeset | 314 | fun theory_name {long} = if long then theory_long_name else theory_base_name;
 | 
| 77886 | 315 | val theory_identifier = #id o identity_of; | 
| 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: 
62889diff
changeset | 316 | |
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 317 | val parents_of = #parents o ancestry_of; | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 318 | val ancestors_of = #ancestors o ancestry_of; | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 319 | |
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 320 | |
| 29093 
1cc36c0ec9eb
refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
 wenzelm parents: 
29069diff
changeset | 321 | (* names *) | 
| 
1cc36c0ec9eb
refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
 wenzelm parents: 
29069diff
changeset | 322 | |
| 
1cc36c0ec9eb
refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
 wenzelm parents: 
29069diff
changeset | 323 | val PureN = "Pure"; | 
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 324 | |
| 61044 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 325 | fun display_name thy_id = | 
| 77889 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
77886diff
changeset | 326 | let | 
| 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
77886diff
changeset | 327 |     val name = theory_id_name {long = false} thy_id;
 | 
| 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
77886diff
changeset | 328 | val final = theory_id_final thy_id; | 
| 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
77886diff
changeset | 329 | in if final then name else name ^ ":" ^ string_of_int (theory_id_stage thy_id) end; | 
| 61044 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 330 | |
| 29093 
1cc36c0ec9eb
refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
 wenzelm parents: 
29069diff
changeset | 331 | fun display_names thy = | 
| 
1cc36c0ec9eb
refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
 wenzelm parents: 
29069diff
changeset | 332 | let | 
| 61044 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 333 | val name = display_name (theory_id thy); | 
| 68192 | 334 | 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: 
60948diff
changeset | 335 | in rev (name :: ancestor_names) end; | 
| 29069 
c7ba485581ae
unified ids for ancestors and checkpoints, removed obsolete history of checkpoints;
 wenzelm parents: 
29001diff
changeset | 336 | |
| 29093 
1cc36c0ec9eb
refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
 wenzelm parents: 
29069diff
changeset | 337 | val pretty_thy = Pretty.str_list "{" "}" o display_names;
 | 
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 338 | |
| 80809 
4a64fc4d1cde
clarified signature: type ML_Pretty.pretty coincides with PolyML.pretty;
 wenzelm parents: 
79081diff
changeset | 339 | val _ = ML_system_pp (fn _ => fn _ => Pretty.to_ML o pretty_thy); | 
| 62663 | 340 | |
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 341 | fun pretty_abbrev_thy thy = | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 342 | let | 
| 29093 
1cc36c0ec9eb
refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
 wenzelm parents: 
29069diff
changeset | 343 | val names = display_names thy; | 
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 344 | val n = length names; | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 345 | 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: 
15801diff
changeset | 346 |   in Pretty.str_list "{" "}" abbrev end;
 | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 347 | |
| 68482 | 348 | fun get_theory long thy name = | 
| 77889 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
77886diff
changeset | 349 | if theory_name long thy <> name then | 
| 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
77886diff
changeset | 350 | (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: 
37852diff
changeset | 351 | SOME thy' => thy' | 
| 
b9783e9e96e1
added Context.get_theory -- avoid referring to accidental theory loader state (cf. Thy_Info.get_theory);
 wenzelm parents: 
37852diff
changeset | 352 |     | NONE => error ("Unknown ancestor theory " ^ quote name))
 | 
| 77886 | 353 | else if theory_id_final (theory_id thy) then thy | 
| 37867 
b9783e9e96e1
added Context.get_theory -- avoid referring to accidental theory loader state (cf. Thy_Info.get_theory);
 wenzelm parents: 
37852diff
changeset | 354 |   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: 
37852diff
changeset | 355 | |
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 356 | |
| 77895 | 357 | (* identity *) | 
| 29093 
1cc36c0ec9eb
refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
 wenzelm parents: 
29069diff
changeset | 358 | |
| 77895 | 359 | fun merge_ids thys = | 
| 79076 | 360 |   fold (identity_of #> (fn {id, ids, ...} => fn acc => Bitset.merge (acc, ids) |> Bitset.insert id))
 | 
| 361 | thys Bitset.empty; | |
| 16533 | 362 | |
| 77886 | 363 | val eq_thy_id = op = o apply2 (#id o rep_theory_id); | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
55547diff
changeset | 364 | val eq_thy = op = o apply2 (#id o identity_of); | 
| 16533 | 365 | |
| 79076 | 366 | val proper_subthy_id = apply2 rep_theory_id #> (fn ({id, ...}, {ids, ...}) => Bitset.member ids id);
 | 
| 60947 | 367 | val proper_subthy = proper_subthy_id o apply2 theory_id; | 
| 16533 | 368 | |
| 60947 | 369 | fun subthy_id p = eq_thy_id p orelse proper_subthy_id p; | 
| 370 | val subthy = subthy_id o apply2 theory_id; | |
| 16533 | 371 | |
| 372 | ||
| 29093 
1cc36c0ec9eb
refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
 wenzelm parents: 
29069diff
changeset | 373 | (* consistent ancestors *) | 
| 
1cc36c0ec9eb
refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
 wenzelm parents: 
29069diff
changeset | 374 | |
| 50431 
955c4aa44f60
check consistent theory names for direct imports as well -- as claimed in the comments (see also 1cc36c0ec9eb);
 wenzelm parents: 
48992diff
changeset | 375 | 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: 
48992diff
changeset | 376 | eq_thy (thy1, thy2) orelse | 
| 77889 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
77886diff
changeset | 377 | (theory_base_name thy1 = theory_base_name thy2 andalso | 
| 50431 
955c4aa44f60
check consistent theory names for direct imports as well -- as claimed in the comments (see also 1cc36c0ec9eb);
 wenzelm parents: 
48992diff
changeset | 378 |       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: 
48992diff
changeset | 379 | |
| 29093 
1cc36c0ec9eb
refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
 wenzelm parents: 
29069diff
changeset | 380 | 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: 
48992diff
changeset | 381 | if member eq_thy_consistent thys thy then | 
| 33033 | 382 |     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: 
29069diff
changeset | 383 | else thy :: thys; | 
| 
1cc36c0ec9eb
refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
 wenzelm parents: 
29069diff
changeset | 384 | |
| 50431 
955c4aa44f60
check consistent theory names for direct imports as well -- as claimed in the comments (see also 1cc36c0ec9eb);
 wenzelm parents: 
48992diff
changeset | 385 | 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: 
29069diff
changeset | 386 | |
| 77895 | 387 | val eq_ancestry = | 
| 388 | apply2 ancestry_of #> | |
| 389 |     (fn ({parents, ancestors}, {parents = parents', ancestors = ancestors'}) =>
 | |
| 390 | eq_list eq_thy (parents, parents') andalso eq_list eq_thy (ancestors, ancestors')); | |
| 391 | ||
| 29093 
1cc36c0ec9eb
refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
 wenzelm parents: 
29069diff
changeset | 392 | |
| 16489 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 393 | |
| 61262 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 394 | (** theory data **) | 
| 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 395 | |
| 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 396 | (* data kinds and access methods *) | 
| 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 397 | |
| 78018 
dfa44d85d751
proper system options to control context tracing/timing;
 wenzelm parents: 
78004diff
changeset | 398 | val data_timing = Unsynchronized.ref false; | 
| 61262 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 399 | |
| 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 400 | local | 
| 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 401 | |
| 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 402 | type kind = | 
| 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 403 |  {pos: Position.T,
 | 
| 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 404 | empty: Any.T, | 
| 77895 | 405 | merge: (theory * Any.T) list -> Any.T}; | 
| 61262 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 406 | |
| 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 407 | val kinds = Synchronized.var "Theory_Data" (Datatab.empty: kind Datatab.table); | 
| 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 408 | |
| 77895 | 409 | fun the_kind k = | 
| 61262 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 410 | (case Datatab.lookup (Synchronized.value kinds) k of | 
| 77895 | 411 | SOME kind => kind | 
| 61262 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 412 | | NONE => raise Fail "Invalid theory data identifier"); | 
| 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 413 | |
| 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 414 | in | 
| 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 415 | |
| 77897 | 416 | fun data_kinds () = | 
| 417 |   Datatab.fold_rev (fn (k, {pos, ...}) => cons (k, pos)) (Synchronized.value kinds) [];
 | |
| 418 | ||
| 77895 | 419 | val invoke_pos = #pos o the_kind; | 
| 420 | val invoke_empty = #empty o the_kind; | |
| 421 | ||
| 422 | fun invoke_merge kind args = | |
| 78018 
dfa44d85d751
proper system options to control context tracing/timing;
 wenzelm parents: 
78004diff
changeset | 423 | if ! data_timing then | 
| 77895 | 424 |     Timing.cond_timeit true ("Theory_Data.merge" ^ Position.here (#pos kind))
 | 
| 425 | (fn () => #merge kind args) | |
| 426 | else #merge kind args; | |
| 61262 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 427 | |
| 77894 | 428 | fun declare_data pos empty merge = | 
| 61262 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 429 | let | 
| 77897 | 430 | val k = data_kind (); | 
| 74561 
8e6c973003c8
discontinued obsolete "val extend = I" for data slots;
 wenzelm parents: 
74461diff
changeset | 431 |     val kind = {pos = pos, empty = empty, merge = merge};
 | 
| 61262 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 432 | val _ = Synchronized.change kinds (Datatab.update (k, kind)); | 
| 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 433 | in k end; | 
| 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 434 | |
| 77895 | 435 | fun lookup_data k thy = Datatab.lookup (data_of thy) k; | 
| 436 | ||
| 77894 | 437 | fun get_data k thy = | 
| 77895 | 438 | (case lookup_data k thy of | 
| 77894 | 439 | SOME x => x | 
| 440 | | NONE => invoke_empty k); | |
| 441 | ||
| 77895 | 442 | fun merge_data [] = Datatab.empty | 
| 443 | | merge_data [thy] = data_of thy | |
| 444 | | merge_data thys = | |
| 445 | let | |
| 446 | fun merge (k, kind) data = | |
| 447 | (case map_filter (fn thy => lookup_data k thy |> Option.map (pair thy)) thys of | |
| 448 | [] => data | |
| 449 | | [(_, x)] => Datatab.default (k, x) data | |
| 450 | | args => Datatab.update (k, invoke_merge kind args) data); | |
| 451 | in Datatab.fold merge (Synchronized.value kinds) (data_of (hd thys)) end; | |
| 61262 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 452 | |
| 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 453 | end; | 
| 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 454 | |
| 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 455 | |
| 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 456 | |
| 16489 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 457 | (** build theories **) | 
| 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 458 | |
| 67623 | 459 | (* create theory *) | 
| 460 | ||
| 77895 | 461 | fun create_thy state ids name stage ancestry data = | 
| 67623 | 462 | let | 
| 77897 | 463 |     val theory_id = make_theory_id {id = new_id (), ids = ids, name = name, stage = stage};
 | 
| 78004 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 464 | val (token, pos, assign) = theory_token (); | 
| 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 465 |     val identity = {theory_id = theory_id, theory_token = token, theory_token_pos = pos};
 | 
| 77996 | 466 | in assign (Thy (state, identity, ancestry, data)) end; | 
| 67623 | 467 | |
| 468 | ||
| 469 | (* primitives *) | |
| 33606 
2b27020ffcb2
local mutex for theory content/identity operations;
 wenzelm parents: 
33517diff
changeset | 470 | |
| 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: 
52682diff
changeset | 471 | val pre_pure_thy = | 
| 77886 | 472 | let | 
| 473 | val state = make_state (); | |
| 474 | val stage = next_stage state; | |
| 79076 | 475 | in create_thy state Bitset.empty PureN stage (make_ancestry [] []) Datatab.empty end; | 
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 476 | |
| 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: 
52682diff
changeset | 477 | 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: 
52682diff
changeset | 478 | |
| 
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: 
52682diff
changeset | 479 | fun change_thy finish f thy = | 
| 16489 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 480 | let | 
| 77895 | 481 |     val {name, stage, ...} = identity_of thy;
 | 
| 77994 | 482 | val Thy (state, _, ancestry, data) = thy; | 
| 74561 
8e6c973003c8
discontinued obsolete "val extend = I" for data slots;
 wenzelm parents: 
74461diff
changeset | 483 | val ancestry' = | 
| 77886 | 484 | if stage_final stage | 
| 74561 
8e6c973003c8
discontinued obsolete "val extend = I" for data slots;
 wenzelm parents: 
74461diff
changeset | 485 | then make_ancestry [thy] (extend_ancestors thy (ancestors_of thy)) | 
| 
8e6c973003c8
discontinued obsolete "val extend = I" for data slots;
 wenzelm parents: 
74461diff
changeset | 486 | else ancestry; | 
| 77895 | 487 | val ids' = merge_ids [thy]; | 
| 77886 | 488 | val stage' = if finish then 0 else next_stage state; | 
| 74561 
8e6c973003c8
discontinued obsolete "val extend = I" for data slots;
 wenzelm parents: 
74461diff
changeset | 489 | val data' = f data; | 
| 77895 | 490 | in create_thy state ids' name stage' ancestry' data' end; | 
| 16489 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 491 | |
| 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: 
52682diff
changeset | 492 | in | 
| 16489 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 493 | |
| 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: 
52682diff
changeset | 494 | val update_thy = change_thy false; | 
| 74561 
8e6c973003c8
discontinued obsolete "val extend = I" for data slots;
 wenzelm parents: 
74461diff
changeset | 495 | val finish_thy = change_thy true I; | 
| 16489 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 496 | |
| 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: 
52682diff
changeset | 497 | end; | 
| 16489 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 498 | |
| 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 499 | |
| 77895 | 500 | (* join: unfinished theory nodes *) | 
| 70361 
34b271c4f400
support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
 wenzelm parents: 
70360diff
changeset | 501 | |
| 77895 | 502 | fun join_thys [] = raise List.Empty | 
| 503 | | join_thys thys = | |
| 504 | let | |
| 505 | val thy0 = hd thys; | |
| 506 | val name0 = theory_long_name thy0; | |
| 507 | val state0 = state_of thy0; | |
| 70361 
34b271c4f400
support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
 wenzelm parents: 
70360diff
changeset | 508 | |
| 77895 | 509 | fun ok thy = | 
| 510 | not (theory_id_final (theory_id thy)) andalso | |
| 511 | theory_long_name thy = name0 andalso | |
| 512 | eq_ancestry (thy0, thy); | |
| 513 | val _ = | |
| 514 | (case filter_out ok thys of | |
| 515 | [] => () | |
| 516 |           | bad => raise THEORY ("Cannot join theories", bad));
 | |
| 70361 
34b271c4f400
support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
 wenzelm parents: 
70360diff
changeset | 517 | |
| 77895 | 518 | val stage = next_stage state0; | 
| 519 | val ids = merge_ids thys; | |
| 520 | val data = merge_data thys; | |
| 521 | in create_thy state0 ids name0 stage (ancestry_of thy0) data end; | |
| 70361 
34b271c4f400
support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
 wenzelm parents: 
70360diff
changeset | 522 | |
| 
34b271c4f400
support join of anonymous theory nodes, e.g. relevant for parallel theory construction;
 wenzelm parents: 
70360diff
changeset | 523 | |
| 77895 | 524 | (* merge: finished theory nodes *) | 
| 61044 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 525 | |
| 77895 | 526 | fun make_parents thys = | 
| 527 | let val thys' = distinct eq_thy thys | |
| 528 | in thys' |> filter_out (fn thy => exists (fn thy' => proper_subthy (thy, thy')) thys') end; | |
| 61044 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 529 | |
| 61262 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 530 | 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: 
52682diff
changeset | 531 |   if name = "" then error ("Bad theory name: " ^ quote name)
 | 
| 77895 | 532 | else if null imports then error "Missing theory imports" | 
| 24369 | 533 | else | 
| 16489 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 534 | let | 
| 77895 | 535 | val parents = make_parents imports; | 
| 29093 
1cc36c0ec9eb
refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
 wenzelm parents: 
29069diff
changeset | 536 | val ancestors = | 
| 77895 | 537 | Library.foldl1 merge_ancestors (map ancestors_of parents) | 
| 29093 
1cc36c0ec9eb
refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
 wenzelm parents: 
29069diff
changeset | 538 | |> fold extend_ancestors parents; | 
| 77895 | 539 | val ancestry = make_ancestry parents ancestors; | 
| 29093 
1cc36c0ec9eb
refined identity and ancestry: explicit draft flat, ids are unnamed, name consistency is checked for ancestry;
 wenzelm parents: 
29069diff
changeset | 540 | |
| 77886 | 541 | val state = make_state (); | 
| 542 | val stage = next_stage state; | |
| 77895 | 543 | val ids = merge_ids parents; | 
| 544 | val data = merge_data parents; | |
| 545 | in create_thy state ids name stage ancestry data |> tap finish_thy end; | |
| 61044 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 546 | |
| 16489 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 547 | |
| 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 548 | (* theory data *) | 
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 549 | |
| 33033 | 550 | structure Theory_Data = | 
| 16489 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 551 | struct | 
| 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 552 | |
| 77894 | 553 | val declare = declare_data; | 
| 16489 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 554 | |
| 77894 | 555 | fun get k dest thy = dest (get_data k thy); | 
| 16489 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 556 | |
| 77894 | 557 | fun put k make x = update_thy (Datatab.update (k, make x)); | 
| 16489 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 558 | |
| 77693 
068ff989c143
clarified theory_sizeof1_data: count bytes, individually for each data entry;
 wenzelm parents: 
74561diff
changeset | 559 | fun sizeof1 k thy = | 
| 
068ff989c143
clarified theory_sizeof1_data: count bytes, individually for each data entry;
 wenzelm parents: 
74561diff
changeset | 560 | Datatab.lookup (data_of thy) k |> Option.map ML_Heap.sizeof1; | 
| 67621 | 561 | |
| 16489 
f66ab8a4e98f
improved treatment of intermediate checkpoints: actual copy
 wenzelm parents: 
16436diff
changeset | 562 | end; | 
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 563 | |
| 77766 | 564 | fun theory_data_sizeof1 thy = | 
| 74234 | 565 | build (data_of thy |> Datatab.fold_rev (fn (k, _) => | 
| 77693 
068ff989c143
clarified theory_sizeof1_data: count bytes, individually for each data entry;
 wenzelm parents: 
74561diff
changeset | 566 | (case Theory_Data.sizeof1 k thy of | 
| 67621 | 567 | NONE => I | 
| 74234 | 568 | | SOME n => (cons (invoke_pos k, n))))); | 
| 67621 | 569 | |
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 570 | |
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 571 | |
| 16533 | 572 | (*** proof context ***) | 
| 573 | ||
| 574 | (* proof data kinds *) | |
| 575 | ||
| 576 | local | |
| 577 | ||
| 51368 
2ea5c7c2d825
tuned signature -- prefer terminology of Scala and Axiom;
 wenzelm parents: 
50431diff
changeset | 578 | val kinds = Synchronized.var "Proof_Data" (Datatab.empty: (theory -> Any.T) Datatab.table); | 
| 16533 | 579 | |
| 59163 | 580 | fun init_data thy = | 
| 581 | Synchronized.value kinds |> Datatab.map (fn _ => fn init => init thy); | |
| 16533 | 582 | |
| 59163 | 583 | fun init_new_data thy = | 
| 584 | Synchronized.value kinds |> Datatab.fold (fn (k, init) => fn data => | |
| 585 | if Datatab.defined data k then data | |
| 586 | 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: 
22827diff
changeset | 587 | |
| 59163 | 588 | fun init_fallback k thy = | 
| 589 | (case Datatab.lookup (Synchronized.value kinds) k of | |
| 590 | SOME init => init thy | |
| 591 | | NONE => raise Fail "Invalid proof data identifier"); | |
| 16533 | 592 | |
| 593 | in | |
| 594 | ||
| 78004 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 595 | fun raw_transfer thy' (ctxt as Prf (_, _, thy, data)) = | 
| 77996 | 596 | if eq_thy (thy, thy') then ctxt | 
| 597 | else if proper_subthy (thy, thy') then | |
| 598 | let | |
| 78004 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 599 | val (token', pos', assign) = proof_token (); | 
| 77996 | 600 | val data' = init_new_data thy' data; | 
| 78004 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 601 | in assign (Prf (token', pos', thy', data')) end | 
| 77996 | 602 | else error "Cannot transfer proof context: not a super theory"; | 
| 22847 
22da6c4bc422
simplified DataFun interfaces: removed name/print, use adhoc value for uninitialized data, init only required for impure data;
 wenzelm parents: 
22827diff
changeset | 603 | |
| 42360 | 604 | structure Proof_Context = | 
| 33031 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 605 | struct | 
| 78004 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 606 | fun theory_of (Prf (_, _, thy, _)) = thy; | 
| 77996 | 607 | fun init_global thy = | 
| 78004 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 608 | let val (token, pos, assign) = proof_token () | 
| 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 609 | in assign (Prf (token, pos, thy, init_data thy)) end; | 
| 77889 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
77886diff
changeset | 610 | fun get_global long thy name = init_global (get_theory long thy name); | 
| 33031 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 611 | end; | 
| 16533 | 612 | |
| 33033 | 613 | structure Proof_Data = | 
| 16533 | 614 | struct | 
| 615 | ||
| 22847 
22da6c4bc422
simplified DataFun interfaces: removed name/print, use adhoc value for uninitialized data, init only required for impure data;
 wenzelm parents: 
22827diff
changeset | 616 | fun declare init = | 
| 16533 | 617 | let | 
| 77897 | 618 | val k = data_kind (); | 
| 43684 | 619 | val _ = Synchronized.change kinds (Datatab.update (k, init)); | 
| 16533 | 620 | in k end; | 
| 621 | ||
| 78004 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 622 | fun get k dest (Prf (_, _, thy, data)) = | 
| 59163 | 623 | (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: 
22827diff
changeset | 624 | SOME x => x | 
| 59163 | 625 | | NONE => init_fallback k thy) |> dest; | 
| 16533 | 626 | |
| 78004 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 627 | fun put k make x (Prf (_, _, thy, data)) = | 
| 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 628 | let | 
| 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 629 | val (token', pos', assign) = proof_token (); | 
| 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 630 | val data' = Datatab.update (k, make x) data; | 
| 
19962431aea8
maintain dynamic position where values are created (again, amending afa6117bace4);
 wenzelm parents: 
78003diff
changeset | 631 | in assign (Prf (token', pos', thy, data')) end; | 
| 16533 | 632 | |
| 633 | end; | |
| 634 | ||
| 635 | end; | |
| 636 | ||
| 637 | ||
| 18632 | 638 | |
| 61044 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 639 | (*** theory certificate ***) | 
| 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 640 | |
| 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 641 | datatype certificate = Certificate of theory | Certificate_Id of theory_id; | 
| 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 642 | |
| 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 643 | fun certificate_theory (Certificate thy) = thy | 
| 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 644 | | certificate_theory (Certificate_Id thy_id) = | 
| 61062 | 645 |       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: 
60948diff
changeset | 646 | |
| 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 647 | fun certificate_theory_id (Certificate thy) = theory_id thy | 
| 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 648 | | certificate_theory_id (Certificate_Id thy_id) = thy_id; | 
| 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 649 | |
| 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 650 | 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: 
60948diff
changeset | 651 | | 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: 
60948diff
changeset | 652 | | eq_certificate _ = false; | 
| 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 653 | |
| 78062 
edb195122938
support for context within morphism (for background theory);
 wenzelm parents: 
78035diff
changeset | 654 | fun err_join (thy_id1, thy_id2) = | 
| 
edb195122938
support for context within morphism (for background theory);
 wenzelm parents: 
78035diff
changeset | 655 |   error ("Cannot join unrelated theory certificates " ^
 | 
| 
edb195122938
support for context within morphism (for background theory);
 wenzelm parents: 
78035diff
changeset | 656 | display_name thy_id1 ^ " and " ^ display_name thy_id2); | 
| 
edb195122938
support for context within morphism (for background theory);
 wenzelm parents: 
78035diff
changeset | 657 | |
| 61044 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 658 | fun join_certificate (cert1, cert2) = | 
| 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 659 | 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: 
60948diff
changeset | 660 | 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: 
60948diff
changeset | 661 | 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: 
60948diff
changeset | 662 | else if proper_subthy_id (thy_id1, thy_id2) then cert2 | 
| 78062 
edb195122938
support for context within morphism (for background theory);
 wenzelm parents: 
78035diff
changeset | 663 | else err_join (thy_id1, thy_id2) | 
| 61044 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 664 | end; | 
| 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 665 | |
| 78062 
edb195122938
support for context within morphism (for background theory);
 wenzelm parents: 
78035diff
changeset | 666 | fun join_certificate_theory (thy1, thy2) = | 
| 
edb195122938
support for context within morphism (for background theory);
 wenzelm parents: 
78035diff
changeset | 667 | let val (thy_id1, thy_id2) = apply2 theory_id (thy1, thy2) in | 
| 
edb195122938
support for context within morphism (for background theory);
 wenzelm parents: 
78035diff
changeset | 668 | if subthy_id (thy_id2, thy_id1) then thy1 | 
| 
edb195122938
support for context within morphism (for background theory);
 wenzelm parents: 
78035diff
changeset | 669 | else if proper_subthy_id (thy_id1, thy_id2) then thy2 | 
| 
edb195122938
support for context within morphism (for background theory);
 wenzelm parents: 
78035diff
changeset | 670 | else err_join (thy_id1, thy_id2) | 
| 
edb195122938
support for context within morphism (for background theory);
 wenzelm parents: 
78035diff
changeset | 671 | end; | 
| 61044 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 672 | |
| 
b7af255dd200
more abstract theory certificate, which is not necessarily the full theory;
 wenzelm parents: 
60948diff
changeset | 673 | |
| 16533 | 674 | (*** generic context ***) | 
| 675 | ||
| 18632 | 676 | fun cases f _ (Theory thy) = f thy | 
| 677 | | cases _ g (Proof prf) = g prf; | |
| 16533 | 678 | |
| 19678 | 679 | fun mapping f g = cases (Theory o f) (Proof o g); | 
| 21660 | 680 | fun mapping_result f g = cases (apsnd Theory o f) (apsnd Proof o g); | 
| 19678 | 681 | |
| 23595 
7ca68a2c8575
the_theory/proof: error instead of exception Fail;
 wenzelm parents: 
23355diff
changeset | 682 | val the_theory = cases I (fn _ => error "Ill-typed context: theory expected"); | 
| 
7ca68a2c8575
the_theory/proof: error instead of exception Fail;
 wenzelm parents: 
23355diff
changeset | 683 | val the_proof = cases (fn _ => error "Ill-typed context: proof expected") I; | 
| 16533 | 684 | |
| 18731 | 685 | fun map_theory f = Theory o f o the_theory; | 
| 686 | fun map_proof f = Proof o f o the_proof; | |
| 687 | ||
| 26486 | 688 | fun map_theory_result f = apsnd Theory o f o the_theory; | 
| 689 | fun map_proof_result f = apsnd Proof o f o the_proof; | |
| 690 | ||
| 18731 | 691 | fun theory_map f = the_theory o f o Theory; | 
| 692 | fun proof_map f = the_proof o f o Proof; | |
| 18665 | 693 | |
| 42360 | 694 | val theory_of = cases I Proof_Context.theory_of; | 
| 695 | val proof_of = cases Proof_Context.init_global I; | |
| 16533 | 696 | |
| 22085 
c138cfd500f7
ML context: full generic context, tuned signature;
 wenzelm parents: 
21962diff
changeset | 697 | |
| 
c138cfd500f7
ML context: full generic context, tuned signature;
 wenzelm parents: 
21962diff
changeset | 698 | |
| 26413 
003dd6155870
added thread data (formerly global ref in ML/ml_context.ML);
 wenzelm parents: 
24559diff
changeset | 699 | (** thread data **) | 
| 
003dd6155870
added thread data (formerly global ref in ML/ml_context.ML);
 wenzelm parents: 
24559diff
changeset | 700 | |
| 62889 | 701 | 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: 
24559diff
changeset | 702 | |
| 62889 | 703 | fun get_generic_context () = Thread_Data.get generic_context_var; | 
| 704 | val put_generic_context = Thread_Data.put generic_context_var; | |
| 705 | fun setmp_generic_context opt_context = Thread_Data.setmp generic_context_var opt_context; | |
| 62876 | 706 | |
| 707 | fun the_generic_context () = | |
| 62889 | 708 | (case get_generic_context () of | 
| 26413 
003dd6155870
added thread data (formerly global ref in ML/ml_context.ML);
 wenzelm parents: 
24559diff
changeset | 709 | SOME context => context | 
| 
003dd6155870
added thread data (formerly global ref in ML/ml_context.ML);
 wenzelm parents: 
24559diff
changeset | 710 | | _ => error "Unknown context"); | 
| 
003dd6155870
added thread data (formerly global ref in ML/ml_context.ML);
 wenzelm parents: 
24559diff
changeset | 711 | |
| 62876 | 712 | val the_global_context = theory_of o the_generic_context; | 
| 713 | 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: 
24559diff
changeset | 714 | |
| 26428 | 715 | end; | 
| 716 | ||
| 717 | fun >>> f = | |
| 718 | let | |
| 62876 | 719 | val (res, context') = f (the_generic_context ()); | 
| 62889 | 720 | val _ = put_generic_context (SOME context'); | 
| 26428 | 721 | in res end; | 
| 722 | ||
| 26421 | 723 | nonfix >>; | 
| 26463 | 724 | fun >> f = >>> (fn context => ((), f context)); | 
| 26413 
003dd6155870
added thread data (formerly global ref in ML/ml_context.ML);
 wenzelm parents: 
24559diff
changeset | 725 | |
| 62889 | 726 | val _ = put_generic_context (SOME (Theory pre_pure_thy)); | 
| 26413 
003dd6155870
added thread data (formerly global ref in ML/ml_context.ML);
 wenzelm parents: 
24559diff
changeset | 727 | |
| 6185 | 728 | end; | 
| 729 | ||
| 33031 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 730 | structure Basic_Context: BASIC_CONTEXT = Context; | 
| 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 731 | open Basic_Context; | 
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 732 | |
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 733 | |
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 734 | |
| 16533 | 735 | (*** type-safe interfaces for data declarations ***) | 
| 736 | ||
| 737 | (** theory data **) | |
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 738 | |
| 61262 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 739 | signature THEORY_DATA'_ARGS = | 
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 740 | sig | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 741 | type T | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 742 | val empty: T | 
| 77895 | 743 | val merge: (theory * T) list -> T | 
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 744 | end; | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 745 | |
| 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: 
34253diff
changeset | 746 | 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: 
34253diff
changeset | 747 | 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: 
34253diff
changeset | 748 | 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: 
34253diff
changeset | 749 | 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: 
34253diff
changeset | 750 | 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: 
34253diff
changeset | 751 | 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: 
34253diff
changeset | 752 | |
| 34253 
5930c6391126
removed further remains of mutable theory data (cf. 25bd3ed2ac9f);
 wenzelm parents: 
34245diff
changeset | 753 | signature THEORY_DATA = | 
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 754 | sig | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 755 | type T | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 756 | val get: theory -> T | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 757 | val put: T -> theory -> theory | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 758 | val map: (T -> T) -> theory -> theory | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 759 | end; | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 760 | |
| 61262 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 761 | functor Theory_Data'(Data: THEORY_DATA'_ARGS): THEORY_DATA = | 
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 762 | struct | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 763 | |
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 764 | type T = Data.T; | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 765 | exception Data of T; | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 766 | |
| 42818 
128cc195ced3
timing of Theory_Data operations, with implicit thread positions when functor is applied;
 wenzelm parents: 
42383diff
changeset | 767 | val kind = | 
| 72060 
efb7fd4a6d1f
subtle change of Theory_Data extend/merge semantics due to Theory.join_theory;
 wenzelm parents: 
70586diff
changeset | 768 | let val pos = Position.thread_data () in | 
| 
efb7fd4a6d1f
subtle change of Theory_Data extend/merge semantics due to Theory.join_theory;
 wenzelm parents: 
70586diff
changeset | 769 | Context.Theory_Data.declare | 
| 
efb7fd4a6d1f
subtle change of Theory_Data extend/merge semantics due to Theory.join_theory;
 wenzelm parents: 
70586diff
changeset | 770 | pos | 
| 
efb7fd4a6d1f
subtle change of Theory_Data extend/merge semantics due to Theory.join_theory;
 wenzelm parents: 
70586diff
changeset | 771 | (Data Data.empty) | 
| 77895 | 772 | (Data o Data.merge o map (fn (thy, Data x) => (thy, x))) | 
| 72060 
efb7fd4a6d1f
subtle change of Theory_Data extend/merge semantics due to Theory.join_theory;
 wenzelm parents: 
70586diff
changeset | 773 | end; | 
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 774 | |
| 33033 | 775 | val get = Context.Theory_Data.get kind (fn Data x => x); | 
| 776 | val put = Context.Theory_Data.put kind Data; | |
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 777 | fun map f thy = put (f (get thy)) thy; | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 778 | |
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 779 | end; | 
| 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 780 | |
| 33517 
d064fa48f305
modernized/simplified functor Theory_Data, Proof_Data, Generic_Data: eliminated Pretty.pp, discontinued mutable data;
 wenzelm parents: 
33033diff
changeset | 781 | functor Theory_Data(Data: THEORY_DATA_ARGS): THEORY_DATA = | 
| 61262 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61062diff
changeset | 782 | 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: 
34253diff
changeset | 783 | ( | 
| 
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: 
34253diff
changeset | 784 | 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: 
34253diff
changeset | 785 | val empty = Data.empty; | 
| 77895 | 786 | fun merge args = Library.foldl (fn (a, (_, b)) => Data.merge (a, b)) (#2 (hd args), tl args) | 
| 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: 
34253diff
changeset | 787 | ); | 
| 33517 
d064fa48f305
modernized/simplified functor Theory_Data, Proof_Data, Generic_Data: eliminated Pretty.pp, discontinued mutable data;
 wenzelm parents: 
33033diff
changeset | 788 | |
| 16533 | 789 | |
| 790 | ||
| 791 | (** proof data **) | |
| 792 | ||
| 793 | signature PROOF_DATA_ARGS = | |
| 794 | sig | |
| 795 | type T | |
| 796 | val init: theory -> T | |
| 797 | end; | |
| 798 | ||
| 799 | signature PROOF_DATA = | |
| 800 | sig | |
| 801 | type T | |
| 33031 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 802 | val get: Proof.context -> T | 
| 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 803 | val put: T -> Proof.context -> Proof.context | 
| 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 wenzelm parents: 
32784diff
changeset | 804 | val map: (T -> T) -> Proof.context -> Proof.context | 
| 16533 | 805 | end; | 
| 806 | ||
| 33517 
d064fa48f305
modernized/simplified functor Theory_Data, Proof_Data, Generic_Data: eliminated Pretty.pp, discontinued mutable data;
 wenzelm parents: 
33033diff
changeset | 807 | functor Proof_Data(Data: PROOF_DATA_ARGS): PROOF_DATA = | 
| 16533 | 808 | struct | 
| 809 | ||
| 810 | type T = Data.T; | |
| 811 | exception Data of T; | |
| 812 | ||
| 33033 | 813 | val kind = Context.Proof_Data.declare (Data o Data.init); | 
| 16533 | 814 | |
| 33033 | 815 | val get = Context.Proof_Data.get kind (fn Data x => x); | 
| 816 | val put = Context.Proof_Data.put kind Data; | |
| 16533 | 817 | fun map f prf = put (f (get prf)) prf; | 
| 818 | ||
| 819 | end; | |
| 820 | ||
| 18632 | 821 | |
| 822 | ||
| 823 | (** generic data **) | |
| 824 | ||
| 825 | signature GENERIC_DATA_ARGS = | |
| 826 | sig | |
| 827 | type T | |
| 828 | val empty: T | |
| 33517 
d064fa48f305
modernized/simplified functor Theory_Data, Proof_Data, Generic_Data: eliminated Pretty.pp, discontinued mutable data;
 wenzelm parents: 
33033diff
changeset | 829 | val merge: T * T -> T | 
| 18632 | 830 | end; | 
| 831 | ||
| 832 | signature GENERIC_DATA = | |
| 833 | sig | |
| 834 | type T | |
| 835 | val get: Context.generic -> T | |
| 836 | val put: T -> Context.generic -> Context.generic | |
| 837 | val map: (T -> T) -> Context.generic -> Context.generic | |
| 838 | end; | |
| 839 | ||
| 33517 
d064fa48f305
modernized/simplified functor Theory_Data, Proof_Data, Generic_Data: eliminated Pretty.pp, discontinued mutable data;
 wenzelm parents: 
33033diff
changeset | 840 | functor Generic_Data(Data: GENERIC_DATA_ARGS): GENERIC_DATA = | 
| 18632 | 841 | struct | 
| 842 | ||
| 33517 
d064fa48f305
modernized/simplified functor Theory_Data, Proof_Data, Generic_Data: eliminated Pretty.pp, discontinued mutable data;
 wenzelm parents: 
33033diff
changeset | 843 | structure Thy_Data = Theory_Data(Data); | 
| 
d064fa48f305
modernized/simplified functor Theory_Data, Proof_Data, Generic_Data: eliminated Pretty.pp, discontinued mutable data;
 wenzelm parents: 
33033diff
changeset | 844 | structure Prf_Data = Proof_Data(type T = Data.T val init = Thy_Data.get); | 
| 18632 | 845 | |
| 846 | type T = Data.T; | |
| 847 | ||
| 33033 | 848 | fun get (Context.Theory thy) = Thy_Data.get thy | 
| 849 | | get (Context.Proof prf) = Prf_Data.get prf; | |
| 18632 | 850 | |
| 33033 | 851 | fun put x (Context.Theory thy) = Context.Theory (Thy_Data.put x thy) | 
| 852 | | put x (Context.Proof prf) = Context.Proof (Prf_Data.put x prf); | |
| 18632 | 853 | |
| 854 | fun map f ctxt = put (f (get ctxt)) ctxt; | |
| 855 | ||
| 856 | end; | |
| 857 | ||
| 16533 | 858 | (*hide private interface*) | 
| 16436 
7eb6b6cbd166
added type theory: generic theory contexts with unique identity,
 wenzelm parents: 
15801diff
changeset | 859 | structure Context: CONTEXT = Context; |