src/Pure/Isar/entity.ML
author wenzelm
Thu, 12 May 2016 22:06:18 +0200
changeset 63090 7aa9ac5165e4
child 77970 31ea5c1f874d
permissions -rw-r--r--
common entity definitions within a global or local theory context;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63090
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/entity.ML
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
     3
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
     4
Entity definitions within a global or local theory context.
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
     5
*)
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
     6
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
     7
signature ENTITY =
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
     8
sig
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
     9
  type 'a data_ops =
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    10
   {get_data: Context.generic -> 'a Name_Space.table,
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    11
    put_data: 'a Name_Space.table -> Context.generic -> Context.generic}
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    12
  val define_global: 'a data_ops -> binding -> 'a -> theory -> string * theory
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    13
  val define: 'a data_ops -> binding -> 'a -> local_theory -> string * local_theory
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    14
end;
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    15
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    16
structure Entity: ENTITY =
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    17
struct
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    18
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    19
(* context data *)
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    20
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    21
type 'a data_ops =
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    22
 {get_data: Context.generic -> 'a Name_Space.table,
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    23
  put_data: 'a Name_Space.table -> Context.generic -> Context.generic};
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    24
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    25
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    26
(* global definition (foundation) *)
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    27
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    28
fun define_global {get_data, put_data} b x thy =
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    29
  let
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    30
    val context = Context.Theory thy;
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    31
    val (name, data') = Name_Space.define context true (b, x) (get_data context);
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    32
  in (name, Context.the_theory (put_data data' context)) end;
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    33
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    34
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    35
(* local definition *)
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    36
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    37
fun alias {get_data, put_data} binding name =
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    38
  Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => fn context =>
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    39
    let
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    40
      val naming = Name_Space.naming_of context;
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    41
      val binding' = Morphism.binding phi binding;
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    42
      val data' = Name_Space.alias_table naming binding' name (get_data context);
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    43
    in put_data data' context end);
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    44
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    45
fun transfer {get_data, put_data} ctxt =
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    46
  let
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    47
    val data0 = get_data (Context.Theory (Proof_Context.theory_of ctxt));
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    48
    val data' = Name_Space.merge_tables (data0, get_data (Context.Proof ctxt));
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    49
  in Context.proof_map (put_data data') ctxt end;
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    50
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    51
fun define ops binding x =
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    52
  Local_Theory.background_theory_result (define_global ops binding x)
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    53
  #-> (fn name =>
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    54
    Local_Theory.map_contexts (K (transfer ops))
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    55
    #> alias ops binding name
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    56
    #> pair name);
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    57
7aa9ac5165e4 common entity definitions within a global or local theory context;
wenzelm
parents:
diff changeset
    58
end;