src/Pure/context.ML
changeset 16489 f66ab8a4e98f
parent 16436 7eb6b6cbd166
child 16533 f1152f75f6fc
equal deleted inserted replaced
16488:38bc902946b2 16489:f66ab8a4e98f
     1 (*  Title:      Pure/context.ML
     1 (*  Title:      Pure/context.ML
     2     ID:         $Id$
     2     ID:         $Id$
     3     Author:     Markus Wenzel, TU Muenchen
     3     Author:     Markus Wenzel, TU Muenchen
     4 
     4 
     5 Generic theory contexts with unique identity, arbitrarily typed data,
     5 Generic theory contexts with unique identity, arbitrarily typed data,
     6 linear and graph development.  Implicit ML theory context.
     6 graph-based development.  Implicit theory context in ML.
     7 *)
     7 *)
     8 
     8 
     9 signature BASIC_CONTEXT =
     9 signature BASIC_CONTEXT =
    10 sig
    10 sig
    11   type theory
    11   type theory
    16 end;
    16 end;
    17 
    17 
    18 signature CONTEXT =
    18 signature CONTEXT =
    19 sig
    19 sig
    20   include BASIC_CONTEXT
    20   include BASIC_CONTEXT
       
    21   exception DATA_FAIL of exn * string
    21   (*theory context*)
    22   (*theory context*)
       
    23   val theory_name: theory -> string
    22   val parents_of: theory -> theory list
    24   val parents_of: theory -> theory list
    23   val ancestors_of: theory -> theory list
    25   val ancestors_of: theory -> theory list
    24   val is_stale: theory -> bool
    26   val is_stale: theory -> bool
    25   val ProtoPureN: string
    27   val ProtoPureN: string
    26   val PureN: string
    28   val PureN: string
    27   val CPureN: string
    29   val CPureN: string
    28   val draftN: string
    30   val draftN: string
    29   val is_draft: theory -> bool
       
    30   val exists_name: string -> theory -> bool
    31   val exists_name: string -> theory -> bool
    31   val names_of: theory -> string list
    32   val names_of: theory -> string list
    32   val pretty_thy: theory -> Pretty.T
    33   val pretty_thy: theory -> Pretty.T
    33   val string_of_thy: theory -> string
    34   val string_of_thy: theory -> string
    34   val pprint_thy: theory -> pprint_args -> unit
    35   val pprint_thy: theory -> pprint_args -> unit
    35   val pretty_abbrev_thy: theory -> Pretty.T
    36   val pretty_abbrev_thy: theory -> Pretty.T
    36   val str_of_thy: theory -> string
    37   val str_of_thy: theory -> string
    37   val check_thy: string -> theory -> theory
    38   val check_thy: string -> theory -> theory
    38   val eq_thy: theory * theory -> bool
    39   val eq_thy: theory * theory -> bool
    39   val subthy: theory * theory -> bool
    40   val subthy: theory * theory -> bool
       
    41   val merge: theory * theory -> theory                     (*exception TERM*)
       
    42   val merge_refs: theory_ref * theory_ref -> theory_ref    (*exception TERM*)
    40   val self_ref: theory -> theory_ref
    43   val self_ref: theory -> theory_ref
    41   val deref: theory_ref -> theory
    44   val deref: theory_ref -> theory
    42   exception DATA_FAIL of exn * string
       
    43   val theory_data: theory -> string list
       
    44   val print_all_data: theory -> unit
       
    45   val copy_thy: theory -> theory
    45   val copy_thy: theory -> theory
    46   val checkpoint_thy: theory -> theory
    46   val checkpoint_thy: theory -> theory
    47   val pre_pure: theory
    47   val finish_thy: theory -> theory
    48   val merge: theory * theory -> theory                     (*exception TERM*)
    48   val theory_data: theory -> string list
    49   val merge_refs: theory_ref * theory_ref -> theory_ref    (*exception TERM*)
    49   val pre_pure_thy: theory
    50   val theory_name: theory -> string
    50   val begin_thy: (theory -> Pretty.pp) -> string -> theory list -> theory
    51   val begin_theory: (theory -> Pretty.pp) -> string -> theory list -> theory
    51   (*ML theory context*)
    52   val end_theory: theory -> theory
       
    53 
       
    54   (*ML context*)
       
    55   val get_context: unit -> theory option
    52   val get_context: unit -> theory option
    56   val set_context: theory option -> unit
    53   val set_context: theory option -> unit
    57   val reset_context: unit -> unit
    54   val reset_context: unit -> unit
    58   val setmp: theory option -> ('a -> 'b) -> 'a -> 'b
    55   val setmp: theory option -> ('a -> 'b) -> 'a -> 'b
    59   val pass: theory option -> ('a -> 'b) -> 'a -> 'b * theory option
    56   val pass: theory option -> ('a -> 'b) -> 'a -> 'b * theory option
    71 signature PRIVATE_CONTEXT =
    68 signature PRIVATE_CONTEXT =
    72 sig
    69 sig
    73   include CONTEXT
    70   include CONTEXT
    74   structure TheoryData:
    71   structure TheoryData:
    75   sig
    72   sig
    76     val declare: string -> Object.T -> (Object.T -> Object.T) ->
    73     val declare: string -> Object.T -> (Object.T -> Object.T) -> (Object.T -> Object.T) ->
    77       (Object.T -> Object.T) -> (Pretty.pp -> Object.T * Object.T -> Object.T) ->
    74       (Pretty.pp -> Object.T * Object.T -> Object.T) -> serial
    78       (theory -> Object.T -> unit) -> serial
       
    79     val init: serial -> theory -> theory
    75     val init: serial -> theory -> theory
    80     val print: serial -> theory -> unit
       
    81     val get: serial -> (Object.T -> 'a) -> theory -> 'a
    76     val get: serial -> (Object.T -> 'a) -> theory -> 'a
    82     val put: serial -> ('a -> Object.T) -> 'a -> theory -> theory
    77     val put: serial -> ('a -> Object.T) -> 'a -> theory -> theory
    83   end;
    78   end
    84 end;
    79 end;
    85 
    80 
    86 structure Context: PRIVATE_CONTEXT =
    81 structure Context: PRIVATE_CONTEXT =
    87 struct
    82 struct
    88 
    83 
    89 (*** theory context ***)
    84 (*** theory context ***)
       
    85 
       
    86 (** theory data **)
       
    87 
       
    88 (* data kinds and access methods *)
       
    89 
       
    90 exception DATA_FAIL of exn * string;
       
    91 
       
    92 local
       
    93 
       
    94 type kind =
       
    95  {name: string,
       
    96   empty: Object.T,
       
    97   copy: Object.T -> Object.T,
       
    98   extend: Object.T -> Object.T,
       
    99   merge: Pretty.pp -> Object.T * Object.T -> Object.T};
       
   100 
       
   101 val kinds = ref (Inttab.empty: kind Inttab.table);
       
   102 
       
   103 fun invoke meth_name meth_fn k =
       
   104   (case Inttab.lookup (! kinds, k) of
       
   105     SOME kind => meth_fn kind |> transform_failure (fn exn =>
       
   106       DATA_FAIL (exn, "Theory data method " ^ #name kind ^ "." ^ meth_name ^ " failed"))
       
   107   | NONE => sys_error ("Invalid theory data identifier " ^ string_of_int k));
       
   108 
       
   109 in
       
   110 
       
   111 fun invoke_name k    = invoke "name" (K o #name) k ();
       
   112 fun invoke_empty k   = invoke "empty" (K o #empty) k ();
       
   113 val invoke_copy      = invoke "copy" #copy;
       
   114 val invoke_extend    = invoke "extend" #extend;
       
   115 fun invoke_merge pp  = invoke "merge" (fn kind => #merge kind pp);
       
   116 
       
   117 fun declare_theory_data name e cp ext mrg =
       
   118   let
       
   119     val k = serial ();
       
   120     val kind = {name = name, empty = e, copy = cp, extend = ext, merge = mrg};
       
   121     val _ = conditional (Inttab.exists (equal name o #name o #2) (! kinds)) (fn () =>
       
   122       warning ("Duplicate declaration of theory data " ^ quote name));
       
   123     val _ = kinds := Inttab.update ((k, kind), ! kinds);
       
   124   in k end;
       
   125 
       
   126 val copy_data = Inttab.map' invoke_copy;
       
   127 val extend_data = Inttab.map' invoke_extend;
       
   128 fun merge_data pp = Inttab.join (SOME oo invoke_merge pp) o pairself extend_data;
       
   129 
       
   130 end;
       
   131 
       
   132 
       
   133 
       
   134 (** datatype theory **)
    90 
   135 
    91 datatype theory =
   136 datatype theory =
    92   Theory of
   137   Theory of
    93   (*identity*)
   138   (*identity*)
    94    {self: theory ref option,
   139    {self: theory ref option,            (*dynamic self reference -- follows theory changes*)
    95     id: serial * string,
   140     id: serial * string,                (*identifier of this theory*)
    96     ids: string Inttab.table} *
   141     ids: string Inttab.table,           (*identifiers of ancestors*)
       
   142     iids: string Inttab.table} *        (*identifiers of intermediate checkpoints*)
    97   (*data*)
   143   (*data*)
    98   Object.T Inttab.table *
   144   Object.T Inttab.table *               (*record of arbitrarily typed data*)
       
   145   (*ancestry of graph development*)
       
   146    {parents: theory list,               (*immediate predecessors*)
       
   147     ancestors: theory list} *           (*all predecessors*)
    99   (*history of linear development*)
   148   (*history of linear development*)
   100    {name: string,
   149    {name: string,                       (*prospective name of finished theory*)
   101     version: int,
   150     version: int,                       (*checkpoint counter*)
   102     intermediates: theory list} *
   151     intermediates: theory list};        (*intermediate checkpoints*)
   103   (*ancestry of graph development*)
       
   104    {parents: theory list,
       
   105     ancestors: theory list};
       
   106 
   152 
   107 exception THEORY of string * theory list;
   153 exception THEORY of string * theory list;
   108 
   154 
   109 fun rep_theory (Theory args) = args;
   155 fun rep_theory (Theory args) = args;
   110 
   156 
   111 val identity_of = #1 o rep_theory;
   157 val identity_of = #1 o rep_theory;
   112 val data_of     = #2 o rep_theory;
   158 val data_of     = #2 o rep_theory;
   113 val history_of  = #3 o rep_theory;
   159 val ancestry_of = #3 o rep_theory;
   114 val ancestry_of = #4 o rep_theory;
   160 val history_of  = #4 o rep_theory;
   115 
   161 
   116 fun make_identity self id ids = {self = self, id = id, ids = ids};
   162 fun make_identity self id ids iids = {self = self, id = id, ids = ids, iids = iids};
       
   163 fun make_ancestry parents ancestors = {parents = parents, ancestors = ancestors};
   117 fun make_history name vers ints = {name = name, version = vers, intermediates = ints};
   164 fun make_history name vers ints = {name = name, version = vers, intermediates = ints};
   118 fun make_ancestry parents ancestors = {parents = parents, ancestors = ancestors};
       
   119 
   165 
   120 val parents_of = #parents o ancestry_of;
   166 val parents_of = #parents o ancestry_of;
   121 val ancestors_of = #ancestors o ancestry_of;
   167 val ancestors_of = #ancestors o ancestry_of;
   122 
   168 val theory_name = #name o history_of;
   123 
   169 
   124 
       
   125 (** theory identity **)
       
   126 
   170 
   127 (* staleness *)
   171 (* staleness *)
   128 
   172 
   129 fun eq_id ((i: int, _), (j, _)) = i = j;
   173 fun eq_id ((i: int, _), (j, _)) = i = j;
   130 
   174 
   132     (Theory ({self = SOME (ref (Theory ({id = id', ...}, _, _, _))), id, ...}, _, _, _)) =
   176     (Theory ({self = SOME (ref (Theory ({id = id', ...}, _, _, _))), id, ...}, _, _, _)) =
   133       not (eq_id (id, id'))
   177       not (eq_id (id, id'))
   134   | is_stale (Theory ({self = NONE, ...}, _, _, _)) = true;
   178   | is_stale (Theory ({self = NONE, ...}, _, _, _)) = true;
   135 
   179 
   136 fun vitalize (thy as Theory ({self = SOME r, ...}, _, _, _)) = (r := thy; thy)
   180 fun vitalize (thy as Theory ({self = SOME r, ...}, _, _, _)) = (r := thy; thy)
   137   | vitalize (thy as Theory ({self = NONE, id, ids}, data, history, ancestry)) =
   181   | vitalize (thy as Theory ({self = NONE, id, ids, iids}, data, ancestry, history)) =
   138       let
   182       let
   139         val r = ref thy;
   183         val r = ref thy;
   140         val thy' = Theory (make_identity (SOME r) id ids, data, history, ancestry);
   184         val thy' = Theory (make_identity (SOME r) id ids iids, data, ancestry, history);
   141       in r := thy'; thy' end;
   185       in r := thy'; thy' end;
   142 
   186 
   143 
   187 
   144 (* names *)
   188 (* names *)
   145 
   189 
   149 
   193 
   150 val draftN = "#";
   194 val draftN = "#";
   151 fun draft_id (_, name) = (name = draftN);
   195 fun draft_id (_, name) = (name = draftN);
   152 val is_draft = draft_id o #id o identity_of;
   196 val is_draft = draft_id o #id o identity_of;
   153 
   197 
   154 fun exists_name name (Theory ({id, ids, ...}, _, _, _)) =
   198 fun exists_name name (Theory ({id, ids, iids, ...}, _, _, _)) =
   155   name = #2 id orelse Inttab.exists (equal name o #2) ids;
   199   name = #2 id orelse
   156 
   200   Inttab.exists (equal name o #2) ids orelse
   157 fun names_of (Theory ({id, ids, ...}, _, _, _)) =
   201   Inttab.exists (equal name o #2) iids;
   158   map #2 (Inttab.dest ids @ [id]);
   202 
       
   203 fun names_of (Theory ({id, ids, iids, ...}, _, _, _)) =
       
   204   rev (#2 id :: Inttab.fold (cons o #2) iids (Inttab.fold (cons o #2) ids []));
   159 
   205 
   160 fun pretty_thy thy =
   206 fun pretty_thy thy =
   161   Pretty.str_list "{" "}" (names_of thy @ (if is_stale thy then ["!"] else []));
   207   Pretty.str_list "{" "}" (names_of thy @ (if is_stale thy then ["!"] else []));
   162 
   208 
   163 val string_of_thy = Pretty.string_of o pretty_thy;
   209 val string_of_thy = Pretty.string_of o pretty_thy;
   175 
   221 
   176 (* consistency *)
   222 (* consistency *)
   177 
   223 
   178 fun check_thy pos thy =
   224 fun check_thy pos thy =
   179   if is_stale thy then
   225   if is_stale thy then
   180     raise TERM ("Stale theory encountered (see " ^ pos ^ "):\n" ^ string_of_thy thy, [])
   226     raise TERM ("Stale theory encountered (in " ^ pos ^ "):\n" ^ string_of_thy thy, [])
   181   else thy;
   227   else thy;
   182 
   228 
   183 fun check_insert id ids =
   229 fun check_ins id ids =
   184   if draft_id id orelse is_some (Inttab.lookup (ids, #1 id)) then ids
   230   if draft_id id orelse is_some (Inttab.lookup (ids, #1 id)) then ids
   185   else if Inttab.exists (equal (#2 id) o #2) ids then
   231   else if Inttab.exists (equal (#2 id) o #2) ids then
   186     raise TERM ("Different versions of theory component " ^ quote (#2 id), [])
   232     raise TERM ("Different versions of theory component " ^ quote (#2 id), [])
   187   else Inttab.update (id, ids);
   233   else Inttab.update (id, ids);
   188 
   234 
   189 fun check_merge thy1 thy2 =
   235 fun check_insert intermediate id (ids, iids) =
   190   let
   236   let val ids' = check_ins id ids and iids' = check_ins id iids
   191     val {id = id1, ids = ids1, ...} = identity_of thy1;
   237   in if intermediate then (ids, iids') else (ids', iids) end;
   192     val {id = id2, ids = ids2, ...} = identity_of thy2;
   238 
   193   in check_insert id2 (Inttab.fold check_insert ids2 (check_insert id1 ids1)) end;
   239 fun check_merge
   194 
   240     (Theory ({id = id1, ids = ids1, iids = iids1, ...}, _, _, history1))
   195 
   241     (Theory ({id = id2, ids = ids2, iids = iids2, ...}, _, _, history2)) =
   196 (* equality and inclusion *)
   242   (Inttab.fold check_ins ids2 ids1, Inttab.fold check_ins iids2 iids1)
       
   243   |> check_insert (#version history1 > 0) id1
       
   244   |> check_insert (#version history2 > 0) id2;
       
   245 
       
   246 
       
   247 (* theory references *)
       
   248 
       
   249 (*theory_ref provides a safe way to store dynamic references to a
       
   250   theory -- a plain theory value would become stale as the self
       
   251   reference moves on*)
       
   252 
       
   253 datatype theory_ref = TheoryRef of theory ref;
       
   254 
       
   255 val self_ref = TheoryRef o the o #self o identity_of o check_thy "Context.self_ref";
       
   256 fun deref (TheoryRef (ref thy)) = thy;
       
   257 
       
   258 
       
   259 (* equality and inclusion *)  (* FIXME proper_subthy; no subthy_internal *)
   197 
   260 
   198 local
   261 local
   199 
   262 
   200 fun exists_ids (Theory ({id, ids, ...}, _, _, _)) (i, _) =
   263 fun exists_ids (Theory ({id, ids, iids, ...}, _, _, _)) (i, _) =
   201   i = #1 id orelse is_some (Inttab.lookup (ids, i));
   264   i = #1 id orelse
       
   265   is_some (Inttab.lookup (ids, i)) orelse
       
   266   is_some (Inttab.lookup (iids, i));
   202 
   267 
   203 fun member_ids (Theory ({id, ...}, _, _, _), thy) = exists_ids thy id;
   268 fun member_ids (Theory ({id, ...}, _, _, _), thy) = exists_ids thy id;
   204 
   269 
   205 fun subset_ids (Theory ({id, ids, ...}, _, _, _), thy) =
   270 fun subset_ids (Theory ({id, ids, iids, ...}, _, _, _), thy) =
   206   exists_ids thy id andalso Inttab.forall (exists_ids thy) ids;
   271   exists_ids thy id andalso
       
   272   Inttab.forall (exists_ids thy) ids andalso
       
   273   Inttab.forall (exists_ids thy) iids;
   207 
   274 
   208 in
   275 in
   209 
   276 
   210 val eq_thy = eq_id o pairself (#id o identity_of o check_thy "Context.eq_thy");
   277 val eq_thy = eq_id o pairself (#id o identity_of o check_thy "Context.eq_thy");
   211 fun subthy thys = eq_thy thys orelse member_ids thys;
   278 fun subthy thys = eq_thy thys orelse member_ids thys;
   212 fun subthy_internal thys = eq_thy thys orelse subset_ids thys;
   279 fun subthy_internal thys = eq_thy thys orelse subset_ids thys;
   213 
   280 
   214 end;
   281 end;
   215 
       
   216 
       
   217 (* external references *)
       
   218 
       
   219 datatype theory_ref = TheoryRef of theory ref;
       
   220 
       
   221 val self_ref = TheoryRef o the o #self o identity_of o check_thy "Context.self_ref";
       
   222 fun deref (TheoryRef (ref thy)) = thy;
       
   223 
       
   224 
       
   225 
       
   226 (** theory data **)
       
   227 
       
   228 (* data kinds and access methods *)
       
   229 
       
   230 exception DATA_FAIL of exn * string;
       
   231 
       
   232 local
       
   233 
       
   234 type kind =
       
   235  {name: string,
       
   236   empty: Object.T,
       
   237   copy: Object.T -> Object.T,
       
   238   extend: Object.T -> Object.T,
       
   239   merge: Pretty.pp -> Object.T * Object.T -> Object.T,
       
   240   print: theory -> Object.T -> unit};
       
   241 
       
   242 val kinds = ref (Inttab.empty: kind Inttab.table);
       
   243 
       
   244 fun invoke meth_name meth_fn k =
       
   245   (case Inttab.lookup (! kinds, k) of
       
   246     SOME kind => meth_fn kind |> transform_failure (fn exn =>
       
   247       DATA_FAIL (exn, "Theory data method " ^ #name kind ^ "." ^ meth_name ^ " failed"))
       
   248   | NONE => sys_error ("Invalid theory data identifier " ^ string_of_int k));
       
   249 
       
   250 in
       
   251 
       
   252 fun invoke_name k    = invoke "name" (K o #name) k ();
       
   253 fun invoke_empty k   = invoke "empty" (K o #empty) k ();
       
   254 val invoke_copy      = invoke "copy" #copy;
       
   255 val invoke_extend    = invoke "extend" #extend;
       
   256 fun invoke_merge pp  = invoke "merge" (fn kind => #merge kind pp);
       
   257 fun invoke_print thy = invoke "print" (fn kind => #print kind thy);
       
   258 
       
   259 fun declare name e cp ext mrg prt =
       
   260   let
       
   261     val k = serial ();
       
   262     val kind = {name = name, empty = e, copy = cp, extend = ext, merge = mrg, print = prt};
       
   263     val _ = conditional (Inttab.exists (equal name o #name o #2) (! kinds)) (fn () =>
       
   264       warning ("Duplicate declaration of theory data " ^ quote name));
       
   265     val _ = kinds := Inttab.update ((k, kind), ! kinds);
       
   266   in k end;
       
   267 
       
   268 val copy_data = Inttab.map' invoke_copy;
       
   269 val extend_data = Inttab.map' invoke_extend;
       
   270 fun merge_data pp = Inttab.join (SOME oo invoke_merge pp) o pairself extend_data;
       
   271 
       
   272 fun theory_data thy =
       
   273   map invoke_name (Inttab.keys (data_of thy))
       
   274   |> map (rpair ()) |> Symtab.make_multi |> Symtab.dest
       
   275   |> map (apsnd length)
       
   276   |> map (fn (name, 1) => name | (name, n) => name ^ enclose "[" "]" (string_of_int n));
       
   277 
       
   278 fun print_all_data thy =
       
   279   List.app (uncurry (invoke_print thy)) (Inttab.dest (data_of thy));
       
   280 
       
   281 end;
       
   282 
       
   283 
       
   284 
       
   285 (** build theories **)
       
   286 
       
   287 (* primitives *)
       
   288 
       
   289 fun create_thy name self id ids data history ancestry =
       
   290   let
       
   291     val ids' = check_insert id ids;
       
   292     val id' = (serial (), name);
       
   293     val _ = check_insert id' ids';
       
   294     val identity' = make_identity self id' ids';
       
   295   in vitalize (Theory (identity', data, history, ancestry)) end;
       
   296 
       
   297 fun copy_thy (thy as Theory ({id, ids, ...}, data, history, ancestry)) =
       
   298   let val _ = check_thy "Context.copy_thy" thy;
       
   299   in create_thy draftN NONE id ids (copy_data data) history ancestry end;
       
   300 
       
   301 fun change_thy name f (thy as Theory ({self, id, ids}, data, history, ancestry)) =
       
   302   let
       
   303     val _ = check_thy "Context.change_thy" thy;
       
   304     val (self', data', ancestry') =
       
   305       if is_draft thy then (self, data, ancestry)
       
   306       else (NONE, extend_data data, make_ancestry [thy] (thy :: #ancestors ancestry));
       
   307   in create_thy name self' id ids (f data') history ancestry' end;
       
   308 
       
   309 fun name_thy name = change_thy name I;
       
   310 val map_thy = change_thy draftN;
       
   311 val extend_thy = map_thy I;
       
   312 
       
   313 fun checkpoint_thy thy =
       
   314   if not (is_draft thy) then thy
       
   315   else
       
   316     let
       
   317       val {name, version, intermediates} = history_of thy;
       
   318       val thy' as Theory (identity', data', _, ancestry') =
       
   319         name_thy (name ^ ":" ^ string_of_int version) thy;
       
   320       val history' = make_history name (version + 1) (thy' :: intermediates);
       
   321     in vitalize (Theory (identity', data', history', ancestry')) end;
       
   322 
       
   323 
       
   324 (* theory data *)
       
   325 
       
   326 structure TheoryData =
       
   327 struct
       
   328 
       
   329 val declare = declare;
       
   330 
       
   331 fun get k dest thy =
       
   332   (case Inttab.lookup (data_of thy, k) of
       
   333     SOME x => (dest x handle Match =>
       
   334       error ("Failed to access theory data " ^ quote (invoke_name k)))
       
   335   | NONE => error ("Uninitialized theory data " ^ quote (invoke_name k)));
       
   336 
       
   337 fun print k thy = invoke_print thy k (get k I thy);
       
   338 fun put k mk x = map_thy (curry Inttab.update (k, mk x));
       
   339 fun init k = put k I (invoke_empty k);
       
   340 
       
   341 end;
       
   342 
       
   343 
       
   344 (* pre_pure *)
       
   345 
       
   346 val pre_pure = create_thy draftN NONE (serial (), draftN) Inttab.empty
       
   347   Inttab.empty (make_history ProtoPureN 0 []) (make_ancestry [] []);
       
   348 
   282 
   349 
   283 
   350 (* trivial merge *)
   284 (* trivial merge *)
   351 
   285 
   352 fun merge (thy1, thy2) =
   286 fun merge (thy1, thy2) =
   357       str_of_thy thy1, str_of_thy thy2], []));
   291       str_of_thy thy1, str_of_thy thy2], []));
   358 
   292 
   359 fun merge_refs (ref1, ref2) = self_ref (merge (deref ref1, deref ref2));
   293 fun merge_refs (ref1, ref2) = self_ref (merge (deref ref1, deref ref2));
   360 
   294 
   361 
   295 
   362 (* non-trivial merge *)
   296 
   363 
   297 (** build theories **)
   364 fun merge_internal pp (thy1, thy2) =
   298 
       
   299 (* primitives *)
       
   300 
       
   301 fun create_thy name self id ids iids data ancestry history =
       
   302   let
       
   303     val intermediate = #version history > 0;
       
   304     val (ids', iids') = check_insert intermediate id (ids, iids);
       
   305     val id' = (serial (), name);
       
   306     val _ = check_insert intermediate id' (ids', iids');
       
   307     val identity' = make_identity self id' ids' iids';
       
   308   in vitalize (Theory (identity', data, ancestry, history)) end;
       
   309 
       
   310 fun change_thy name f (thy as Theory ({self, id, ids, iids}, data, ancestry, history)) =
       
   311   let
       
   312     val _ = check_thy "Context.change_thy" thy;
       
   313     val (self', data', ancestry') =
       
   314       if is_draft thy then (self, data, ancestry)    (*destructive change!*)
       
   315       else if #version history > 0
       
   316       then (NONE, copy_data data, ancestry)
       
   317       else (NONE, extend_data data, make_ancestry [thy] (thy :: #ancestors ancestry));
       
   318     val data'' = f data';
       
   319   in create_thy name self' id ids iids data'' ancestry' history end;
       
   320 
       
   321 fun name_thy name = change_thy name I;
       
   322 val modify_thy = change_thy draftN;
       
   323 val extend_thy = modify_thy I;
       
   324 
       
   325 fun copy_thy (thy as Theory ({id, ids, iids, ...}, data, ancestry, history)) =
       
   326   (check_thy "Context.copy_thy" thy;
       
   327     create_thy draftN NONE id ids iids (copy_data data) ancestry history);
       
   328 
       
   329 val pre_pure_thy = create_thy draftN NONE (serial (), draftN) Inttab.empty Inttab.empty
       
   330   Inttab.empty (make_ancestry [] []) (make_history ProtoPureN 0 []);
       
   331 
       
   332 
       
   333 (* named theory nodes *)
       
   334 
       
   335 fun merge_thys pp (thy1, thy2) =
   365   if subthy_internal (thy2, thy1) then thy1
   336   if subthy_internal (thy2, thy1) then thy1
   366   else if subthy_internal (thy1, thy2) then thy2
   337   else if subthy_internal (thy1, thy2) then thy2
   367   else if exists_name CPureN thy1 <> exists_name CPureN thy2 then
   338   else if exists_name CPureN thy1 <> exists_name CPureN thy2 then
   368     error "Cannot merge Pure and CPure developments"
   339     error "Cannot merge Pure and CPure developments"
   369   else
   340   else
   370     let
   341     let
   371       val ids = check_merge thy1 thy2;
   342       val (ids, iids) = check_merge thy1 thy2;
   372       val data = merge_data (pp thy1) (data_of thy1, data_of thy2);
   343       val data = merge_data (pp thy1) (data_of thy1, data_of thy2);
       
   344       val ancestry = make_ancestry [] [];
   373       val history = make_history "" 0 [];
   345       val history = make_history "" 0 [];
   374       val ancestry = make_ancestry [] [];
   346     in create_thy draftN NONE (serial (), draftN) ids iids data ancestry history end;
   375     in create_thy draftN NONE (serial (), draftN) ids data history ancestry end;
   347 
   376 
   348 fun begin_thy pp name imports =
   377 
   349   if name = draftN then error ("Illegal theory name: " ^ quote draftN)
   378 (* named theory nodes *)
       
   379 
       
   380 val theory_name = #name o history_of;
       
   381 
       
   382 fun begin_theory pp name imports =
       
   383   if name = draftN then
       
   384     error ("Illegal theory name: " ^ quote draftN)
       
   385   else if exists is_draft imports then
       
   386     error "Attempt to import draft theories"
       
   387   else
   350   else
   388     let
   351     let
   389       val parents = gen_distinct eq_thy imports;
   352       val parents = gen_distinct eq_thy imports;  (* FIXME maximals *)
   390       val ancestors = gen_distinct eq_thy (parents @ List.concat (map ancestors_of parents));
   353       val ancestors = gen_distinct eq_thy (parents @ List.concat (map ancestors_of parents));
   391       val Theory ({id, ids, ...}, data, _, _) =
   354       val Theory ({id, ids, iids, ...}, data, _, _) =
   392         (case parents of
   355         (case parents of
   393           [] => error "No parent theories"
   356           [] => error "No parent theories"
   394         | thy :: thys => extend_thy (Library.foldl (merge_internal pp) (thy, thys)));
   357         | thy :: thys => extend_thy (Library.foldl (merge_thys pp) (thy, thys)));
       
   358       val ancestry = make_ancestry parents ancestors;
   395       val history = make_history name 0 [];
   359       val history = make_history name 0 [];
   396       val ancestry = make_ancestry parents ancestors;
   360     in create_thy draftN NONE id ids iids data ancestry history end;
   397     in create_thy draftN NONE id ids data history ancestry end;
   361 
   398 
   362 
   399 fun end_theory thy =
   363 (* undoable checkpoints *)
   400   thy
   364 
   401 (*|> purge_thy  FIXME *)
   365 fun checkpoint_thy thy =
   402   |> name_thy (theory_name thy);
   366   if not (is_draft thy) then thy
   403 
   367   else
   404 
   368     let
   405 
   369       val {name, version, intermediates} = history_of thy;
   406 
   370       val thy' as Theory (identity', data', ancestry', _) =
   407 (*** ML theory context ***)
   371         name_thy (name ^ ":" ^ string_of_int version) thy;
       
   372       val history' = make_history name (version + 1) (thy' :: intermediates);
       
   373     in vitalize (Theory (identity', data', ancestry', history')) end;
       
   374 
       
   375 fun finish_thy thy =
       
   376   let
       
   377     val {name, version, intermediates} = history_of thy;
       
   378     val rs = map (the o #self o identity_of o check_thy "Context.finish_thy") intermediates;
       
   379     val thy' as Theory ({self, id, ids, ...}, data', ancestry', _) = name_thy name thy;
       
   380     val identity' = make_identity self id ids Inttab.empty;
       
   381     val history' = make_history name 0 [];
       
   382     val thy'' = vitalize (Theory (identity', data', ancestry', history'));
       
   383   in List.app (fn r => r := thy'') rs; thy'' end;
       
   384 
       
   385 
       
   386 (* theory data *)
       
   387 
       
   388 fun theory_data thy =
       
   389   map invoke_name (Inttab.keys (data_of thy))
       
   390   |> map (rpair ()) |> Symtab.make_multi |> Symtab.dest
       
   391   |> map (apsnd length)
       
   392   |> map (fn (name, 1) => name | (name, n) => name ^ enclose "[" "]" (string_of_int n));
       
   393 
       
   394 structure TheoryData =
       
   395 struct
       
   396 
       
   397 val declare = declare_theory_data;
       
   398 
       
   399 fun get k dest thy =
       
   400   (case Inttab.lookup (data_of thy, k) of
       
   401     SOME x => (dest x handle Match =>
       
   402       error ("Failed to access theory data " ^ quote (invoke_name k)))
       
   403   | NONE => error ("Uninitialized theory data " ^ quote (invoke_name k)));
       
   404 
       
   405 fun put k mk x = modify_thy (curry Inttab.update (k, mk x));
       
   406 fun init k = put k I (invoke_empty k);
       
   407 
       
   408 end;
       
   409 
       
   410 
       
   411 
       
   412 (** ML theory context **)
   408 
   413 
   409 local
   414 local
   410   val current_theory = ref (NONE: theory option);
   415   val current_theory = ref (NONE: theory option);
   411 in
   416 in
   412   fun get_context () = ! current_theory;
   417   fun get_context () = ! current_theory;
   487 sig
   492 sig
   488   type T
   493   type T
   489   val init: theory -> theory
   494   val init: theory -> theory
   490   val print: theory -> unit
   495   val print: theory -> unit
   491   val get: theory -> T
   496   val get: theory -> T
       
   497   val get_sg: theory -> T    (*obsolete*)
   492   val put: T -> theory -> theory
   498   val put: T -> theory -> theory
   493   val map: (T -> T) -> theory -> theory
   499   val map: (T -> T) -> theory -> theory
   494 end;
   500 end;
   495 
   501 
   496 functor TheoryDataFun(Data: THEORY_DATA_ARGS): THEORY_DATA =
   502 functor TheoryDataFun(Data: THEORY_DATA_ARGS): THEORY_DATA =
   503 
   509 
   504 val kind = TheoryData.declare Data.name
   510 val kind = TheoryData.declare Data.name
   505   (Data Data.empty)
   511   (Data Data.empty)
   506   (fn Data x => Data (Data.copy x))
   512   (fn Data x => Data (Data.copy x))
   507   (fn Data x => Data (Data.extend x))
   513   (fn Data x => Data (Data.extend x))
   508   (fn pp => fn (Data x1, Data x2) => Data (Data.merge pp (x1, x2)))
   514   (fn pp => fn (Data x1, Data x2) => Data (Data.merge pp (x1, x2)));
   509   (fn thy => fn Data x => Data.print thy x);
       
   510 
   515 
   511 val init = TheoryData.init kind;
   516 val init = TheoryData.init kind;
   512 val print = TheoryData.print kind;
       
   513 val get = TheoryData.get kind (fn Data x => x);
   517 val get = TheoryData.get kind (fn Data x => x);
   514 val get_sg = get;    (*obsolete*)
   518 val get_sg = get;
       
   519 fun print thy = Data.print thy (get thy);
   515 val put = TheoryData.put kind Data;
   520 val put = TheoryData.put kind Data;
   516 fun map f thy = put (f (get thy)) thy;
   521 fun map f thy = put (f (get thy)) thy;
   517 
   522 
   518 end;
   523 end;
   519 
   524