src/Pure/Isar/locale.ML
author ballarin
Tue, 29 Sep 2009 22:15:54 +0200
changeset 32804 ca430e6aee1c
parent 32803 fc02b4b9eccc
parent 32801 6f97a67e8da8
child 32845 d2d0b9b1a69d
permissions -rw-r--r--
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
     1
(*  Title:      Pure/Isar/locale.ML
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
     2
    Author:     Clemens Ballarin, TU Muenchen
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
     3
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
     4
Locales -- managed Isar proof contexts, based on Pure predicates.
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
     5
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
     6
Draws basic ideas from Florian Kammueller's original version of
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
     7
locales, but uses the richer infrastructure of Isar instead of the raw
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
     8
meta-logic.  Furthermore, structured import of contexts (with merge
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
     9
and rename operations) are provided, as well as type-inference of the
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    10
signature parts, and predicate definitions of the specification text.
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    11
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    12
Interpretation enables the reuse of theorems of locales in other
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    13
contexts, namely those defined by theories, structured proofs and
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    14
locales themselves.
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    15
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    16
See also:
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    17
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    18
[1] Clemens Ballarin. Locales and Locale Expressions in Isabelle/Isar.
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    19
    In Stefano Berardi et al., Types for Proofs and Programs: International
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    20
    Workshop, TYPES 2003, Torino, Italy, LNCS 3085, pages 34-50, 2004.
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    21
[2] Clemens Ballarin. Interpretation of Locales in Isabelle: Managing
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    22
    Dependencies between Locales. Technical Report TUM-I0607, Technische
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    23
    Universitaet Muenchen, 2006.
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    24
[3] Clemens Ballarin. Interpretation of Locales in Isabelle: Theories and
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    25
    Proof Contexts. In J.M. Borwein and W.M. Farmer, MKM 2006, LNAI 4108,
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    26
    pages 31-43, 2006.
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    27
*)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    28
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    29
signature LOCALE =
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    30
sig
29576
669b560fc2b9 wrecked old locale package and related modules
haftmann
parents: 29544
diff changeset
    31
  (* Locale specification *)
30344
10a67c5ddddb more uniform handling of binding in targets and derived elements;
wenzelm
parents: 30223
diff changeset
    32
  val register_locale: binding ->
30755
7ef503d216c2 simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
wenzelm
parents: 30754
diff changeset
    33
    (string * sort) list * ((string * typ) * mixfix) list ->
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    34
    term option * term list ->
29441
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
    35
    thm option * thm option -> thm list ->
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
    36
    declaration list * declaration list ->
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
    37
    (string * (Attrib.binding * (thm list * Attrib.src list) list) list) list ->
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
    38
    (string * morphism) list -> theory -> theory
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    39
  val intern: theory -> xstring -> string
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    40
  val extern: theory -> string -> xstring
29392
aa3a30b625d1 tuned signature; internal code reorganisation
haftmann
parents: 29363
diff changeset
    41
  val defined: theory -> string -> bool
30755
7ef503d216c2 simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
wenzelm
parents: 30754
diff changeset
    42
  val params_of: theory -> string -> ((string * typ) * mixfix) list
29441
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
    43
  val intros_of: theory -> string -> thm option * thm option
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
    44
  val axioms_of: theory -> string -> thm list
29544
bc50244cd1d7 exported depedencies; tuned signature
haftmann
parents: 29502
diff changeset
    45
  val instance_of: theory -> string -> morphism -> term list
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    46
  val specification_of: theory -> string -> term option * term list
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    47
  val declarations_of: theory -> string -> declaration list * declaration list
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    48
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    49
  (* Storing results *)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    50
  val add_thmss: string -> string -> (Attrib.binding * (thm list * Attrib.src list) list) list ->
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    51
    Proof.context -> Proof.context
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    52
  val add_type_syntax: string -> declaration -> Proof.context -> Proof.context
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    53
  val add_term_syntax: string -> declaration -> Proof.context -> Proof.context
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    54
  val add_declaration: string -> declaration -> Proof.context -> Proof.context
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    55
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    56
  (* Activation *)
30764
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
    57
  val activate_declarations: string * morphism -> Proof.context -> Proof.context
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
    58
  val activate_facts: string * morphism -> Context.generic -> Context.generic
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    59
  val init: string -> theory -> Proof.context
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    60
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    61
  (* Reasoning about locales *)
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
    62
  val get_witnesses: Proof.context -> thm list
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
    63
  val get_intros: Proof.context -> thm list
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
    64
  val get_unfolds: Proof.context -> thm list
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
    65
  val witness_add: attribute
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
    66
  val intro_add: attribute
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
    67
  val unfold_add: attribute
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    68
  val intro_locales_tac: bool -> Proof.context -> thm list -> tactic
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    69
31988
801aabf9f376 tuned locale interface
haftmann
parents: 30777
diff changeset
    70
  (* Registrations and dependencies *)
32804
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
    71
  val add_registration: string * morphism -> morphism * bool -> morphism -> theory -> theory
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
    72
  val amend_registration: string * morphism -> morphism * bool -> morphism -> theory -> theory
32113
bafffa63ebfd dropped add_registration interface in locale
haftmann
parents: 32074
diff changeset
    73
  val add_registration_eqs: string * morphism -> thm list -> morphism -> theory -> theory
32800
57fcca4e7c0e Improved comments and api names.
ballarin
parents: 32113
diff changeset
    74
  val amend_registration_legacy: morphism -> string * morphism -> theory -> theory
32074
76d6ba08a05f simplification of locale interfaces
haftmann
parents: 31988
diff changeset
    75
  val add_dependency: string -> string * morphism -> morphism -> theory -> theory
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    76
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    77
  (* Diagnostic *)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    78
  val print_locales: theory -> unit
30344
10a67c5ddddb more uniform handling of binding in targets and derived elements;
wenzelm
parents: 30223
diff changeset
    79
  val print_locale: theory -> bool -> xstring -> unit
32804
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
    80
  val print_registrations: theory -> string -> unit
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    81
end;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    82
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    83
structure Locale: LOCALE =
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    84
struct
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    85
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    86
datatype ctxt = datatype Element.ctxt;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    87
29392
aa3a30b625d1 tuned signature; internal code reorganisation
haftmann
parents: 29363
diff changeset
    88
aa3a30b625d1 tuned signature; internal code reorganisation
haftmann
parents: 29363
diff changeset
    89
(*** Theory data ***)
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    90
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    91
datatype locale = Loc of {
29392
aa3a30b625d1 tuned signature; internal code reorganisation
haftmann
parents: 29363
diff changeset
    92
  (** static part **)
30755
7ef503d216c2 simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
wenzelm
parents: 30754
diff changeset
    93
  parameters: (string * sort) list * ((string * typ) * mixfix) list,
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    94
    (* type and term parameters *)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    95
  spec: term option * term list,
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
    96
    (* assumptions (as a single predicate expression) and defines *)
29441
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
    97
  intros: thm option * thm option,
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
    98
  axioms: thm list,
29392
aa3a30b625d1 tuned signature; internal code reorganisation
haftmann
parents: 29363
diff changeset
    99
  (** dynamic part **)
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   100
  decls: (declaration * stamp) list * (declaration * stamp) list,
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   101
    (* type and term syntax declarations *)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   102
  notes: ((string * (Attrib.binding * (thm list * Attrib.src list) list) list) * stamp) list,
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   103
    (* theorem declarations *)
29544
bc50244cd1d7 exported depedencies; tuned signature
haftmann
parents: 29502
diff changeset
   104
  dependencies: ((string * morphism) * stamp) list
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   105
    (* locale dependencies (sublocale relation) *)
29392
aa3a30b625d1 tuned signature; internal code reorganisation
haftmann
parents: 29363
diff changeset
   106
};
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   107
29441
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
   108
fun mk_locale ((parameters, spec, intros, axioms), ((decls, notes), dependencies)) =
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
   109
  Loc {parameters = parameters, intros = intros, axioms = axioms, spec = spec,
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
   110
    decls = decls, notes = notes, dependencies = dependencies};
30754
c896167de3d5 minor tuning;
wenzelm
parents: 30725
diff changeset
   111
29441
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
   112
fun map_locale f (Loc {parameters, spec, intros, axioms, decls, notes, dependencies}) =
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
   113
  mk_locale (f ((parameters, spec, intros, axioms), ((decls, notes), dependencies)));
30754
c896167de3d5 minor tuning;
wenzelm
parents: 30725
diff changeset
   114
29441
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
   115
fun merge_locale (Loc {parameters, spec, intros, axioms, decls = (decls1, decls2),
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
   116
  notes, dependencies}, Loc {decls = (decls1', decls2'), notes = notes',
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
   117
    dependencies = dependencies', ...}) = mk_locale
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
   118
      ((parameters, spec, intros, axioms),
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
   119
        (((merge (eq_snd op =) (decls1, decls1'), merge (eq_snd op =) (decls2, decls2')),
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
   120
          merge (eq_snd op =) (notes, notes')),
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
   121
            merge (eq_snd op =) (dependencies, dependencies')));
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   122
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   123
structure Locales = TheoryDataFun
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   124
(
29392
aa3a30b625d1 tuned signature; internal code reorganisation
haftmann
parents: 29363
diff changeset
   125
  type T = locale NameSpace.table;
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   126
  val empty = NameSpace.empty_table;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   127
  val copy = I;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   128
  val extend = I;
29392
aa3a30b625d1 tuned signature; internal code reorganisation
haftmann
parents: 29363
diff changeset
   129
  fun merge _ = NameSpace.join_tables (K merge_locale);
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   130
);
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   131
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   132
val intern = NameSpace.intern o #1 o Locales.get;
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   133
val extern = NameSpace.extern o #1 o Locales.get;
29392
aa3a30b625d1 tuned signature; internal code reorganisation
haftmann
parents: 29363
diff changeset
   134
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   135
val get_locale = Symtab.lookup o #2 o Locales.get;
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   136
val defined = Symtab.defined o #2 o Locales.get;
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   137
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   138
fun the_locale thy name =
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   139
  (case get_locale thy name of
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   140
    SOME (Loc loc) => loc
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   141
  | NONE => error ("Unknown locale " ^ quote name));
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   142
30344
10a67c5ddddb more uniform handling of binding in targets and derived elements;
wenzelm
parents: 30223
diff changeset
   143
fun register_locale binding parameters spec intros axioms decls notes dependencies thy =
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   144
  thy |> Locales.map (NameSpace.define (Sign.naming_of thy)
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   145
    (binding,
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   146
      mk_locale ((parameters, spec, intros, axioms),
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   147
        ((pairself (map (fn decl => (decl, stamp ()))) decls, map (fn n => (n, stamp ())) notes),
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   148
          map (fn d => (d, stamp ())) dependencies))) #> snd);
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   149
29392
aa3a30b625d1 tuned signature; internal code reorganisation
haftmann
parents: 29363
diff changeset
   150
fun change_locale name =
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   151
  Locales.map o apsnd o Symtab.map_entry name o map_locale o apsnd;
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   152
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   153
fun print_locales thy =
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   154
  Pretty.strs ("locales:" :: map #1 (NameSpace.extern_table (Locales.get thy)))
29392
aa3a30b625d1 tuned signature; internal code reorganisation
haftmann
parents: 29363
diff changeset
   155
  |> Pretty.writeln;
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   156
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   157
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   158
(*** Primitive operations ***)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   159
29392
aa3a30b625d1 tuned signature; internal code reorganisation
haftmann
parents: 29363
diff changeset
   160
fun params_of thy = snd o #parameters o the_locale thy;
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   161
29441
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
   162
fun intros_of thy = #intros o the_locale thy;
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
   163
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
   164
fun axioms_of thy = #axioms o the_locale thy;
28d7d7572b81 explicit bookkeeping of intro rules and axiom
haftmann
parents: 29392
diff changeset
   165
29392
aa3a30b625d1 tuned signature; internal code reorganisation
haftmann
parents: 29363
diff changeset
   166
fun instance_of thy name morph = params_of thy name |>
30755
7ef503d216c2 simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
wenzelm
parents: 30754
diff changeset
   167
  map (Morphism.term morph o Free o #1);
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   168
29392
aa3a30b625d1 tuned signature; internal code reorganisation
haftmann
parents: 29363
diff changeset
   169
fun specification_of thy = #spec o the_locale thy;
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   170
29392
aa3a30b625d1 tuned signature; internal code reorganisation
haftmann
parents: 29363
diff changeset
   171
fun declarations_of thy name = the_locale thy name |>
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   172
  #decls |> pairself (map fst);
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   173
29544
bc50244cd1d7 exported depedencies; tuned signature
haftmann
parents: 29502
diff changeset
   174
fun dependencies_of thy name = the_locale thy name |>
bc50244cd1d7 exported depedencies; tuned signature
haftmann
parents: 29502
diff changeset
   175
  #dependencies |> map fst;
bc50244cd1d7 exported depedencies; tuned signature
haftmann
parents: 29502
diff changeset
   176
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   177
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   178
(*** Activate context elements of locale ***)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   179
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   180
(** Identifiers: activated locales in theory or proof context **)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   181
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   182
fun ident_eq thy ((n: string, ts), (m, ss)) = (m = n) andalso Pattern.matchess thy (ss, ts);
32804
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   183
(* FIXME: this is ident_le, smaller term is more general *)
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   184
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   185
local
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   186
30754
c896167de3d5 minor tuning;
wenzelm
parents: 30725
diff changeset
   187
datatype 'a delayed = Ready of 'a | ToDo of 'a delayed * 'a delayed;
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   188
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   189
structure Identifiers = GenericDataFun
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   190
(
30754
c896167de3d5 minor tuning;
wenzelm
parents: 30725
diff changeset
   191
  type T = (string * term list) list delayed;
c896167de3d5 minor tuning;
wenzelm
parents: 30725
diff changeset
   192
  val empty = Ready [];
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   193
  val extend = I;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   194
  fun merge _ = ToDo;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   195
);
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   196
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   197
fun finish thy (ToDo (i1, i2)) = merge (ident_eq thy) (finish thy i1, finish thy i2)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   198
  | finish _ (Ready ids) = ids;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   199
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   200
val _ = Context.>> (Context.map_theory (Theory.at_begin (fn thy =>
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   201
  (case Identifiers.get (Context.Theory thy) of
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   202
    Ready _ => NONE
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   203
  | ids => SOME (Context.theory_map (Identifiers.put (Ready (finish thy ids))) thy)))));
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   204
30764
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   205
in
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   206
30764
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   207
val get_idents = (fn Ready ids => ids) o Identifiers.get;
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   208
val put_idents = Identifiers.put o Ready;
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   209
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   210
end;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   211
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   212
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   213
(** Resolve locale dependencies in a depth-first fashion **)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   214
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   215
local
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   216
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   217
val roundup_bound = 120;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   218
32803
fc02b4b9eccc Archive registrations by external view.
ballarin
parents: 32800
diff changeset
   219
fun add thy depth export (name, morph) (deps, marked) =
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   220
  if depth > roundup_bound
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   221
  then error "Roundup bound exceeded (sublocale relation probably not terminating)."
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   222
  else
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   223
    let
30754
c896167de3d5 minor tuning;
wenzelm
parents: 30725
diff changeset
   224
      val dependencies = dependencies_of thy name;
32803
fc02b4b9eccc Archive registrations by external view.
ballarin
parents: 32800
diff changeset
   225
      val instance = instance_of thy name (morph $> export);
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   226
    in
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   227
      if member (ident_eq thy) marked (name, instance)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   228
      then (deps, marked)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   229
      else
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   230
        let
30754
c896167de3d5 minor tuning;
wenzelm
parents: 30725
diff changeset
   231
          val dependencies' = map (fn (name, morph') => (name, morph' $> morph)) dependencies;
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   232
          val marked' = (name, instance) :: marked;
32803
fc02b4b9eccc Archive registrations by external view.
ballarin
parents: 32800
diff changeset
   233
          val (deps', marked'') = fold_rev (add thy (depth + 1) export) dependencies' ([], marked');
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   234
        in
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   235
          ((name, morph) :: deps' @ deps, marked'')
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   236
        end
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   237
    end;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   238
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   239
in
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   240
32804
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   241
(* Note that while identifiers always have the external (exported) view, activate_dep is
32803
fc02b4b9eccc Archive registrations by external view.
ballarin
parents: 32800
diff changeset
   242
  is presented with the internal view. *)
fc02b4b9eccc Archive registrations by external view.
ballarin
parents: 32800
diff changeset
   243
fc02b4b9eccc Archive registrations by external view.
ballarin
parents: 32800
diff changeset
   244
fun roundup thy activate_dep export (name, morph) (marked, input) =
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   245
  let
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   246
    (* Find all dependencies incuding new ones (which are dependencies enriching
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   247
      existing registrations). *)
32803
fc02b4b9eccc Archive registrations by external view.
ballarin
parents: 32800
diff changeset
   248
    val (dependencies, marked') = add thy 0 export (name, morph) ([], []);
32800
57fcca4e7c0e Improved comments and api names.
ballarin
parents: 32113
diff changeset
   249
    (* Filter out fragments from marked; these won't be activated. *)
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   250
    val dependencies' = filter_out (fn (name, morph) =>
32803
fc02b4b9eccc Archive registrations by external view.
ballarin
parents: 32800
diff changeset
   251
      member (ident_eq thy) marked (name, instance_of thy name (morph $> export))) dependencies;
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   252
  in
30773
6cc9964436c3 simplified roundup activation interface;
wenzelm
parents: 30764
diff changeset
   253
    (merge (ident_eq thy) (marked, marked'), input |> fold_rev activate_dep dependencies')
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   254
  end;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   255
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   256
end;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   257
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   258
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   259
(* Declarations, facts and entire locale content *)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   260
30764
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   261
fun activate_decls (name, morph) context =
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   262
  let
30764
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   263
    val thy = Context.theory_of context;
29392
aa3a30b625d1 tuned signature; internal code reorganisation
haftmann
parents: 29363
diff changeset
   264
    val {decls = (typ_decls, term_decls), ...} = the_locale thy name;
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   265
  in
30764
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   266
    context
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   267
    |> fold_rev (fn (decl, _) => decl morph) typ_decls
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   268
    |> fold_rev (fn (decl, _) => decl morph) term_decls
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   269
  end;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   270
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   271
fun activate_notes activ_elem transfer thy (name, morph) input =
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   272
  let
29392
aa3a30b625d1 tuned signature; internal code reorganisation
haftmann
parents: 29363
diff changeset
   273
    val {notes, ...} = the_locale thy name;
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   274
    fun activate ((kind, facts), _) input =
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   275
      let
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   276
        val facts' = facts |> Element.facts_map (Element.morph_ctxt (transfer input $> morph))
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   277
      in activ_elem (Notes (kind, facts')) input end;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   278
  in
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   279
    fold_rev activate notes input
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   280
  end;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   281
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   282
fun activate_all name thy activ_elem transfer (marked, input) =
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   283
  let
30764
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   284
    val {parameters = (_, params), spec = (asm, defs), ...} = the_locale thy name;
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   285
    val input' = input |>
30755
7ef503d216c2 simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
wenzelm
parents: 30754
diff changeset
   286
      (not (null params) ?
7ef503d216c2 simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
wenzelm
parents: 30754
diff changeset
   287
        activ_elem (Fixes (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) params))) |>
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   288
      (* FIXME type parameters *)
30764
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   289
      (case asm of SOME A => activ_elem (Assumes [(Attrib.empty_binding, [(A, [])])]) | _ => I) |>
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   290
      (if not (null defs)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   291
        then activ_elem (Defines (map (fn def => (Attrib.empty_binding, (def, []))) defs))
30764
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   292
        else I);
30773
6cc9964436c3 simplified roundup activation interface;
wenzelm
parents: 30764
diff changeset
   293
    val activate = activate_notes activ_elem transfer thy;
30764
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   294
  in
32803
fc02b4b9eccc Archive registrations by external view.
ballarin
parents: 32800
diff changeset
   295
    roundup thy activate Morphism.identity (name, Morphism.identity) (marked, input')
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   296
  end;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   297
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   298
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   299
(** Public activation functions **)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   300
30775
71f777103225 added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
wenzelm
parents: 30773
diff changeset
   301
fun activate_declarations dep = Context.proof_map (fn context =>
30764
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   302
  let
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   303
    val thy = Context.theory_of context;
32803
fc02b4b9eccc Archive registrations by external view.
ballarin
parents: 32800
diff changeset
   304
  in roundup thy activate_decls Morphism.identity dep (get_idents context, context) |-> put_idents end);
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   305
30764
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   306
fun activate_facts dep context =
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   307
  let
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   308
    val thy = Context.theory_of context;
30775
71f777103225 added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
wenzelm
parents: 30773
diff changeset
   309
    val activate = activate_notes Element.init (Element.transfer_morphism o Context.theory_of) thy;
32803
fc02b4b9eccc Archive registrations by external view.
ballarin
parents: 32800
diff changeset
   310
  in roundup thy activate Morphism.identity dep (get_idents context, context) |-> put_idents end;
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   311
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   312
fun init name thy =
30775
71f777103225 added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
wenzelm
parents: 30773
diff changeset
   313
  activate_all name thy Element.init (Element.transfer_morphism o Context.theory_of)
30764
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   314
    ([], Context.Proof (ProofContext.init thy)) |-> put_idents |> Context.proof_of;
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   315
30764
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   316
fun print_locale thy show_facts raw_name =
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   317
  let
30764
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   318
    val name = intern thy raw_name;
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   319
    val ctxt = init name thy;
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   320
    fun cons_elem (elem as Notes _) = show_facts ? cons elem
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   321
      | cons_elem elem = cons elem;
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   322
    val elems =
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   323
      activate_all name thy cons_elem (K (Element.transfer_morphism thy)) ([], [])
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   324
      |> snd |> rev;
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   325
  in
30764
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   326
    Pretty.big_list "locale elements:" (map (Pretty.chunks o Element.pretty_ctxt ctxt) elems)
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   327
    |> Pretty.writeln
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   328
  end;
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   329
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   330
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   331
(*** Registrations: interpretations in theories ***)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   332
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   333
structure Registrations = TheoryDataFun
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   334
(
32801
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   335
  type T = ((string * (morphism * morphism)) * stamp) list *
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   336
    (* registrations, in reverse order of declaration *)
32801
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   337
    (stamp * ((morphism * bool) * stamp) list) list;
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   338
    (* alist of mixin lists, per list mixins in reverse order of declaration *)
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   339
  val empty = ([], []);
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   340
  val extend = I;
29392
aa3a30b625d1 tuned signature; internal code reorganisation
haftmann
parents: 29363
diff changeset
   341
  val copy = I;
32801
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   342
  fun merge _ ((r1, m1), (r2, m2)) : T =
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   343
    (Library.merge (eq_snd op =) (r1, r2),
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   344
      AList.join (op =) (K (Library.merge (eq_snd op =))) (m1, m2));
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   345
    (* FIXME consolidate with dependencies, consider one data slot only *)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   346
);
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   347
32801
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   348
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   349
(* Primitive operations *)
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   350
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   351
fun compose_mixins mixins =
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   352
  fold_rev Morphism.compose (map (fst o fst) mixins) Morphism.identity;
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   353
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   354
fun reg_morph mixins ((name, (base, export)), stamp) =
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   355
  let val mix = the_default [] (AList.lookup (op =) mixins stamp) |> compose_mixins;
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   356
  in (name, base $> mix $> export) end;
31988
801aabf9f376 tuned locale interface
haftmann
parents: 30777
diff changeset
   357
801aabf9f376 tuned locale interface
haftmann
parents: 30777
diff changeset
   358
fun these_registrations thy name = Registrations.get thy
32801
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   359
  |>> filter (curry (op =) name o fst o fst)
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   360
  |-> (fn regs => fn mixins => map (reg_morph mixins) regs);
31988
801aabf9f376 tuned locale interface
haftmann
parents: 30777
diff changeset
   361
801aabf9f376 tuned locale interface
haftmann
parents: 30777
diff changeset
   362
fun all_registrations thy = Registrations.get thy
32801
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   363
  |-> (fn regs => fn mixins => map (reg_morph mixins) regs);
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   364
32804
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   365
fun get_mixins thy (name, morph) =
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   366
  let
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   367
    val (regs, mixins) = Registrations.get thy;
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   368
  in
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   369
    case find_first (fn ((name', (morph', export')), _) => ident_eq thy
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   370
      ((name', instance_of thy name' (morph' $> export')), (name, instance_of thy name morph))) (rev regs) of
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   371
      NONE => []
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   372
    | SOME (_, stamp) => the_default [] (AList.lookup (op =) mixins stamp)
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   373
  end;
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   374
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   375
fun collect_mixins thy (name, morph) =
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   376
  roundup thy (fn dep => fn mixins =>
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   377
    merge (eq_snd op =) (mixins, get_mixins thy dep)) Morphism.identity (name, morph) ([], [])
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   378
  |> snd |> filter (snd o fst);  (* only inheritable mixins *)
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   379
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   380
fun activate_notes' activ_elem transfer thy export (name, morph) input =
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   381
  let
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   382
    val {notes, ...} = the_locale thy name;
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   383
    fun activate ((kind, facts), _) input =
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   384
      let
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   385
        val mixin = collect_mixins thy (name, morph $> export) |> compose_mixins;
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   386
        val facts' = facts |> Element.facts_map (Element.morph_ctxt (transfer input $> morph $> mixin $> export))
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   387
      in activ_elem (Notes (kind, facts')) input end;
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   388
  in
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   389
    fold_rev activate notes input
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   390
  end;
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   391
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   392
fun activate_facts' export dep context =
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   393
  let
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   394
    val thy = Context.theory_of context;
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   395
    val activate = activate_notes' Element.init (Element.transfer_morphism o Context.theory_of) thy export;
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   396
  in roundup thy activate export dep (get_idents context, context) |-> put_idents end;
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   397
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   398
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   399
(* Diagnostic *)
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   400
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   401
fun print_registrations thy raw_name =
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   402
  let
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   403
    val name = intern thy raw_name;
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   404
    val name' = extern thy name;
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   405
    val ctxt = ProofContext.init thy;
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   406
    fun prt_qual (qual, mand) = Pretty.str (qual ^ (if mand then "!" else "?"));
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   407
    fun prt_quals qs = Pretty.separate "." (map prt_qual qs) |> Pretty.block;
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   408
    val prt_term = Pretty.quote o Syntax.pretty_term ctxt;
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   409
    fun prt_term' t = if !show_types
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   410
      then Pretty.block [prt_term t, Pretty.brk 1, Pretty.str "::",
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   411
        Pretty.brk 1, (Pretty.quote o Syntax.pretty_typ ctxt) (type_of t)]
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   412
      else prt_term t;
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   413
    fun prt_inst ts =
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   414
      Pretty.block (Pretty.breaks (Pretty.str name' :: map prt_term' ts));
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   415
    fun prt_reg (name, morph) =
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   416
      let
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   417
        val qs = Binding.name "x" |> Morphism.binding morph |> Binding.prefix_of;
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   418
        val ts = instance_of thy name morph;
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   419
      in
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   420
        case qs of
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   421
           [] => prt_inst ts
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   422
         | qs => Pretty.block [prt_quals qs, Pretty.brk 1, Pretty.str ":",
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   423
             Pretty.brk 1, prt_inst ts]
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   424
      end;
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   425
  in
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   426
    (case these_registrations thy name of
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   427
        [] => Pretty.str ("no interpretations")
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   428
      | regs => Pretty.big_list "interpretations:" (map prt_reg regs))
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   429
    |> Pretty.writeln
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   430
  end;
32801
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   431
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   432
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   433
(* Add and extend registrations *)
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   434
32804
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   435
fun amend_registration (name, morph) mixin export thy =
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   436
  let
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   437
    val regs = Registrations.get thy |> fst;
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   438
    val base = instance_of thy name (morph $> export);
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   439
    fun match ((name', (morph', export')), _) =
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   440
      name = name' andalso eq_list (op aconv) (base, instance_of thy name' (morph' $> export'));
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   441
  in
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   442
    case find_first match (rev regs) of
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   443
        NONE => error ("No interpretation of locale " ^
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   444
          quote (extern thy name) ^ " and\nparameter instantiation " ^
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   445
          space_implode " " (map (quote o Syntax.string_of_term_global thy) base) ^
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   446
          " available")
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   447
      | SOME (_, stamp') => Registrations.map 
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   448
          (apsnd (AList.map_default (op =) (stamp', []) (cons (mixin, stamp ())))) thy
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   449
    (* FIXME deal with inheritance: propagate to existing children *)
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   450
  end;
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   451
32801
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   452
(* Note that a registration that would be subsumed by an existing one will not be
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   453
   generated, and it will not be possible to amend it. *)
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   454
32804
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   455
fun add_registration (name, base_morph) mixin export thy =
32801
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   456
  let
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   457
    val base = instance_of thy name base_morph;
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   458
  in
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   459
    if member (ident_eq thy) (get_idents (Context.Theory thy)) (name, base)
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   460
    then thy
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   461
    else
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   462
      let
32804
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   463
        fun add_reg (dep', morph') =
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   464
          let
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   465
            val mixins = collect_mixins thy (dep', morph' $> export);
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   466
            val stamp = stamp ();
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   467
          in
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   468
            Registrations.map (apfst (cons ((dep', (morph', export)), stamp))
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   469
              #> apsnd (cons (stamp, mixins)))
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   470
          end
32801
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   471
      in
32804
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   472
        (get_idents (Context.Theory thy), thy)
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   473
        (* add new registrations with inherited mixins *)
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   474
        |> roundup thy add_reg export (name, base_morph)
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   475
        |> snd
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   476
        (* add mixin *)
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   477
        |> amend_registration (name, base_morph) mixin export
32801
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   478
        (* activate import hierarchy as far as not already active *)
32804
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   479
        |> Context.theory_map (activate_facts' export (name, base_morph))
32801
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   480
      end
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   481
  end;
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   482
32800
57fcca4e7c0e Improved comments and api names.
ballarin
parents: 32113
diff changeset
   483
fun amend_registration_legacy morph (name, base_morph) thy =
57fcca4e7c0e Improved comments and api names.
ballarin
parents: 32113
diff changeset
   484
  (* legacy, never modify base morphism *)
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   485
  let
32801
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   486
    val regs = Registrations.get thy |> fst |> map fst;
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   487
    val base = instance_of thy name base_morph;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   488
    fun match (name', (morph', _)) =
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   489
      name = name' andalso eq_list (op aconv) (base, instance_of thy name' morph');
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   490
    val i = find_index match (rev regs);
30764
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   491
    val _ =
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   492
      if i = ~1 then error ("No registration of locale " ^
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   493
        quote (extern thy name) ^ " and parameter instantiation " ^
30764
3e3e7aa0cc7a simplified Locale.activate operations, using generic context;
wenzelm
parents: 30763
diff changeset
   494
        space_implode " " (map (quote o Syntax.string_of_term_global thy) base) ^ " available")
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   495
      else ();
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   496
  in
32801
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   497
    Registrations.map ((apfst o nth_map (length regs - 1 - i))
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   498
      (fn ((name, (base, export)), stamp) => ((name, (base $> morph, export)), stamp))) thy
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   499
  end;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   500
32074
76d6ba08a05f simplification of locale interfaces
haftmann
parents: 31988
diff changeset
   501
fun add_registration_eqs (dep, proto_morph) eqns export thy =
76d6ba08a05f simplification of locale interfaces
haftmann
parents: 31988
diff changeset
   502
  let
76d6ba08a05f simplification of locale interfaces
haftmann
parents: 31988
diff changeset
   503
    val morph = if null eqns then proto_morph
76d6ba08a05f simplification of locale interfaces
haftmann
parents: 31988
diff changeset
   504
      else proto_morph $> Element.eq_morphism thy eqns;
76d6ba08a05f simplification of locale interfaces
haftmann
parents: 31988
diff changeset
   505
  in
32113
bafffa63ebfd dropped add_registration interface in locale
haftmann
parents: 32074
diff changeset
   506
    (get_idents (Context.Theory thy), thy)
bafffa63ebfd dropped add_registration interface in locale
haftmann
parents: 32074
diff changeset
   507
    |> roundup thy (fn (dep', morph') =>
32804
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   508
        Registrations.map (apfst (cons ((dep', (morph', export)), stamp ())))) export (dep, morph)
32113
bafffa63ebfd dropped add_registration interface in locale
haftmann
parents: 32074
diff changeset
   509
    |> snd
32074
76d6ba08a05f simplification of locale interfaces
haftmann
parents: 31988
diff changeset
   510
    |> Context.theory_map (activate_facts (dep, morph $> export))
76d6ba08a05f simplification of locale interfaces
haftmann
parents: 31988
diff changeset
   511
  end;
76d6ba08a05f simplification of locale interfaces
haftmann
parents: 31988
diff changeset
   512
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   513
31988
801aabf9f376 tuned locale interface
haftmann
parents: 30777
diff changeset
   514
(*** Dependencies ***)
801aabf9f376 tuned locale interface
haftmann
parents: 30777
diff changeset
   515
32074
76d6ba08a05f simplification of locale interfaces
haftmann
parents: 31988
diff changeset
   516
fun add_dependency loc (dep, morph) export thy =
31988
801aabf9f376 tuned locale interface
haftmann
parents: 30777
diff changeset
   517
  thy
32074
76d6ba08a05f simplification of locale interfaces
haftmann
parents: 31988
diff changeset
   518
  |> (change_locale loc o apsnd) (cons ((dep, morph $> export), stamp ()))
32804
ca430e6aee1c Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
parents: 32803 32801
diff changeset
   519
  |> (fn thy => fold_rev (Context.theory_map o activate_facts' export)
31988
801aabf9f376 tuned locale interface
haftmann
parents: 30777
diff changeset
   520
      (all_registrations thy) thy);
32801
6f97a67e8da8 Explicit management of registration mixins.
ballarin
parents: 32800
diff changeset
   521
  (* FIXME deal with inheritance: propagate mixins to new children *)
31988
801aabf9f376 tuned locale interface
haftmann
parents: 30777
diff changeset
   522
801aabf9f376 tuned locale interface
haftmann
parents: 30777
diff changeset
   523
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   524
(*** Storing results ***)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   525
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   526
(* Theorems *)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   527
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   528
fun add_thmss loc kind args ctxt =
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   529
  let
30777
9960ff945c52 simplified Element.activate(_i): singleton version;
wenzelm
parents: 30775
diff changeset
   530
    val (Notes args', ctxt') = Element.activate_i (Notes (kind, args)) ctxt;
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   531
    val ctxt'' = ctxt' |> ProofContext.theory (
29392
aa3a30b625d1 tuned signature; internal code reorganisation
haftmann
parents: 29363
diff changeset
   532
      (change_locale loc o apfst o apsnd) (cons (args', stamp ()))
aa3a30b625d1 tuned signature; internal code reorganisation
haftmann
parents: 29363
diff changeset
   533
        #>
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   534
      (* Registrations *)
31988
801aabf9f376 tuned locale interface
haftmann
parents: 30777
diff changeset
   535
      (fn thy => fold_rev (fn (_, morph) =>
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   536
            let
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   537
              val args'' = snd args' |> Element.facts_map (Element.morph_ctxt morph) |>
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   538
                Attrib.map_facts (Attrib.attribute_i thy)
30438
c2d49315b93b eliminated qualified_names naming policy: qualified names are only permitted via explicit Binding.qualify/qualified_name etc. (NB: user-level outer syntax should never do this);
wenzelm
parents: 30344
diff changeset
   539
            in PureThy.note_thmss kind args'' #> snd end)
31988
801aabf9f376 tuned locale interface
haftmann
parents: 30777
diff changeset
   540
        (these_registrations thy loc) thy))
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   541
  in ctxt'' end;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   542
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   543
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   544
(* Declarations *)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   545
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   546
local
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   547
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   548
fun decl_attrib decl phi = Thm.declaration_attribute (K (decl phi));
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   549
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   550
fun add_decls add loc decl =
30223
24d975352879 renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
wenzelm
parents: 29576
diff changeset
   551
  ProofContext.theory ((change_locale loc o apfst o apfst) (add (decl, stamp ()))) #>
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   552
  add_thmss loc Thm.internalK
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   553
    [((Binding.empty, [Attrib.internal (decl_attrib decl)]), [([Drule.dummy_thm], [])])];
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   554
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   555
in
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   556
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   557
val add_type_syntax = add_decls (apfst o cons);
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   558
val add_term_syntax = add_decls (apsnd o cons);
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   559
val add_declaration = add_decls (K I);
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   560
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   561
end;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   562
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   563
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   564
(*** Reasoning about locales ***)
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   565
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   566
(* Storage for witnesses, intro and unfold rules *)
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   567
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   568
structure Thms = GenericDataFun
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   569
(
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   570
  type T = thm list * thm list * thm list;
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   571
  val empty = ([], [], []);
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   572
  val extend = I;
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   573
  fun merge _ ((witnesses1, intros1, unfolds1), (witnesses2, intros2, unfolds2)) =
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   574
   (Thm.merge_thms (witnesses1, witnesses2),
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   575
    Thm.merge_thms (intros1, intros2),
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   576
    Thm.merge_thms (unfolds1, unfolds2));
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   577
);
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   578
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   579
val get_witnesses = #1 o Thms.get o Context.Proof;
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   580
val get_intros = #2 o Thms.get o Context.Proof;
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   581
val get_unfolds = #3 o Thms.get o Context.Proof;
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   582
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   583
val witness_add =
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   584
  Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (Thm.add_thm th x, y, z)));
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   585
val intro_add =
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   586
  Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (x, Thm.add_thm th y, z)));
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   587
val unfold_add =
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   588
  Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (x, y, Thm.add_thm th z)));
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   589
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   590
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   591
(* Tactic *)
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   592
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   593
fun intro_locales_tac eager ctxt =
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   594
  Method.intros_tac
30725
c23a5b3cd1b9 register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents: 30585
diff changeset
   595
    (get_witnesses ctxt @ get_intros ctxt @ (if eager then get_unfolds ctxt else []));
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   596
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   597
val _ = Context.>> (Context.map_theory
30515
bca05b17b618 simplified method setup;
wenzelm
parents: 30510
diff changeset
   598
 (Method.setup (Binding.name "intro_locales") (Scan.succeed (METHOD o intro_locales_tac false))
bca05b17b618 simplified method setup;
wenzelm
parents: 30510
diff changeset
   599
    "back-chain introduction rules of locales without unfolding predicates" #>
bca05b17b618 simplified method setup;
wenzelm
parents: 30510
diff changeset
   600
  Method.setup (Binding.name "unfold_locales") (Scan.succeed (METHOD o intro_locales_tac true))
bca05b17b618 simplified method setup;
wenzelm
parents: 30510
diff changeset
   601
    "back-chain all introduction rules of locales"));
29361
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   602
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   603
end;
764d51ab0198 locale -> old_locale, new_locale -> locale
haftmann
parents:
diff changeset
   604