| author | haftmann | 
| Tue, 31 Aug 2010 14:21:06 +0200 | |
| changeset 38926 | 24f82786cc57 | 
| parent 38797 | abe92b33ac9f | 
| child 39134 | 917b4b6ba3d2 | 
| permissions | -rw-r--r-- | 
| 29361 | 1  | 
(* Title: Pure/Isar/locale.ML  | 
2  | 
Author: Clemens Ballarin, TU Muenchen  | 
|
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 | 5  | 
|
6  | 
Draws basic ideas from Florian Kammueller's original version of  | 
|
7  | 
locales, but uses the richer infrastructure of Isar instead of the raw  | 
|
8  | 
meta-logic. Furthermore, structured import of contexts (with merge  | 
|
9  | 
and rename operations) are provided, as well as type-inference of the  | 
|
10  | 
signature parts, and predicate definitions of the specification text.  | 
|
11  | 
||
12  | 
Interpretation enables the reuse of theorems of locales in other  | 
|
13  | 
contexts, namely those defined by theories, structured proofs and  | 
|
14  | 
locales themselves.  | 
|
15  | 
||
16  | 
See also:  | 
|
17  | 
||
18  | 
[1] Clemens Ballarin. Locales and Locale Expressions in Isabelle/Isar.  | 
|
19  | 
In Stefano Berardi et al., Types for Proofs and Programs: International  | 
|
20  | 
Workshop, TYPES 2003, Torino, Italy, LNCS 3085, pages 34-50, 2004.  | 
|
21  | 
[2] Clemens Ballarin. Interpretation of Locales in Isabelle: Managing  | 
|
22  | 
Dependencies between Locales. Technical Report TUM-I0607, Technische  | 
|
23  | 
Universitaet Muenchen, 2006.  | 
|
24  | 
[3] Clemens Ballarin. Interpretation of Locales in Isabelle: Theories and  | 
|
25  | 
Proof Contexts. In J.M. Borwein and W.M. Farmer, MKM 2006, LNAI 4108,  | 
|
26  | 
pages 31-43, 2006.  | 
|
27  | 
*)  | 
|
28  | 
||
29  | 
signature LOCALE =  | 
|
30  | 
sig  | 
|
| 29576 | 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 | 34  | 
term option * term list ->  | 
| 29441 | 35  | 
thm option * thm option -> thm list ->  | 
| 
35798
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
36  | 
declaration list ->  | 
| 
30725
 
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 | 39  | 
val intern: theory -> xstring -> string  | 
40  | 
val extern: theory -> string -> xstring  | 
|
| 29392 | 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 | 43  | 
val intros_of: theory -> string -> thm option * thm option  | 
44  | 
val axioms_of: theory -> string -> thm list  | 
|
| 29544 | 45  | 
val instance_of: theory -> string -> morphism -> term list  | 
| 29361 | 46  | 
val specification_of: theory -> string -> term option * term list  | 
47  | 
||
48  | 
(* Storing results *)  | 
|
49  | 
val add_thmss: string -> string -> (Attrib.binding * (thm list * Attrib.src list) list) list ->  | 
|
50  | 
Proof.context -> Proof.context  | 
|
51  | 
val add_declaration: string -> declaration -> Proof.context -> Proof.context  | 
|
| 
35798
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
52  | 
val add_syntax_declaration: string -> declaration -> Proof.context -> Proof.context  | 
| 29361 | 53  | 
|
54  | 
(* Activation *)  | 
|
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
55  | 
val activate_declarations: string * morphism -> Proof.context -> Proof.context  | 
| 
38316
 
88e774d09fbc
Revert performance improvement of 8ed3a5fb4d25 since it breaks notes element declarations.
 
ballarin 
parents: 
38211 
diff
changeset
 | 
56  | 
val activate_facts: morphism option -> string * morphism -> Context.generic -> Context.generic  | 
| 29361 | 57  | 
val init: string -> theory -> Proof.context  | 
58  | 
||
59  | 
(* Reasoning about locales *)  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
60  | 
val get_witnesses: Proof.context -> thm list  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
61  | 
val get_intros: Proof.context -> thm list  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
62  | 
val get_unfolds: Proof.context -> thm list  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
63  | 
val witness_add: attribute  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
64  | 
val intro_add: attribute  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
65  | 
val unfold_add: attribute  | 
| 29361 | 66  | 
val intro_locales_tac: bool -> Proof.context -> thm list -> tactic  | 
67  | 
||
| 31988 | 68  | 
(* Registrations and dependencies *)  | 
| 
32845
 
d2d0b9b1a69d
Avoid administrative overhead for identity mixins.
 
ballarin 
parents: 
32804 
diff
changeset
 | 
69  | 
val add_registration: string * morphism -> (morphism * bool) option ->  | 
| 38107 | 70  | 
morphism -> Context.generic -> Context.generic  | 
| 
32845
 
d2d0b9b1a69d
Avoid administrative overhead for identity mixins.
 
ballarin 
parents: 
32804 
diff
changeset
 | 
71  | 
val amend_registration: string * morphism -> morphism * bool ->  | 
| 38107 | 72  | 
morphism -> Context.generic -> Context.generic  | 
| 38111 | 73  | 
val registrations_of: Context.generic -> string -> (string * morphism) list  | 
| 32074 | 74  | 
val add_dependency: string -> string * morphism -> morphism -> theory -> theory  | 
| 29361 | 75  | 
|
76  | 
(* Diagnostic *)  | 
|
| 37471 | 77  | 
val all_locales: theory -> string list  | 
| 29361 | 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  | 
| 38109 | 80  | 
val print_registrations: Proof.context -> string -> unit  | 
| 
37897
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
81  | 
val locale_deps: theory ->  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
82  | 
    { params: ((string * typ) * mixfix) list, axioms: term list, registrations: term list list } Graph.T
 | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
83  | 
* term list list Symtab.table Symtab.table  | 
| 29361 | 84  | 
end;  | 
85  | 
||
86  | 
structure Locale: LOCALE =  | 
|
87  | 
struct  | 
|
88  | 
||
89  | 
datatype ctxt = datatype Element.ctxt;  | 
|
90  | 
||
| 29392 | 91  | 
|
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
92  | 
(*** Locales ***)  | 
| 29361 | 93  | 
|
94  | 
datatype locale = Loc of {
 | 
|
| 29392 | 95  | 
(** static part **)  | 
| 
30755
 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 
wenzelm 
parents: 
30754 
diff
changeset
 | 
96  | 
parameters: (string * sort) list * ((string * typ) * mixfix) list,  | 
| 29361 | 97  | 
(* type and term parameters *)  | 
98  | 
spec: term option * term list,  | 
|
99  | 
(* assumptions (as a single predicate expression) and defines *)  | 
|
| 29441 | 100  | 
intros: thm option * thm option,  | 
101  | 
axioms: thm list,  | 
|
| 29392 | 102  | 
(** dynamic part **)  | 
| 36096 | 103  | 
syntax_decls: (declaration * serial) list,  | 
| 
35798
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
104  | 
(* syntax declarations *)  | 
| 36096 | 105  | 
notes: ((string * (Attrib.binding * (thm list * Attrib.src list) list) list) * serial) list,  | 
| 29361 | 106  | 
(* theorem declarations *)  | 
| 36651 | 107  | 
dependencies: ((string * (morphism * morphism)) * serial) list  | 
| 29361 | 108  | 
(* locale dependencies (sublocale relation) *)  | 
| 29392 | 109  | 
};  | 
| 29361 | 110  | 
|
| 
35798
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
111  | 
fun mk_locale ((parameters, spec, intros, axioms), ((syntax_decls, notes), dependencies)) =  | 
| 29441 | 112  | 
  Loc {parameters = parameters, intros = intros, axioms = axioms, spec = spec,
 | 
| 
35798
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
113  | 
syntax_decls = syntax_decls, notes = notes, dependencies = dependencies};  | 
| 
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
114  | 
|
| 
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
115  | 
fun map_locale f (Loc {parameters, spec, intros, axioms, syntax_decls, notes, dependencies}) =
 | 
| 
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
116  | 
mk_locale (f ((parameters, spec, intros, axioms), ((syntax_decls, notes), dependencies)));  | 
| 30754 | 117  | 
|
| 
35798
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
118  | 
fun merge_locale  | 
| 
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
119  | 
 (Loc {parameters, spec, intros, axioms, syntax_decls, notes, dependencies},
 | 
| 
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
120  | 
  Loc {syntax_decls = syntax_decls', notes = notes', dependencies = dependencies', ...}) =
 | 
| 
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
121  | 
mk_locale  | 
| 29441 | 122  | 
((parameters, spec, intros, axioms),  | 
| 
35798
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
123  | 
((merge (eq_snd op =) (syntax_decls, syntax_decls'),  | 
| 29441 | 124  | 
merge (eq_snd op =) (notes, notes')),  | 
125  | 
merge (eq_snd op =) (dependencies, dependencies')));  | 
|
| 29361 | 126  | 
|
| 33522 | 127  | 
structure Locales = Theory_Data  | 
| 29361 | 128  | 
(  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
129  | 
type T = locale Name_Space.table;  | 
| 33159 | 130  | 
val empty : T = Name_Space.empty_table "locale";  | 
| 29361 | 131  | 
val extend = I;  | 
| 33522 | 132  | 
val merge = Name_Space.join_tables (K merge_locale);  | 
| 29361 | 133  | 
);  | 
134  | 
||
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
135  | 
val intern = Name_Space.intern o #1 o Locales.get;  | 
| 
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
136  | 
val extern = Name_Space.extern o #1 o Locales.get;  | 
| 29392 | 137  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
138  | 
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
 | 
139  | 
val defined = Symtab.defined o #2 o Locales.get;  | 
| 29361 | 140  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
141  | 
fun the_locale thy name =  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
142  | 
(case get_locale thy name of  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
143  | 
SOME (Loc loc) => loc  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
144  | 
  | NONE => error ("Unknown locale " ^ quote name));
 | 
| 29361 | 145  | 
|
| 
35798
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
146  | 
fun register_locale binding parameters spec intros axioms syntax_decls notes dependencies thy =  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
147  | 
thy |> Locales.map (Name_Space.define true (Sign.naming_of thy)  | 
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
148  | 
(binding,  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
149  | 
mk_locale ((parameters, spec, intros, axioms),  | 
| 36096 | 150  | 
((map (fn decl => (decl, serial ())) syntax_decls, map (fn n => (n, serial ())) notes),  | 
| 36651 | 151  | 
map (fn d => (d |> apsnd (rpair Morphism.identity), serial ())) dependencies))) #> snd);  | 
152  | 
(* FIXME *)  | 
|
| 29361 | 153  | 
|
| 29392 | 154  | 
fun change_locale name =  | 
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
155  | 
Locales.map o apsnd o Symtab.map_entry name o map_locale o apsnd;  | 
| 29361 | 156  | 
|
157  | 
||
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
158  | 
(** Primitive operations **)  | 
| 29361 | 159  | 
|
| 29392 | 160  | 
fun params_of thy = snd o #parameters o the_locale thy;  | 
| 29361 | 161  | 
|
| 29441 | 162  | 
fun intros_of thy = #intros o the_locale thy;  | 
163  | 
||
164  | 
fun axioms_of thy = #axioms o the_locale thy;  | 
|
165  | 
||
| 29392 | 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 | 168  | 
|
| 29392 | 169  | 
fun specification_of thy = #spec o the_locale thy;  | 
| 29361 | 170  | 
|
| 29544 | 171  | 
fun dependencies_of thy name = the_locale thy name |>  | 
172  | 
#dependencies |> map fst;  | 
|
173  | 
||
| 
37133
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
174  | 
(* Print instance and qualifiers *)  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
175  | 
|
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
176  | 
fun pretty_reg thy (name, morph) =  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
177  | 
let  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
178  | 
val name' = extern thy name;  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
179  | 
val ctxt = ProofContext.init_global thy;  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
180  | 
fun prt_qual (qual, mand) = Pretty.str (qual ^ (if mand then "!" else "?"));  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
181  | 
fun prt_quals qs = Pretty.separate "." (map prt_qual qs) |> Pretty.block;  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
182  | 
val prt_term = Pretty.quote o Syntax.pretty_term ctxt;  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
183  | 
fun prt_term' t = if !show_types  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
184  | 
then Pretty.block [prt_term t, Pretty.brk 1, Pretty.str "::",  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
185  | 
Pretty.brk 1, (Pretty.quote o Syntax.pretty_typ ctxt) (type_of t)]  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
186  | 
else prt_term t;  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
187  | 
fun prt_inst ts =  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
188  | 
Pretty.block (Pretty.breaks (Pretty.str name' :: map prt_term' ts));  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
189  | 
|
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
190  | 
val qs = Binding.name "x" |> Morphism.binding morph |> Binding.prefix_of;  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
191  | 
val ts = instance_of thy name morph;  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
192  | 
in  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
193  | 
case qs of  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
194  | 
[] => prt_inst ts  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
195  | 
| qs => Pretty.block [prt_quals qs, Pretty.brk 1, Pretty.str ":",  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
196  | 
Pretty.brk 1, prt_inst ts]  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
197  | 
end;  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
198  | 
|
| 29361 | 199  | 
|
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
200  | 
(*** Identifiers: activated locales in theory or proof context ***)  | 
| 29361 | 201  | 
|
| 
37105
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
202  | 
(* subsumption *)  | 
| 
37103
 
6ea25bb157e1
Consistently use equality for registration lookup.
 
ballarin 
parents: 
37102 
diff
changeset
 | 
203  | 
fun ident_le thy ((n: string, ts), (m, ss)) = (m = n) andalso Pattern.matchess thy (ss, ts);  | 
| 
37105
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
204  | 
(* smaller term is more general *)  | 
| 
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
205  | 
|
| 
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
206  | 
(* total order *)  | 
| 
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
207  | 
fun ident_ord ((n: string, ts), (m, ss)) =  | 
| 
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
208  | 
case fast_string_ord (m, n) of  | 
| 
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
209  | 
EQUAL => list_ord Term_Ord.fast_term_ord (ts, ss)  | 
| 
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
210  | 
| ord => ord;  | 
| 29361 | 211  | 
|
212  | 
local  | 
|
213  | 
||
| 30754 | 214  | 
datatype 'a delayed = Ready of 'a | ToDo of 'a delayed * 'a delayed;  | 
| 29361 | 215  | 
|
| 33519 | 216  | 
structure Identifiers = Generic_Data  | 
| 29361 | 217  | 
(  | 
| 30754 | 218  | 
type T = (string * term list) list delayed;  | 
219  | 
val empty = Ready [];  | 
|
| 29361 | 220  | 
val extend = I;  | 
| 33519 | 221  | 
val merge = ToDo;  | 
| 29361 | 222  | 
);  | 
223  | 
||
| 
37103
 
6ea25bb157e1
Consistently use equality for registration lookup.
 
ballarin 
parents: 
37102 
diff
changeset
 | 
224  | 
fun finish thy (ToDo (i1, i2)) = merge (ident_le thy) (finish thy i1, finish thy i2)  | 
| 29361 | 225  | 
| finish _ (Ready ids) = ids;  | 
226  | 
||
227  | 
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
 | 
228  | 
(case Identifiers.get (Context.Theory thy) of  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
229  | 
Ready _ => NONE  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
230  | 
| ids => SOME (Context.theory_map (Identifiers.put (Ready (finish thy ids))) thy)))));  | 
| 29361 | 231  | 
|
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
232  | 
in  | 
| 29361 | 233  | 
|
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
234  | 
val get_idents = (fn Ready ids => ids) o Identifiers.get;  | 
| 
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
235  | 
val put_idents = Identifiers.put o Ready;  | 
| 29361 | 236  | 
|
237  | 
end;  | 
|
238  | 
||
239  | 
||
240  | 
(** Resolve locale dependencies in a depth-first fashion **)  | 
|
241  | 
||
242  | 
local  | 
|
243  | 
||
244  | 
val roundup_bound = 120;  | 
|
245  | 
||
| 32803 | 246  | 
fun add thy depth export (name, morph) (deps, marked) =  | 
| 29361 | 247  | 
if depth > roundup_bound  | 
248  | 
then error "Roundup bound exceeded (sublocale relation probably not terminating)."  | 
|
249  | 
else  | 
|
250  | 
let  | 
|
| 30754 | 251  | 
val dependencies = dependencies_of thy name;  | 
| 32803 | 252  | 
val instance = instance_of thy name (morph $> export);  | 
| 29361 | 253  | 
in  | 
| 
37103
 
6ea25bb157e1
Consistently use equality for registration lookup.
 
ballarin 
parents: 
37102 
diff
changeset
 | 
254  | 
if member (ident_le thy) marked (name, instance)  | 
| 29361 | 255  | 
then (deps, marked)  | 
256  | 
else  | 
|
257  | 
let  | 
|
| 36651 | 258  | 
val dependencies' = map (fn (name, (morph', export')) => (name, morph' $> export' $> morph)) dependencies;  | 
| 29361 | 259  | 
val marked' = (name, instance) :: marked;  | 
| 32803 | 260  | 
val (deps', marked'') = fold_rev (add thy (depth + 1) export) dependencies' ([], marked');  | 
| 29361 | 261  | 
in  | 
262  | 
((name, morph) :: deps' @ deps, marked'')  | 
|
263  | 
end  | 
|
264  | 
end;  | 
|
265  | 
||
266  | 
in  | 
|
267  | 
||
| 33541 | 268  | 
(* Note that while identifiers always have the external (exported) view, activate_dep  | 
| 32803 | 269  | 
is presented with the internal view. *)  | 
270  | 
||
271  | 
fun roundup thy activate_dep export (name, morph) (marked, input) =  | 
|
| 29361 | 272  | 
let  | 
273  | 
(* Find all dependencies incuding new ones (which are dependencies enriching  | 
|
274  | 
existing registrations). *)  | 
|
| 32803 | 275  | 
val (dependencies, marked') = add thy 0 export (name, morph) ([], []);  | 
| 32800 | 276  | 
(* Filter out fragments from marked; these won't be activated. *)  | 
| 29361 | 277  | 
val dependencies' = filter_out (fn (name, morph) =>  | 
| 
37103
 
6ea25bb157e1
Consistently use equality for registration lookup.
 
ballarin 
parents: 
37102 
diff
changeset
 | 
278  | 
member (ident_le thy) marked (name, instance_of thy name (morph $> export))) dependencies;  | 
| 29361 | 279  | 
in  | 
| 
37103
 
6ea25bb157e1
Consistently use equality for registration lookup.
 
ballarin 
parents: 
37102 
diff
changeset
 | 
280  | 
(merge (ident_le thy) (marked, marked'), input |> fold_rev activate_dep dependencies')  | 
| 29361 | 281  | 
end;  | 
282  | 
||
283  | 
end;  | 
|
284  | 
||
285  | 
||
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
286  | 
(*** Registrations: interpretations in theories or proof contexts ***)  | 
| 29361 | 287  | 
|
| 
37105
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
288  | 
structure Idtab = Table(type key = string * term list val ord = ident_ord);  | 
| 
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
289  | 
|
| 38107 | 290  | 
structure Registrations = Generic_Data  | 
| 29361 | 291  | 
(  | 
| 
37105
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
292  | 
type T = ((morphism * morphism) * serial) Idtab.table *  | 
| 
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
293  | 
(* registrations, indexed by locale name and instance;  | 
| 36090 | 294  | 
serial points to mixin list *)  | 
| 
37105
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
295  | 
(((morphism * bool) * serial) list) Inttab.table;  | 
| 
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
296  | 
(* table of mixin lists, per list mixins in reverse order of declaration;  | 
| 36090 | 297  | 
lists indexed by registration serial,  | 
298  | 
entries for empty lists may be omitted *)  | 
|
| 
37105
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
299  | 
val empty = (Idtab.empty, Inttab.empty);  | 
| 29361 | 300  | 
val extend = I;  | 
| 
37133
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
301  | 
fun merge ((reg1, mix1), (reg2, mix2)) : T =  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
302  | 
(Idtab.join (fn id => fn (r1 as (_, s1), r2 as (_, s2)) =>  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
303  | 
if s1 = s2 then raise Idtab.SAME else raise Idtab.DUP id) (reg1, reg2),  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
304  | 
Inttab.join (K (Library.merge (eq_snd op =))) (mix1, mix2))  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
305  | 
handle Idtab.DUP id =>  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
306  | 
(* distinct interpretations with same base: merge their mixins *)  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
307  | 
let  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
308  | 
val (_, s1) = Idtab.lookup reg1 id |> the;  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
309  | 
val (morph2, s2) = Idtab.lookup reg2 id |> the;  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
310  | 
val reg2' = Idtab.update (id, (morph2, s1)) reg2;  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
311  | 
val mix2' = case Inttab.lookup mix2 s2 of  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
312  | 
NONE => mix2 |  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
313  | 
SOME mxs => Inttab.delete s2 mix2 |> Inttab.update_new (s1, mxs);  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
314  | 
val _ = warning "Removed duplicate interpretation after retrieving its mixins.";  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
315  | 
(* FIXME print interpretations,  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
316  | 
which is not straightforward without theory context *)  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
317  | 
in merge ((reg1, mix1), (reg2', mix2')) end;  | 
| 29361 | 318  | 
(* FIXME consolidate with dependencies, consider one data slot only *)  | 
319  | 
);  | 
|
320  | 
||
| 32801 | 321  | 
|
322  | 
(* Primitive operations *)  | 
|
323  | 
||
| 
37104
 
3877a6c45d57
Avoid recomputation of registration instance for lookup.
 
ballarin 
parents: 
37103 
diff
changeset
 | 
324  | 
fun add_reg thy export (name, morph) =  | 
| 
37105
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
325  | 
Registrations.map (apfst (Idtab.insert (K false)  | 
| 
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
326  | 
((name, instance_of thy name (morph $> export)), ((morph, export), serial ()))));  | 
| 31988 | 327  | 
|
| 
36095
 
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
 
ballarin 
parents: 
36094 
diff
changeset
 | 
328  | 
fun add_mixin serial' mixin =  | 
| 
 
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
 
ballarin 
parents: 
36094 
diff
changeset
 | 
329  | 
(* registration to be amended identified by its serial id *)  | 
| 
37105
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
330  | 
Registrations.map (apsnd (Inttab.map_default (serial', []) (cons (mixin, serial ()))));  | 
| 32801 | 331  | 
|
| 38107 | 332  | 
fun get_mixins context (name, morph) =  | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
333  | 
let  | 
| 38107 | 334  | 
val thy = Context.theory_of context;  | 
335  | 
val (regs, mixins) = Registrations.get context;  | 
|
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
336  | 
in  | 
| 
37105
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
337  | 
case Idtab.lookup regs (name, instance_of thy name morph) of  | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
338  | 
NONE => []  | 
| 
37105
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
339  | 
| SOME (_, serial) => the_default [] (Inttab.lookup mixins serial)  | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
340  | 
end;  | 
| 
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
341  | 
|
| 
36091
 
055934ed89b0
A rough implementation of full mixin inheritance; additional unit tests.
 
ballarin 
parents: 
36090 
diff
changeset
 | 
342  | 
fun compose_mixins mixins =  | 
| 
 
055934ed89b0
A rough implementation of full mixin inheritance; additional unit tests.
 
ballarin 
parents: 
36090 
diff
changeset
 | 
343  | 
fold_rev Morphism.compose (map (fst o fst) mixins) Morphism.identity;  | 
| 
 
055934ed89b0
A rough implementation of full mixin inheritance; additional unit tests.
 
ballarin 
parents: 
36090 
diff
changeset
 | 
344  | 
|
| 38107 | 345  | 
fun collect_mixins context (name, morph) =  | 
346  | 
let  | 
|
347  | 
val thy = Context.theory_of context;  | 
|
348  | 
in  | 
|
349  | 
roundup thy (fn dep => fn mixins => merge (eq_snd op =) (mixins, get_mixins context dep))  | 
|
350  | 
Morphism.identity (name, morph) ([(name, instance_of thy name morph)], [])  | 
|
351  | 
|> snd |> filter (snd o fst) (* only inheritable mixins *)  | 
|
352  | 
|> (fn x => merge (eq_snd op =) (x, get_mixins context (name, morph)))  | 
|
353  | 
|> compose_mixins  | 
|
354  | 
end;  | 
|
| 
36091
 
055934ed89b0
A rough implementation of full mixin inheritance; additional unit tests.
 
ballarin 
parents: 
36090 
diff
changeset
 | 
355  | 
|
| 38107 | 356  | 
fun get_registrations context select = Registrations.get context  | 
| 37491 | 357  | 
|>> Idtab.dest |>> select  | 
| 
36095
 
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
 
ballarin 
parents: 
36094 
diff
changeset
 | 
358  | 
(* with inherited mixins *)  | 
| 
37105
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
359  | 
|-> (fn regs => fn _ => map (fn ((name, _), ((base, export) ,_)) =>  | 
| 38107 | 360  | 
(name, base $> (collect_mixins context (name, base $> export)) $> export)) regs);  | 
| 
36091
 
055934ed89b0
A rough implementation of full mixin inheritance; additional unit tests.
 
ballarin 
parents: 
36090 
diff
changeset
 | 
361  | 
|
| 38111 | 362  | 
fun registrations_of context name =  | 
| 38107 | 363  | 
get_registrations context (filter (curry (op =) name o fst o fst));  | 
| 37491 | 364  | 
|
| 38107 | 365  | 
fun all_registrations context = get_registrations context I;  | 
| 37973 | 366  | 
|
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
367  | 
|
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
368  | 
(*** Activate context elements of locale ***)  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
369  | 
|
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
370  | 
(* Declarations, facts and entire locale content *)  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
371  | 
|
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
372  | 
fun activate_syntax_decls (name, morph) context =  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
373  | 
let  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
374  | 
val thy = Context.theory_of context;  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
375  | 
    val {syntax_decls, ...} = the_locale thy name;
 | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
376  | 
in  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
377  | 
context  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
378  | 
|> fold_rev (fn (decl, _) => decl morph) syntax_decls  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
379  | 
end;  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
380  | 
|
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
381  | 
fun activate_notes activ_elem transfer context export' (name, morph) input =  | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
382  | 
let  | 
| 38107 | 383  | 
val thy = Context.theory_of context;  | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
384  | 
    val {notes, ...} = the_locale thy name;
 | 
| 
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
385  | 
fun activate ((kind, facts), _) input =  | 
| 
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
386  | 
let  | 
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
387  | 
val mixin = case export' of NONE => Morphism.identity  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
388  | 
| SOME export => collect_mixins context (name, morph $> export) $> export;  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
389  | 
val facts' = facts |> Element.facts_map (Element.morph_ctxt (transfer input $> morph $> mixin))  | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
390  | 
in activ_elem (Notes (kind, facts')) input end;  | 
| 
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
391  | 
in  | 
| 
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
392  | 
fold_rev activate notes input  | 
| 
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
393  | 
end;  | 
| 
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
394  | 
|
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
395  | 
fun activate_all name thy activ_elem transfer (marked, input) =  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
396  | 
let  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
397  | 
    val {parameters = (_, params), spec = (asm, defs), ...} = the_locale thy name;
 | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
398  | 
val input' = input |>  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
399  | 
(not (null params) ?  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
400  | 
activ_elem (Fixes (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) params))) |>  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
401  | 
(* FIXME type parameters *)  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
402  | 
(case asm of SOME A => activ_elem (Assumes [(Attrib.empty_binding, [(A, [])])]) | _ => I) |>  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
403  | 
(if not (null defs)  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
404  | 
then activ_elem (Defines (map (fn def => (Attrib.empty_binding, (def, []))) defs))  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
405  | 
else I);  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
406  | 
val activate = activate_notes activ_elem transfer (Context.Theory thy) NONE;  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
407  | 
in  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
408  | 
roundup thy activate Morphism.identity (name, Morphism.identity) (marked, input')  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
409  | 
end;  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
410  | 
|
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
411  | 
|
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
412  | 
(** Public activation functions **)  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
413  | 
|
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
414  | 
fun activate_declarations dep = Context.proof_map (fn context =>  | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
415  | 
let  | 
| 
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
416  | 
val thy = Context.theory_of context;  | 
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
417  | 
in  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
418  | 
roundup thy activate_syntax_decls Morphism.identity dep (get_idents context, context)  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
419  | 
|-> put_idents  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
420  | 
end);  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
421  | 
|
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
422  | 
fun activate_facts export dep context =  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
423  | 
let  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
424  | 
val thy = Context.theory_of context;  | 
| 
38316
 
88e774d09fbc
Revert performance improvement of 8ed3a5fb4d25 since it breaks notes element declarations.
 
ballarin 
parents: 
38211 
diff
changeset
 | 
425  | 
val activate = activate_notes Element.init (Element.transfer_morphism o Context.theory_of) context export;  | 
| 
 
88e774d09fbc
Revert performance improvement of 8ed3a5fb4d25 since it breaks notes element declarations.
 
ballarin 
parents: 
38211 
diff
changeset
 | 
426  | 
in  | 
| 
 
88e774d09fbc
Revert performance improvement of 8ed3a5fb4d25 since it breaks notes element declarations.
 
ballarin 
parents: 
38211 
diff
changeset
 | 
427  | 
roundup thy activate (case export of NONE => Morphism.identity | SOME export => export)  | 
| 
 
88e774d09fbc
Revert performance improvement of 8ed3a5fb4d25 since it breaks notes element declarations.
 
ballarin 
parents: 
38211 
diff
changeset
 | 
428  | 
dep (get_idents context, context)  | 
| 
 
88e774d09fbc
Revert performance improvement of 8ed3a5fb4d25 since it breaks notes element declarations.
 
ballarin 
parents: 
38211 
diff
changeset
 | 
429  | 
|-> put_idents  | 
| 
 
88e774d09fbc
Revert performance improvement of 8ed3a5fb4d25 since it breaks notes element declarations.
 
ballarin 
parents: 
38211 
diff
changeset
 | 
430  | 
end;  | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
431  | 
|
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
432  | 
fun init name thy =  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
433  | 
activate_all name thy Element.init (Element.transfer_morphism o Context.theory_of)  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
434  | 
([], Context.Proof (ProofContext.init_global thy)) |-> put_idents |> Context.proof_of;  | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
435  | 
|
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
436  | 
|
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
437  | 
(*** Add and extend registrations ***)  | 
| 32801 | 438  | 
|
| 38107 | 439  | 
fun amend_registration (name, morph) mixin export context =  | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
440  | 
let  | 
| 38107 | 441  | 
val thy = Context.theory_of context;  | 
442  | 
val regs = Registrations.get context |> fst;  | 
|
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
443  | 
val base = instance_of thy name (morph $> export);  | 
| 
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
444  | 
in  | 
| 
37105
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
445  | 
case Idtab.lookup regs (name, base) of  | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
446  | 
        NONE => error ("No interpretation of locale " ^
 | 
| 
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
447  | 
quote (extern thy name) ^ " and\nparameter instantiation " ^  | 
| 
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
448  | 
space_implode " " (map (quote o Syntax.string_of_term_global thy) base) ^  | 
| 
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
449  | 
" available")  | 
| 38107 | 450  | 
| SOME (_, serial') => add_mixin serial' mixin context  | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
451  | 
end;  | 
| 
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
452  | 
|
| 32801 | 453  | 
(* Note that a registration that would be subsumed by an existing one will not be  | 
454  | 
generated, and it will not be possible to amend it. *)  | 
|
455  | 
||
| 38107 | 456  | 
fun add_registration (name, base_morph) mixin export context =  | 
| 32801 | 457  | 
let  | 
| 38107 | 458  | 
val thy = Context.theory_of context;  | 
| 
37101
 
7099a9ed3be2
Reapply mixin patch: base for performance improvements.
 
ballarin 
parents: 
36919 
diff
changeset
 | 
459  | 
val mix = case mixin of NONE => Morphism.identity  | 
| 
 
7099a9ed3be2
Reapply mixin patch: base for performance improvements.
 
ballarin 
parents: 
36919 
diff
changeset
 | 
460  | 
| SOME (mix, _) => mix;  | 
| 37102 | 461  | 
val morph = base_morph $> mix;  | 
462  | 
val inst = instance_of thy name morph;  | 
|
| 32801 | 463  | 
in  | 
| 38107 | 464  | 
if member (ident_le thy) (get_idents context) (name, inst)  | 
465  | 
then context  | 
|
| 32801 | 466  | 
else  | 
| 38107 | 467  | 
(get_idents context, context)  | 
| 
36095
 
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
 
ballarin 
parents: 
36094 
diff
changeset
 | 
468  | 
(* add new registrations with inherited mixins *)  | 
| 
37104
 
3877a6c45d57
Avoid recomputation of registration instance for lookup.
 
ballarin 
parents: 
37103 
diff
changeset
 | 
469  | 
|> roundup thy (add_reg thy export) export (name, morph)  | 
| 
36095
 
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
 
ballarin 
parents: 
36094 
diff
changeset
 | 
470  | 
|> snd  | 
| 
 
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
 
ballarin 
parents: 
36094 
diff
changeset
 | 
471  | 
(* add mixin *)  | 
| 
 
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
 
ballarin 
parents: 
36094 
diff
changeset
 | 
472  | 
|> (case mixin of NONE => I  | 
| 37102 | 473  | 
| SOME mixin => amend_registration (name, morph) mixin export)  | 
| 
36095
 
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
 
ballarin 
parents: 
36094 
diff
changeset
 | 
474  | 
(* activate import hierarchy as far as not already active *)  | 
| 
38316
 
88e774d09fbc
Revert performance improvement of 8ed3a5fb4d25 since it breaks notes element declarations.
 
ballarin 
parents: 
38211 
diff
changeset
 | 
475  | 
|> activate_facts (SOME export) (name, morph)  | 
| 32801 | 476  | 
end;  | 
477  | 
||
| 29361 | 478  | 
|
| 31988 | 479  | 
(*** Dependencies ***)  | 
480  | 
||
| 37102 | 481  | 
fun add_dependency loc (name, morph) export thy =  | 
482  | 
let  | 
|
483  | 
val thy' = (change_locale loc o apsnd) (cons ((name, (morph, export)), serial ())) thy;  | 
|
| 38107 | 484  | 
val context' = Context.Theory thy';  | 
| 37102 | 485  | 
val (_, regs) = fold_rev (roundup thy' cons export)  | 
| 
38783
 
f171050ad3f8
For sublocale it is sufficient to reconsider ancestors of the target.
 
ballarin 
parents: 
38316 
diff
changeset
 | 
486  | 
(registrations_of context' loc) (get_idents (context'), []);  | 
| 37102 | 487  | 
in  | 
488  | 
thy'  | 
|
| 38107 | 489  | 
|> fold_rev (fn dep => Context.theory_map (add_registration dep NONE export)) regs  | 
| 37102 | 490  | 
end;  | 
| 31988 | 491  | 
|
492  | 
||
| 29361 | 493  | 
(*** Storing results ***)  | 
494  | 
||
495  | 
(* Theorems *)  | 
|
496  | 
||
497  | 
fun add_thmss loc kind args ctxt =  | 
|
498  | 
let  | 
|
| 
30777
 
9960ff945c52
simplified Element.activate(_i): singleton version;
 
wenzelm 
parents: 
30775 
diff
changeset
 | 
499  | 
val (Notes args', ctxt') = Element.activate_i (Notes (kind, args)) ctxt;  | 
| 
38757
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38756 
diff
changeset
 | 
500  | 
val ctxt'' = ctxt' |> ProofContext.background_theory  | 
| 
 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38756 
diff
changeset
 | 
501  | 
((change_locale loc o apfst o apsnd) (cons (args', serial ()))  | 
| 29392 | 502  | 
#>  | 
| 29361 | 503  | 
(* Registrations *)  | 
| 31988 | 504  | 
(fn thy => fold_rev (fn (_, morph) =>  | 
| 29361 | 505  | 
let  | 
506  | 
val args'' = snd args' |> Element.facts_map (Element.morph_ctxt morph) |>  | 
|
507  | 
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
 | 
508  | 
in PureThy.note_thmss kind args'' #> snd end)  | 
| 38111 | 509  | 
(registrations_of (Context.Theory thy) loc) thy))  | 
| 29361 | 510  | 
in ctxt'' end;  | 
511  | 
||
512  | 
||
513  | 
(* Declarations *)  | 
|
514  | 
||
| 
35798
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
515  | 
fun add_declaration loc decl =  | 
| 
33643
 
b275f26a638b
eliminated obsolete "internal" kind -- collapsed to unspecific "";
 
wenzelm 
parents: 
33541 
diff
changeset
 | 
516  | 
add_thmss loc ""  | 
| 
35798
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
517  | 
[((Binding.conceal Binding.empty,  | 
| 
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
518  | 
[Attrib.internal (fn phi => Thm.declaration_attribute (K (decl phi)))]),  | 
| 33278 | 519  | 
[([Drule.dummy_thm], [])])];  | 
| 29361 | 520  | 
|
| 
35798
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
521  | 
fun add_syntax_declaration loc decl =  | 
| 
38756
 
d07959fabde6
renamed ProofContext.theory(_result) to ProofContext.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38316 
diff
changeset
 | 
522  | 
ProofContext.background_theory ((change_locale loc o apfst o apfst) (cons (decl, serial ())))  | 
| 
35798
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
523  | 
#> add_declaration loc decl;  | 
| 29361 | 524  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
525  | 
|
| 29361 | 526  | 
(*** Reasoning about locales ***)  | 
527  | 
||
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
528  | 
(* Storage for witnesses, intro and unfold rules *)  | 
| 29361 | 529  | 
|
| 33519 | 530  | 
structure Thms = Generic_Data  | 
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
531  | 
(  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
532  | 
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
 | 
533  | 
val empty = ([], [], []);  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
534  | 
val extend = I;  | 
| 33519 | 535  | 
fun merge ((witnesses1, intros1, unfolds1), (witnesses2, intros2, unfolds2)) =  | 
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
536  | 
(Thm.merge_thms (witnesses1, witnesses2),  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
537  | 
Thm.merge_thms (intros1, intros2),  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
538  | 
Thm.merge_thms (unfolds1, unfolds2));  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
539  | 
);  | 
| 29361 | 540  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
541  | 
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
 | 
542  | 
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
 | 
543  | 
val get_unfolds = #3 o Thms.get o Context.Proof;  | 
| 29361 | 544  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
545  | 
val witness_add =  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
546  | 
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
 | 
547  | 
val intro_add =  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
548  | 
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
 | 
549  | 
val unfold_add =  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
550  | 
Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (x, y, Thm.add_thm th z)));  | 
| 29361 | 551  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
552  | 
|
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
553  | 
(* Tactic *)  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
554  | 
|
| 
36093
 
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
 
ballarin 
parents: 
36091 
diff
changeset
 | 
555  | 
fun gen_intro_locales_tac intros_tac eager ctxt =  | 
| 
 
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
 
ballarin 
parents: 
36091 
diff
changeset
 | 
556  | 
intros_tac  | 
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
557  | 
(get_witnesses ctxt @ get_intros ctxt @ (if eager then get_unfolds ctxt else []));  | 
| 29361 | 558  | 
|
| 
36093
 
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
 
ballarin 
parents: 
36091 
diff
changeset
 | 
559  | 
val intro_locales_tac = gen_intro_locales_tac Method.intros_tac;  | 
| 
 
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
 
ballarin 
parents: 
36091 
diff
changeset
 | 
560  | 
val try_intro_locales_tac= gen_intro_locales_tac Method.try_intros_tac;  | 
| 
 
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
 
ballarin 
parents: 
36091 
diff
changeset
 | 
561  | 
|
| 29361 | 562  | 
val _ = Context.>> (Context.map_theory  | 
| 
36093
 
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
 
ballarin 
parents: 
36091 
diff
changeset
 | 
563  | 
(Method.setup (Binding.name "intro_locales") (Scan.succeed (METHOD o try_intro_locales_tac false))  | 
| 30515 | 564  | 
"back-chain introduction rules of locales without unfolding predicates" #>  | 
| 
36093
 
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
 
ballarin 
parents: 
36091 
diff
changeset
 | 
565  | 
Method.setup (Binding.name "unfold_locales") (Scan.succeed (METHOD o try_intro_locales_tac true))  | 
| 30515 | 566  | 
"back-chain all introduction rules of locales"));  | 
| 29361 | 567  | 
|
| 37471 | 568  | 
|
569  | 
(*** diagnostic commands and interfaces ***)  | 
|
570  | 
||
| 
37897
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
571  | 
val all_locales = Symtab.keys o snd o Locales.get;  | 
| 37471 | 572  | 
|
573  | 
fun print_locales thy =  | 
|
| 
37897
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
574  | 
  Pretty.strs ("locales:" :: map #1 (Name_Space.extern_table (Locales.get thy)))
 | 
| 37471 | 575  | 
|> Pretty.writeln;  | 
576  | 
||
577  | 
fun print_locale thy show_facts raw_name =  | 
|
578  | 
let  | 
|
579  | 
val name = intern thy raw_name;  | 
|
580  | 
val ctxt = init name thy;  | 
|
581  | 
fun cons_elem (elem as Notes _) = show_facts ? cons elem  | 
|
582  | 
| cons_elem elem = cons elem;  | 
|
583  | 
val elems =  | 
|
584  | 
activate_all name thy cons_elem (K (Element.transfer_morphism thy)) ([], [])  | 
|
585  | 
|> snd |> rev;  | 
|
586  | 
in  | 
|
587  | 
Pretty.big_list "locale elements:" (map (Pretty.chunks o Element.pretty_ctxt ctxt) elems)  | 
|
588  | 
|> Pretty.writeln  | 
|
589  | 
end;  | 
|
590  | 
||
| 38109 | 591  | 
fun print_registrations ctxt raw_name =  | 
592  | 
let  | 
|
593  | 
val thy = ProofContext.theory_of ctxt;  | 
|
594  | 
in  | 
|
| 38111 | 595  | 
(case registrations_of (Context.Proof ctxt) (* FIXME *) (intern thy raw_name) of  | 
| 38109 | 596  | 
        [] => Pretty.str ("no interpretations")
 | 
597  | 
| regs => Pretty.big_list "interpretations:" (map (pretty_reg thy) (rev regs)))  | 
|
598  | 
|> Pretty.writeln  | 
|
599  | 
end;  | 
|
| 37471 | 600  | 
|
| 
37897
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
601  | 
fun locale_deps thy =  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
602  | 
let  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
603  | 
val names = all_locales thy  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
604  | 
fun add_locale_node name =  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
605  | 
let  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
606  | 
val params = params_of thy name;  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
607  | 
val axioms = (these o Option.map (Logic.strip_imp_prems o Thm.prop_of) o fst o intros_of thy) name;  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
608  | 
val registrations = map (instance_of thy name o snd)  | 
| 38111 | 609  | 
(registrations_of (Context.Theory thy) name);  | 
| 
37897
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
610  | 
in  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
611  | 
        Graph.new_node (name, { params = params, axioms = axioms, registrations = registrations })
 | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
612  | 
end;  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
613  | 
fun add_locale_deps name =  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
614  | 
let  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
615  | 
val dependencies = (map o apsnd) (instance_of thy name o op $>)  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
616  | 
(dependencies_of thy name);  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
617  | 
in  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
618  | 
fold (fn (super, ts) => fn (gr, deps) => (gr |> Graph.add_edge (super, name),  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
619  | 
deps |> Symtab.map_default (super, Symtab.empty) (Symtab.cons_list (name, ts))))  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
620  | 
dependencies  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
621  | 
end;  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
622  | 
in  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
623  | 
Graph.empty  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
624  | 
|> fold add_locale_node names  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
625  | 
|> rpair Symtab.empty  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
626  | 
|> fold add_locale_deps names  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
627  | 
end;  | 
| 
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
628  | 
|
| 29361 | 629  | 
end;  |