| author | wenzelm | 
| Wed, 17 Jun 2009 17:06:07 +0200 | |
| changeset 31686 | e54ae15335a1 | 
| parent 30777 | 9960ff945c52 | 
| child 31988 | 801aabf9f376 | 
| 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 ->  | 
| 
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 | 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  | 
val declarations_of: theory -> string -> declaration list * declaration list  | 
|
| 29544 | 48  | 
val dependencies_of: theory -> string -> (string * morphism) list  | 
| 29361 | 49  | 
|
50  | 
(* Storing results *)  | 
|
51  | 
val add_thmss: string -> string -> (Attrib.binding * (thm list * Attrib.src list) list) list ->  | 
|
52  | 
Proof.context -> Proof.context  | 
|
53  | 
val add_type_syntax: string -> declaration -> Proof.context -> Proof.context  | 
|
54  | 
val add_term_syntax: string -> declaration -> Proof.context -> Proof.context  | 
|
55  | 
val add_declaration: string -> declaration -> Proof.context -> Proof.context  | 
|
| 29544 | 56  | 
val add_dependency: string -> string * morphism -> theory -> theory  | 
| 29361 | 57  | 
|
58  | 
(* Activation *)  | 
|
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
59  | 
val activate_declarations: string * morphism -> Proof.context -> Proof.context  | 
| 
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
60  | 
val activate_facts: string * morphism -> Context.generic -> Context.generic  | 
| 29361 | 61  | 
val init: string -> theory -> Proof.context  | 
62  | 
||
63  | 
(* Reasoning about locales *)  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
64  | 
val get_witnesses: Proof.context -> thm list  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
65  | 
val get_intros: Proof.context -> thm list  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
66  | 
val get_unfolds: Proof.context -> thm list  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
67  | 
val witness_add: attribute  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
68  | 
val intro_add: attribute  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
69  | 
val unfold_add: attribute  | 
| 29361 | 70  | 
val intro_locales_tac: bool -> Proof.context -> thm list -> tactic  | 
71  | 
||
72  | 
(* Registrations *)  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
73  | 
val add_registration: string * (morphism * morphism) -> theory -> theory  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
74  | 
val amend_registration: morphism -> string * morphism -> theory -> theory  | 
| 29544 | 75  | 
val get_registrations: theory -> (string * morphism) list  | 
| 29361 | 76  | 
|
77  | 
(* Diagnostic *)  | 
|
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  | 
| 29361 | 80  | 
end;  | 
81  | 
||
82  | 
structure Locale: LOCALE =  | 
|
83  | 
struct  | 
|
84  | 
||
85  | 
datatype ctxt = datatype Element.ctxt;  | 
|
86  | 
||
| 29392 | 87  | 
|
88  | 
(*** Theory data ***)  | 
|
| 29361 | 89  | 
|
90  | 
datatype locale = Loc of {
 | 
|
| 29392 | 91  | 
(** static part **)  | 
| 
30755
 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 
wenzelm 
parents: 
30754 
diff
changeset
 | 
92  | 
parameters: (string * sort) list * ((string * typ) * mixfix) list,  | 
| 29361 | 93  | 
(* type and term parameters *)  | 
94  | 
spec: term option * term list,  | 
|
95  | 
(* assumptions (as a single predicate expression) and defines *)  | 
|
| 29441 | 96  | 
intros: thm option * thm option,  | 
97  | 
axioms: thm list,  | 
|
| 29392 | 98  | 
(** dynamic part **)  | 
| 29361 | 99  | 
decls: (declaration * stamp) list * (declaration * stamp) list,  | 
100  | 
(* type and term syntax declarations *)  | 
|
101  | 
notes: ((string * (Attrib.binding * (thm list * Attrib.src list) list) list) * stamp) list,  | 
|
102  | 
(* theorem declarations *)  | 
|
| 29544 | 103  | 
dependencies: ((string * morphism) * stamp) list  | 
| 29361 | 104  | 
(* locale dependencies (sublocale relation) *)  | 
| 29392 | 105  | 
};  | 
| 29361 | 106  | 
|
| 29441 | 107  | 
fun mk_locale ((parameters, spec, intros, axioms), ((decls, notes), dependencies)) =  | 
108  | 
  Loc {parameters = parameters, intros = intros, axioms = axioms, spec = spec,
 | 
|
109  | 
decls = decls, notes = notes, dependencies = dependencies};  | 
|
| 30754 | 110  | 
|
| 29441 | 111  | 
fun map_locale f (Loc {parameters, spec, intros, axioms, decls, notes, dependencies}) =
 | 
112  | 
mk_locale (f ((parameters, spec, intros, axioms), ((decls, notes), dependencies)));  | 
|
| 30754 | 113  | 
|
| 29441 | 114  | 
fun merge_locale (Loc {parameters, spec, intros, axioms, decls = (decls1, decls2),
 | 
115  | 
  notes, dependencies}, Loc {decls = (decls1', decls2'), notes = notes',
 | 
|
116  | 
dependencies = dependencies', ...}) = mk_locale  | 
|
117  | 
((parameters, spec, intros, axioms),  | 
|
118  | 
(((merge (eq_snd op =) (decls1, decls1'), merge (eq_snd op =) (decls2, decls2')),  | 
|
119  | 
merge (eq_snd op =) (notes, notes')),  | 
|
120  | 
merge (eq_snd op =) (dependencies, dependencies')));  | 
|
| 29361 | 121  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
122  | 
structure Locales = TheoryDataFun  | 
| 29361 | 123  | 
(  | 
| 29392 | 124  | 
type T = locale NameSpace.table;  | 
| 29361 | 125  | 
val empty = NameSpace.empty_table;  | 
126  | 
val copy = I;  | 
|
127  | 
val extend = I;  | 
|
| 29392 | 128  | 
fun merge _ = NameSpace.join_tables (K merge_locale);  | 
| 29361 | 129  | 
);  | 
130  | 
||
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
131  | 
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
 | 
132  | 
val extern = NameSpace.extern o #1 o Locales.get;  | 
| 29392 | 133  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
134  | 
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
 | 
135  | 
val defined = Symtab.defined o #2 o Locales.get;  | 
| 29361 | 136  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
137  | 
fun the_locale thy name =  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
138  | 
(case get_locale thy name of  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
139  | 
SOME (Loc loc) => loc  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
140  | 
  | NONE => error ("Unknown locale " ^ quote name));
 | 
| 29361 | 141  | 
|
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30223 
diff
changeset
 | 
142  | 
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
 | 
143  | 
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
 | 
144  | 
(binding,  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
145  | 
mk_locale ((parameters, spec, intros, axioms),  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
146  | 
((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
 | 
147  | 
map (fn d => (d, stamp ())) dependencies))) #> snd);  | 
| 29361 | 148  | 
|
| 29392 | 149  | 
fun change_locale name =  | 
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
150  | 
Locales.map o apsnd o Symtab.map_entry name o map_locale o apsnd;  | 
| 29361 | 151  | 
|
152  | 
fun print_locales thy =  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
153  | 
  Pretty.strs ("locales:" :: map #1 (NameSpace.extern_table (Locales.get thy)))
 | 
| 29392 | 154  | 
|> Pretty.writeln;  | 
| 29361 | 155  | 
|
156  | 
||
157  | 
(*** Primitive operations ***)  | 
|
158  | 
||
| 29392 | 159  | 
fun params_of thy = snd o #parameters o the_locale thy;  | 
| 29361 | 160  | 
|
| 29441 | 161  | 
fun intros_of thy = #intros o the_locale thy;  | 
162  | 
||
163  | 
fun axioms_of thy = #axioms o the_locale thy;  | 
|
164  | 
||
| 29392 | 165  | 
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
 | 
166  | 
map (Morphism.term morph o Free o #1);  | 
| 29361 | 167  | 
|
| 29392 | 168  | 
fun specification_of thy = #spec o the_locale thy;  | 
| 29361 | 169  | 
|
| 29392 | 170  | 
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
 | 
171  | 
#decls |> pairself (map fst);  | 
| 29361 | 172  | 
|
| 29544 | 173  | 
fun dependencies_of thy name = the_locale thy name |>  | 
174  | 
#dependencies |> map fst;  | 
|
175  | 
||
| 29361 | 176  | 
|
177  | 
(*** Activate context elements of locale ***)  | 
|
178  | 
||
179  | 
(** Identifiers: activated locales in theory or proof context **)  | 
|
180  | 
||
181  | 
fun ident_eq thy ((n: string, ts), (m, ss)) = (m = n) andalso Pattern.matchess thy (ss, ts);  | 
|
182  | 
||
183  | 
local  | 
|
184  | 
||
| 30754 | 185  | 
datatype 'a delayed = Ready of 'a | ToDo of 'a delayed * 'a delayed;  | 
| 29361 | 186  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
187  | 
structure Identifiers = GenericDataFun  | 
| 29361 | 188  | 
(  | 
| 30754 | 189  | 
type T = (string * term list) list delayed;  | 
190  | 
val empty = Ready [];  | 
|
| 29361 | 191  | 
val extend = I;  | 
192  | 
fun merge _ = ToDo;  | 
|
193  | 
);  | 
|
194  | 
||
195  | 
fun finish thy (ToDo (i1, i2)) = merge (ident_eq thy) (finish thy i1, finish thy i2)  | 
|
196  | 
| finish _ (Ready ids) = ids;  | 
|
197  | 
||
198  | 
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
 | 
199  | 
(case Identifiers.get (Context.Theory thy) of  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
200  | 
Ready _ => NONE  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
201  | 
| ids => SOME (Context.theory_map (Identifiers.put (Ready (finish thy ids))) thy)))));  | 
| 29361 | 202  | 
|
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
203  | 
in  | 
| 29361 | 204  | 
|
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
205  | 
val get_idents = (fn Ready ids => ids) o Identifiers.get;  | 
| 
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
206  | 
val put_idents = Identifiers.put o Ready;  | 
| 29361 | 207  | 
|
208  | 
end;  | 
|
209  | 
||
210  | 
||
211  | 
(** Resolve locale dependencies in a depth-first fashion **)  | 
|
212  | 
||
213  | 
local  | 
|
214  | 
||
215  | 
val roundup_bound = 120;  | 
|
216  | 
||
217  | 
fun add thy depth (name, morph) (deps, marked) =  | 
|
218  | 
if depth > roundup_bound  | 
|
219  | 
then error "Roundup bound exceeded (sublocale relation probably not terminating)."  | 
|
220  | 
else  | 
|
221  | 
let  | 
|
| 30754 | 222  | 
val dependencies = dependencies_of thy name;  | 
| 29361 | 223  | 
val instance = instance_of thy name morph;  | 
224  | 
in  | 
|
225  | 
if member (ident_eq thy) marked (name, instance)  | 
|
226  | 
then (deps, marked)  | 
|
227  | 
else  | 
|
228  | 
let  | 
|
| 30754 | 229  | 
val dependencies' = map (fn (name, morph') => (name, morph' $> morph)) dependencies;  | 
| 29361 | 230  | 
val marked' = (name, instance) :: marked;  | 
231  | 
val (deps', marked'') = fold_rev (add thy (depth + 1)) dependencies' ([], marked');  | 
|
232  | 
in  | 
|
233  | 
((name, morph) :: deps' @ deps, marked'')  | 
|
234  | 
end  | 
|
235  | 
end;  | 
|
236  | 
||
237  | 
in  | 
|
238  | 
||
239  | 
fun roundup thy activate_dep (name, morph) (marked, input) =  | 
|
240  | 
let  | 
|
241  | 
(* Find all dependencies incuding new ones (which are dependencies enriching  | 
|
242  | 
existing registrations). *)  | 
|
| 30754 | 243  | 
val (dependencies, marked') = add thy 0 (name, morph) ([], []);  | 
| 29361 | 244  | 
(* Filter out exisiting fragments. *)  | 
245  | 
val dependencies' = filter_out (fn (name, morph) =>  | 
|
246  | 
member (ident_eq thy) marked (name, instance_of thy name morph)) dependencies;  | 
|
247  | 
in  | 
|
| 30773 | 248  | 
(merge (ident_eq thy) (marked, marked'), input |> fold_rev activate_dep dependencies')  | 
| 29361 | 249  | 
end;  | 
250  | 
||
251  | 
end;  | 
|
252  | 
||
253  | 
||
254  | 
(* Declarations, facts and entire locale content *)  | 
|
255  | 
||
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
256  | 
fun activate_decls (name, morph) context =  | 
| 29361 | 257  | 
let  | 
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
258  | 
val thy = Context.theory_of context;  | 
| 29392 | 259  | 
    val {decls = (typ_decls, term_decls), ...} = the_locale thy name;
 | 
| 29361 | 260  | 
in  | 
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
261  | 
context  | 
| 
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
262  | 
|> fold_rev (fn (decl, _) => decl morph) typ_decls  | 
| 
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
263  | 
|> fold_rev (fn (decl, _) => decl morph) term_decls  | 
| 29361 | 264  | 
end;  | 
265  | 
||
266  | 
fun activate_notes activ_elem transfer thy (name, morph) input =  | 
|
267  | 
let  | 
|
| 29392 | 268  | 
    val {notes, ...} = the_locale thy name;
 | 
| 29361 | 269  | 
fun activate ((kind, facts), _) input =  | 
270  | 
let  | 
|
271  | 
val facts' = facts |> Element.facts_map (Element.morph_ctxt (transfer input $> morph))  | 
|
272  | 
in activ_elem (Notes (kind, facts')) input end;  | 
|
273  | 
in  | 
|
274  | 
fold_rev activate notes input  | 
|
275  | 
end;  | 
|
276  | 
||
277  | 
fun activate_all name thy activ_elem transfer (marked, input) =  | 
|
278  | 
let  | 
|
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
279  | 
    val {parameters = (_, params), spec = (asm, defs), ...} = the_locale thy name;
 | 
| 
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
280  | 
val input' = input |>  | 
| 
30755
 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 
wenzelm 
parents: 
30754 
diff
changeset
 | 
281  | 
(not (null params) ?  | 
| 
 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 
wenzelm 
parents: 
30754 
diff
changeset
 | 
282  | 
activ_elem (Fixes (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) params))) |>  | 
| 29361 | 283  | 
(* FIXME type parameters *)  | 
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
284  | 
(case asm of SOME A => activ_elem (Assumes [(Attrib.empty_binding, [(A, [])])]) | _ => I) |>  | 
| 29361 | 285  | 
(if not (null defs)  | 
286  | 
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
 | 
287  | 
else I);  | 
| 30773 | 288  | 
val activate = activate_notes activ_elem transfer thy;  | 
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
289  | 
in  | 
| 30773 | 290  | 
roundup thy activate (name, Morphism.identity) (marked, input')  | 
| 29361 | 291  | 
end;  | 
292  | 
||
293  | 
||
294  | 
(** Public activation functions **)  | 
|
295  | 
||
| 
30775
 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 
wenzelm 
parents: 
30773 
diff
changeset
 | 
296  | 
fun activate_declarations dep = Context.proof_map (fn context =>  | 
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
297  | 
let  | 
| 
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
298  | 
val thy = Context.theory_of context;  | 
| 30773 | 299  | 
val context' = roundup thy activate_decls dep (get_idents context, context) |-> put_idents;  | 
| 
30775
 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 
wenzelm 
parents: 
30773 
diff
changeset
 | 
300  | 
in context' end);  | 
| 29361 | 301  | 
|
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
302  | 
fun activate_facts dep context =  | 
| 
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
303  | 
let  | 
| 
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
304  | 
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
 | 
305  | 
val activate = activate_notes Element.init (Element.transfer_morphism o Context.theory_of) thy;  | 
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
306  | 
in roundup thy activate dep (get_idents context, context) |-> put_idents end;  | 
| 29361 | 307  | 
|
308  | 
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
 | 
309  | 
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
 | 
310  | 
([], Context.Proof (ProofContext.init thy)) |-> put_idents |> Context.proof_of;  | 
| 29361 | 311  | 
|
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
312  | 
fun print_locale thy show_facts raw_name =  | 
| 29361 | 313  | 
let  | 
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
314  | 
val name = intern thy raw_name;  | 
| 
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
315  | 
val ctxt = init name thy;  | 
| 
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
316  | 
fun cons_elem (elem as Notes _) = show_facts ? cons elem  | 
| 
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
317  | 
| cons_elem elem = cons elem;  | 
| 
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
318  | 
val elems =  | 
| 
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
319  | 
activate_all name thy cons_elem (K (Element.transfer_morphism thy)) ([], [])  | 
| 
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
320  | 
|> snd |> rev;  | 
| 29361 | 321  | 
in  | 
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
322  | 
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
 | 
323  | 
|> Pretty.writeln  | 
| 
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
324  | 
end;  | 
| 29361 | 325  | 
|
326  | 
||
327  | 
(*** Registrations: interpretations in theories ***)  | 
|
328  | 
||
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
329  | 
structure Registrations = TheoryDataFun  | 
| 29361 | 330  | 
(  | 
| 29544 | 331  | 
type T = ((string * (morphism * morphism)) * stamp) list;  | 
| 29392 | 332  | 
(* FIXME mixins need to be stamped *)  | 
| 29361 | 333  | 
(* registrations, in reverse order of declaration *)  | 
334  | 
val empty = [];  | 
|
335  | 
val extend = I;  | 
|
| 29392 | 336  | 
val copy = I;  | 
| 29361 | 337  | 
fun merge _ data : T = Library.merge (eq_snd op =) data;  | 
338  | 
(* FIXME consolidate with dependencies, consider one data slot only *)  | 
|
339  | 
);  | 
|
340  | 
||
| 29392 | 341  | 
val get_registrations =  | 
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
342  | 
Registrations.get #> map (#1 #> apsnd op $>);  | 
| 29361 | 343  | 
|
| 29392 | 344  | 
fun add_registration (name, (base_morph, export)) thy =  | 
| 30773 | 345  | 
roundup thy (fn (name', morph') => Registrations.map (cons ((name', (morph', export)), stamp ())))  | 
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
346  | 
(name, base_morph) (get_idents (Context.Theory thy), thy) |> snd;  | 
| 
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
347  | 
(* FIXME |-> put_global_idents ?*)  | 
| 29361 | 348  | 
|
| 29392 | 349  | 
fun amend_registration morph (name, base_morph) thy =  | 
| 29361 | 350  | 
let  | 
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
351  | 
val regs = map #1 (Registrations.get thy);  | 
| 29361 | 352  | 
val base = instance_of thy name base_morph;  | 
353  | 
fun match (name', (morph', _)) =  | 
|
354  | 
name = name' andalso eq_list (op aconv) (base, instance_of thy name' morph');  | 
|
355  | 
val i = find_index match (rev regs);  | 
|
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
356  | 
val _ =  | 
| 
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
357  | 
      if i = ~1 then error ("No registration of locale " ^
 | 
| 29361 | 358  | 
quote (extern thy name) ^ " and parameter instantiation " ^  | 
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
359  | 
space_implode " " (map (quote o Syntax.string_of_term_global thy) base) ^ " available")  | 
| 29361 | 360  | 
else ();  | 
361  | 
in  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
362  | 
Registrations.map (nth_map (length regs - 1 - i)  | 
| 29361 | 363  | 
(fn ((name, (base, export)), stamp) => ((name, (base $> morph, export)), stamp))) thy  | 
364  | 
end;  | 
|
365  | 
||
366  | 
||
367  | 
(*** Storing results ***)  | 
|
368  | 
||
369  | 
(* Theorems *)  | 
|
370  | 
||
371  | 
fun add_thmss loc kind args ctxt =  | 
|
372  | 
let  | 
|
| 
30777
 
9960ff945c52
simplified Element.activate(_i): singleton version;
 
wenzelm 
parents: 
30775 
diff
changeset
 | 
373  | 
val (Notes args', ctxt') = Element.activate_i (Notes (kind, args)) ctxt;  | 
| 29361 | 374  | 
val ctxt'' = ctxt' |> ProofContext.theory (  | 
| 29392 | 375  | 
(change_locale loc o apfst o apsnd) (cons (args', stamp ()))  | 
376  | 
#>  | 
|
| 29361 | 377  | 
(* Registrations *)  | 
378  | 
(fn thy => fold_rev (fn (name, morph) =>  | 
|
379  | 
let  | 
|
380  | 
val args'' = snd args' |> Element.facts_map (Element.morph_ctxt morph) |>  | 
|
381  | 
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
 | 
382  | 
in PureThy.note_thmss kind args'' #> snd end)  | 
| 29392 | 383  | 
(get_registrations thy |> filter (fn (name, _) => name = loc)) thy))  | 
| 29361 | 384  | 
in ctxt'' end;  | 
385  | 
||
386  | 
||
387  | 
(* Declarations *)  | 
|
388  | 
||
389  | 
local  | 
|
390  | 
||
391  | 
fun decl_attrib decl phi = Thm.declaration_attribute (K (decl phi));  | 
|
392  | 
||
393  | 
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
 | 
394  | 
ProofContext.theory ((change_locale loc o apfst o apfst) (add (decl, stamp ()))) #>  | 
| 29361 | 395  | 
add_thmss loc Thm.internalK  | 
396  | 
[((Binding.empty, [Attrib.internal (decl_attrib decl)]), [([Drule.dummy_thm], [])])];  | 
|
397  | 
||
398  | 
in  | 
|
399  | 
||
400  | 
val add_type_syntax = add_decls (apfst o cons);  | 
|
401  | 
val add_term_syntax = add_decls (apsnd o cons);  | 
|
402  | 
val add_declaration = add_decls (K I);  | 
|
403  | 
||
404  | 
end;  | 
|
405  | 
||
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
406  | 
|
| 29361 | 407  | 
(* Dependencies *)  | 
408  | 
||
| 29392 | 409  | 
fun add_dependency loc dep = (change_locale loc o apsnd) (cons (dep, stamp ()));  | 
| 29361 | 410  | 
|
411  | 
||
412  | 
(*** Reasoning about locales ***)  | 
|
413  | 
||
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
414  | 
(* Storage for witnesses, intro and unfold rules *)  | 
| 29361 | 415  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
416  | 
structure Thms = GenericDataFun  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
417  | 
(  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
418  | 
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
 | 
419  | 
val empty = ([], [], []);  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
420  | 
val extend = I;  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
421  | 
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
 | 
422  | 
(Thm.merge_thms (witnesses1, witnesses2),  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
423  | 
Thm.merge_thms (intros1, intros2),  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
424  | 
Thm.merge_thms (unfolds1, unfolds2));  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
425  | 
);  | 
| 29361 | 426  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
427  | 
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
 | 
428  | 
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
 | 
429  | 
val get_unfolds = #3 o Thms.get o Context.Proof;  | 
| 29361 | 430  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
431  | 
val witness_add =  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
432  | 
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
 | 
433  | 
val intro_add =  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
434  | 
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
 | 
435  | 
val unfold_add =  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
436  | 
Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (x, y, Thm.add_thm th z)));  | 
| 29361 | 437  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
438  | 
|
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
439  | 
(* Tactic *)  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
440  | 
|
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
441  | 
fun intro_locales_tac eager ctxt =  | 
| 29361 | 442  | 
Method.intros_tac  | 
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
443  | 
(get_witnesses ctxt @ get_intros ctxt @ (if eager then get_unfolds ctxt else []));  | 
| 29361 | 444  | 
|
445  | 
val _ = Context.>> (Context.map_theory  | 
|
| 30515 | 446  | 
(Method.setup (Binding.name "intro_locales") (Scan.succeed (METHOD o intro_locales_tac false))  | 
447  | 
"back-chain introduction rules of locales without unfolding predicates" #>  | 
|
448  | 
Method.setup (Binding.name "unfold_locales") (Scan.succeed (METHOD o intro_locales_tac true))  | 
|
449  | 
"back-chain all introduction rules of locales"));  | 
|
| 29361 | 450  | 
|
451  | 
end;  | 
|
452  |