| author | haftmann | 
| Fri, 12 Oct 2007 14:42:30 +0200 | |
| changeset 25002 | c3d9cb390471 | 
| parent 24952 | f336c36f41a0 | 
| child 25038 | 522abf8a5f87 | 
| permissions | -rw-r--r-- | 
| 12014 | 1  | 
(* Title: Pure/Isar/locale.ML  | 
| 11896 | 2  | 
ID: $Id$  | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
3  | 
Author: Clemens Ballarin, TU Muenchen; Markus Wenzel, LMU/TU Muenchen  | 
| 11896 | 4  | 
|
| 12058 | 5  | 
Locales -- Isar proof contexts as meta-level predicates, with local  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
6  | 
syntax and implicit structures.  | 
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
7  | 
|
| 
19942
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
8  | 
Draws basic ideas from Florian Kammueller's original version of  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
9  | 
locales, but uses the richer infrastructure of Isar instead of the raw  | 
| 
20035
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
10  | 
meta-logic. Furthermore, structured import of contexts (with merge  | 
| 
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
11  | 
and rename operations) are provided, as well as type-inference of the  | 
| 13375 | 12  | 
signature parts, and predicate definitions of the specification text.  | 
| 14446 | 13  | 
|
| 17437 | 14  | 
Interpretation enables the reuse of theorems of locales in other  | 
15  | 
contexts, namely those defined by theories, structured proofs and  | 
|
16  | 
locales themselves.  | 
|
17  | 
||
| 14446 | 18  | 
See also:  | 
19  | 
||
20  | 
[1] Clemens Ballarin. Locales and Locale Expressions in Isabelle/Isar.  | 
|
21  | 
In Stefano Berardi et al., Types for Proofs and Programs: International  | 
|
| 15099 | 22  | 
Workshop, TYPES 2003, Torino, Italy, LNCS 3085, pages 34-50, 2004.  | 
| 
19942
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
23  | 
[2] Clemens Ballarin. Interpretation of Locales in Isabelle: Managing  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
24  | 
Dependencies between Locales. Technical Report TUM-I0607, Technische  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
25  | 
Universitaet Muenchen, 2006.  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
26  | 
[3] Clemens Ballarin. Interpretation of Locales in Isabelle: Theories and  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
27  | 
Proof Contexts. In J.M. Borwein and W.M. Farmer, MKM 2006, LNAI 4108,  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
28  | 
pages 31-43, 2006.  | 
| 11896 | 29  | 
*)  | 
30  | 
||
| 
16169
 
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
 
ballarin 
parents: 
16144 
diff
changeset
 | 
31  | 
(* TODO:  | 
| 
 
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
 
ballarin 
parents: 
16144 
diff
changeset
 | 
32  | 
- beta-eta normalisation of interpretation parameters  | 
| 
 
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
 
ballarin 
parents: 
16144 
diff
changeset
 | 
33  | 
- dangling type frees in locales  | 
| 
16620
 
2a7f46324218
Proper treatment of beta-redexes in witness theorems.
 
ballarin 
parents: 
16458 
diff
changeset
 | 
34  | 
- test subsumption of interpretations when merging theories  | 
| 
16169
 
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
 
ballarin 
parents: 
16144 
diff
changeset
 | 
35  | 
*)  | 
| 
 
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
 
ballarin 
parents: 
16144 
diff
changeset
 | 
36  | 
|
| 11896 | 37  | 
signature LOCALE =  | 
38  | 
sig  | 
|
| 12273 | 39  | 
datatype expr =  | 
40  | 
Locale of string |  | 
|
| 
16102
 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 
ballarin 
parents: 
16028 
diff
changeset
 | 
41  | 
Rename of expr * (string * mixfix option) option list |  | 
| 12273 | 42  | 
Merge of expr list  | 
43  | 
val empty: expr  | 
|
| 18137 | 44  | 
datatype 'a element = Elem of 'a | Expr of expr  | 
| 21035 | 45  | 
  val map_elem: ('a -> 'b) -> 'a element -> 'b element
 | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
46  | 
|
| 16458 | 47  | 
val intern: theory -> xstring -> string  | 
| 24751 | 48  | 
val intern_expr: theory -> expr -> expr  | 
| 16458 | 49  | 
val extern: theory -> string -> xstring  | 
| 
19991
 
0c9650664d47
Locales no longer generate views.  The following functions have changed
 
ballarin 
parents: 
19984 
diff
changeset
 | 
50  | 
val init: string -> theory -> Proof.context  | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
51  | 
|
| 
18795
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
52  | 
(* The specification of a locale *)  | 
| 
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
53  | 
val parameters_of: theory -> string ->  | 
| 18917 | 54  | 
((string * typ) * mixfix) list  | 
| 19276 | 55  | 
val parameters_of_expr: theory -> expr ->  | 
56  | 
((string * typ) * mixfix) list  | 
|
| 
18795
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
57  | 
val local_asms_of: theory -> string ->  | 
| 
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
58  | 
((string * Attrib.src list) * term list) list  | 
| 
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
59  | 
val global_asms_of: theory -> string ->  | 
| 
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
60  | 
((string * Attrib.src list) * term list) list  | 
| 22523 | 61  | 
val intros: theory -> string ->  | 
62  | 
thm list * thm list  | 
|
| 
18795
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
63  | 
|
| 
18899
 
a8e913c93578
theorem(_in_locale): Element.statement, Obtain.statement;
 
wenzelm 
parents: 
18890 
diff
changeset
 | 
64  | 
(* Processing of locale statements *)  | 
| 18137 | 65  | 
val read_context_statement: xstring option -> Element.context element list ->  | 
| 19585 | 66  | 
(string * string list) list list -> Proof.context ->  | 
| 
19991
 
0c9650664d47
Locales no longer generate views.  The following functions have changed
 
ballarin 
parents: 
19984 
diff
changeset
 | 
67  | 
string option * Proof.context * Proof.context * (term * term list) list list  | 
| 21035 | 68  | 
val read_context_statement_i: string option -> Element.context element list ->  | 
69  | 
(string * string list) list list -> Proof.context ->  | 
|
70  | 
string option * Proof.context * Proof.context * (term * term list) list list  | 
|
| 18137 | 71  | 
val cert_context_statement: string option -> Element.context_i element list ->  | 
| 19585 | 72  | 
(term * term list) list list -> Proof.context ->  | 
| 
19991
 
0c9650664d47
Locales no longer generate views.  The following functions have changed
 
ballarin 
parents: 
19984 
diff
changeset
 | 
73  | 
string option * Proof.context * Proof.context * (term * term list) list list  | 
| 19780 | 74  | 
val read_expr: expr -> Element.context list -> Proof.context ->  | 
75  | 
Element.context_i list * Proof.context  | 
|
76  | 
val cert_expr: expr -> Element.context_i list -> Proof.context ->  | 
|
77  | 
Element.context_i list * Proof.context  | 
|
| 15596 | 78  | 
|
79  | 
(* Diagnostic functions *)  | 
|
| 12758 | 80  | 
val print_locales: theory -> unit  | 
| 18137 | 81  | 
val print_locale: theory -> bool -> expr -> Element.context list -> unit  | 
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
82  | 
val print_registrations: bool -> string -> Proof.context -> unit  | 
| 18137 | 83  | 
|
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
84  | 
val add_locale: string option -> bstring -> expr -> Element.context list -> theory  | 
| 20965 | 85  | 
-> string * Proof.context  | 
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
86  | 
val add_locale_i: string option -> bstring -> expr -> Element.context_i list -> theory  | 
| 20965 | 87  | 
-> string * Proof.context  | 
| 15596 | 88  | 
|
| 21225 | 89  | 
(* Tactics *)  | 
90  | 
val intro_locales_tac: bool -> Proof.context -> thm list -> tactic  | 
|
91  | 
||
| 15696 | 92  | 
(* Storing results *)  | 
| 18806 | 93  | 
val add_thmss: string -> string ->  | 
94  | 
((string * Attrib.src list) * (thm list * Attrib.src list) list) list ->  | 
|
| 21582 | 95  | 
Proof.context -> Proof.context  | 
| 24020 | 96  | 
val add_type_syntax: string -> declaration -> Proof.context -> Proof.context  | 
97  | 
val add_term_syntax: string -> declaration -> Proof.context -> Proof.context  | 
|
98  | 
val add_declaration: string -> declaration -> Proof.context -> Proof.context  | 
|
| 18137 | 99  | 
|
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
100  | 
val interpretation_i: (Proof.context -> Proof.context) ->  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
101  | 
(bool * string) * Attrib.src list -> expr ->  | 
| 23918 | 102  | 
(typ Symtab.table * term Symtab.table) * term list ->  | 
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
103  | 
theory -> Proof.state  | 
| 21005 | 104  | 
val interpretation: (Proof.context -> Proof.context) ->  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
105  | 
(bool * string) * Attrib.src list -> expr ->  | 
| 23918 | 106  | 
string option list * string list ->  | 
| 21005 | 107  | 
theory -> Proof.state  | 
108  | 
val interpretation_in_locale: (Proof.context -> Proof.context) ->  | 
|
109  | 
xstring * expr -> theory -> Proof.state  | 
|
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
110  | 
val interpret_i: (Proof.state -> Proof.state Seq.seq) ->  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
111  | 
(bool * string) * Attrib.src list -> expr ->  | 
| 23918 | 112  | 
(typ Symtab.table * term Symtab.table) * term list ->  | 
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
113  | 
bool -> Proof.state -> Proof.state  | 
| 21005 | 114  | 
val interpret: (Proof.state -> Proof.state Seq.seq) ->  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
115  | 
(bool * string) * Attrib.src list -> expr ->  | 
| 23918 | 116  | 
string option list * string list ->  | 
| 21005 | 117  | 
bool -> Proof.state -> Proof.state  | 
| 11896 | 118  | 
end;  | 
| 12839 | 119  | 
|
| 12289 | 120  | 
structure Locale: LOCALE =  | 
| 11896 | 121  | 
struct  | 
122  | 
||
| 
23228
 
05fb9b2b16d7
added gen_merge_lists/merge_lists/merge_alists (legacy operations from library.ML);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
123  | 
(* legacy operations *)  | 
| 
 
05fb9b2b16d7
added gen_merge_lists/merge_lists/merge_alists (legacy operations from library.ML);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
124  | 
|
| 
 
05fb9b2b16d7
added gen_merge_lists/merge_lists/merge_alists (legacy operations from library.ML);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
125  | 
fun gen_merge_lists _ xs [] = xs  | 
| 
 
05fb9b2b16d7
added gen_merge_lists/merge_lists/merge_alists (legacy operations from library.ML);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
126  | 
| gen_merge_lists _ [] ys = ys  | 
| 
 
05fb9b2b16d7
added gen_merge_lists/merge_lists/merge_alists (legacy operations from library.ML);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
127  | 
| gen_merge_lists eq xs ys = xs @ filter_out (member eq xs) ys;  | 
| 
 
05fb9b2b16d7
added gen_merge_lists/merge_lists/merge_alists (legacy operations from library.ML);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
128  | 
|
| 
 
05fb9b2b16d7
added gen_merge_lists/merge_lists/merge_alists (legacy operations from library.ML);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
129  | 
fun merge_lists xs ys = gen_merge_lists (op =) xs ys;  | 
| 
 
05fb9b2b16d7
added gen_merge_lists/merge_lists/merge_alists (legacy operations from library.ML);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
130  | 
fun merge_alists xs = gen_merge_lists (eq_fst (op =)) xs;  | 
| 
 
05fb9b2b16d7
added gen_merge_lists/merge_lists/merge_alists (legacy operations from library.ML);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
131  | 
|
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
132  | 
fun legacy_unvarifyT thm =  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
133  | 
let  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
134  | 
val cT = Thm.ctyp_of (Thm.theory_of_thm thm);  | 
| 22700 | 135  | 
val tvars = rev (Thm.fold_terms Term.add_tvars thm []);  | 
| 20307 | 136  | 
val tfrees = map (fn ((x, _), S) => SOME (cT (TFree (x, S)))) tvars;  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
137  | 
in Drule.instantiate' tfrees [] thm end;  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
138  | 
|
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
139  | 
fun legacy_unvarify raw_thm =  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
140  | 
let  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
141  | 
val thm = legacy_unvarifyT raw_thm;  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
142  | 
val ct = Thm.cterm_of (Thm.theory_of_thm thm);  | 
| 22700 | 143  | 
val vars = rev (Thm.fold_terms Term.add_vars thm []);  | 
| 20307 | 144  | 
val frees = map (fn ((x, _), T) => SOME (ct (Free (x, T)))) vars;  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
145  | 
in Drule.instantiate' [] frees thm end;  | 
| 19780 | 146  | 
|
| 20307 | 147  | 
|
| 12273 | 148  | 
(** locale elements and expressions **)  | 
| 11896 | 149  | 
|
| 18137 | 150  | 
datatype ctxt = datatype Element.ctxt;  | 
| 17355 | 151  | 
|
| 12273 | 152  | 
datatype expr =  | 
153  | 
Locale of string |  | 
|
| 
16102
 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 
ballarin 
parents: 
16028 
diff
changeset
 | 
154  | 
Rename of expr * (string * mixfix option) option list |  | 
| 12273 | 155  | 
Merge of expr list;  | 
| 11896 | 156  | 
|
| 12273 | 157  | 
val empty = Merge [];  | 
158  | 
||
| 18137 | 159  | 
datatype 'a element =  | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
160  | 
Elem of 'a | Expr of expr;  | 
| 12273 | 161  | 
|
| 18137 | 162  | 
fun map_elem f (Elem e) = Elem (f e)  | 
163  | 
| map_elem _ (Expr e) = Expr e;  | 
|
164  | 
||
| 24020 | 165  | 
type decl = declaration * stamp;  | 
| 
21499
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
166  | 
|
| 12070 | 167  | 
type locale =  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
168  | 
 {axiom: Element.witness list,
 | 
| 
19942
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
169  | 
(* For locales that define predicates this is [A [A]], where A is the locale  | 
| 20317 | 170  | 
specification. Otherwise [].  | 
171  | 
Only required to generate the right witnesses for locales with predicates. *)  | 
|
| 22573 | 172  | 
imports: expr, (*dynamic imports*)  | 
| 19783 | 173  | 
elems: (Element.context_i * stamp) list,  | 
174  | 
(* Static content, neither Fixes nor Constrains elements *)  | 
|
| 
21499
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
175  | 
params: ((string * typ) * mixfix) list, (*all params*)  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
176  | 
lparams: string list, (*local params*)  | 
| 
21665
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
177  | 
decls: decl list * decl list, (*type/term_syntax declarations*)  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
178  | 
regs: ((string * string list) * Element.witness list) list,  | 
| 19780 | 179  | 
(* Registrations: indentifiers and witnesses of locales interpreted in the locale. *)  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
180  | 
intros: thm list * thm list}  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
181  | 
(* Introduction rules: of delta predicate and locale predicate. *)  | 
| 11896 | 182  | 
|
| 15703 | 183  | 
(* CB: an internal (Int) locale element was either imported or included,  | 
184  | 
an external (Ext) element appears directly in the locale text. *)  | 
|
185  | 
||
186  | 
datatype ('a, 'b) int_ext = Int of 'a | Ext of 'b;
 | 
|
187  | 
||
188  | 
||
189  | 
||
| 
16736
 
1e792b32abef
Preparations for interpretation of locales in locales.
 
ballarin 
parents: 
16620 
diff
changeset
 | 
190  | 
(** management of registrations in theories and proof contexts **)  | 
| 11896 | 191  | 
|
| 15837 | 192  | 
structure Registrations :  | 
193  | 
sig  | 
|
194  | 
type T  | 
|
195  | 
val empty: T  | 
|
196  | 
val join: T * T -> T  | 
|
| 
20069
 
77a6b62418bb
Witness theorems of interpretations now transfered to current theory.
 
ballarin 
parents: 
20059 
diff
changeset
 | 
197  | 
val dest: theory -> T ->  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
198  | 
(term list *  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
199  | 
(((bool * string) * Attrib.src list) * Element.witness list *  | 
| 23918 | 200  | 
Thm.thm Termtab.table)) list  | 
| 16458 | 201  | 
val lookup: theory -> T * term list ->  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
202  | 
(((bool * string) * Attrib.src list) * Element.witness list *  | 
| 23918 | 203  | 
Thm.thm Termtab.table) option  | 
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
204  | 
val insert: theory -> term list * ((bool * string) * Attrib.src list) -> T ->  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
205  | 
T * (term list * (((bool * string) * Attrib.src list) * Element.witness list)) list  | 
| 19780 | 206  | 
val add_witness: term list -> Element.witness -> T -> T  | 
| 23918 | 207  | 
val add_equation: term list -> Thm.thm -> T -> T  | 
| 15837 | 208  | 
end =  | 
209  | 
struct  | 
|
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
210  | 
(* A registration is indexed by parameter instantiation. Its components are:  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
211  | 
- parameters and prefix  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
212  | 
(if the Boolean flag is set, only accesses containing the prefix are generated,  | 
| 23418 | 213  | 
otherwise the prefix may be omitted when accessing theorems etc.)  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
214  | 
- theorems (actually witnesses) instantiating locale assumptions  | 
| 23918 | 215  | 
- theorems (equations) interpreting derived concepts and indexed by lhs  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
216  | 
*)  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
217  | 
type T = (((bool * string) * Attrib.src list) * Element.witness list *  | 
| 23918 | 218  | 
Thm.thm Termtab.table) Termtab.table;  | 
| 15837 | 219  | 
|
220  | 
val empty = Termtab.empty;  | 
|
221  | 
||
222  | 
(* term list represented as single term, for simultaneous matching *)  | 
|
223  | 
fun termify ts =  | 
|
| 18343 | 224  | 
    Term.list_comb (Const ("", map fastype_of ts ---> propT), ts);
 | 
| 15837 | 225  | 
fun untermify t =  | 
226  | 
let fun ut (Const _) ts = ts  | 
|
227  | 
| ut (s $ t) ts = ut s (t::ts)  | 
|
228  | 
in ut t [] end;  | 
|
229  | 
||
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
230  | 
(* joining of registrations:  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
231  | 
- prefix and attributes of right theory;  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
232  | 
- witnesses are equal, no attempt to subsumption testing;  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
233  | 
- union of equalities, if conflicting (i.e. two eqns with equal lhs)  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
234  | 
eqn of right theory takes precedence *)  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
235  | 
fun join (r1, r2) = Termtab.join (fn _ => fn ((_, _, e1), (n, w, e2)) =>  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
236  | 
(n, w, Termtab.join (fn _ => fn (_, e) => e) (e1, e2))) (r1, r2);  | 
| 15837 | 237  | 
|
| 
20069
 
77a6b62418bb
Witness theorems of interpretations now transfered to current theory.
 
ballarin 
parents: 
20059 
diff
changeset
 | 
238  | 
fun dest_transfer thy regs =  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
239  | 
Termtab.dest regs |> map (apsnd (fn (n, ws, es) =>  | 
| 23918 | 240  | 
(n, map (Element.transfer_witness thy) ws, Termtab.map (Thm.transfer thy) es)));  | 
| 
20069
 
77a6b62418bb
Witness theorems of interpretations now transfered to current theory.
 
ballarin 
parents: 
20059 
diff
changeset
 | 
241  | 
|
| 
 
77a6b62418bb
Witness theorems of interpretations now transfered to current theory.
 
ballarin 
parents: 
20059 
diff
changeset
 | 
242  | 
fun dest thy regs = dest_transfer thy regs |> map (apfst untermify);  | 
| 15837 | 243  | 
|
244  | 
(* registrations that subsume t *)  | 
|
| 17203 | 245  | 
fun subsumers thy t regs =  | 
| 
20069
 
77a6b62418bb
Witness theorems of interpretations now transfered to current theory.
 
ballarin 
parents: 
20059 
diff
changeset
 | 
246  | 
filter (fn (t', _) => Pattern.matches thy (t', t)) (dest_transfer thy regs);  | 
| 15837 | 247  | 
|
248  | 
(* look up registration, pick one that subsumes the query *)  | 
|
| 19780 | 249  | 
fun lookup thy (regs, ts) =  | 
| 15837 | 250  | 
let  | 
251  | 
val t = termify ts;  | 
|
| 19780 | 252  | 
val subs = subsumers thy t regs;  | 
| 
21483
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
253  | 
in  | 
| 
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
254  | 
(case subs of  | 
| 15837 | 255  | 
[] => NONE  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
256  | 
| ((t', (attn, thms, eqns)) :: _) =>  | 
| 
21483
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
257  | 
let  | 
| 19780 | 258  | 
val (tinst, inst) = Pattern.match thy (t', t) (Vartab.empty, Vartab.empty);  | 
| 15837 | 259  | 
(* thms contain Frees, not Vars *)  | 
| 
21665
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
260  | 
val tinst' = tinst |> Vartab.dest (* FIXME Vartab.map (!?) *)  | 
| 19810 | 261  | 
|> map (fn ((x, 0), (_, T)) => (x, Logic.legacy_unvarifyT T))  | 
| 15837 | 262  | 
|> Symtab.make;  | 
263  | 
val inst' = inst |> Vartab.dest  | 
|
| 19810 | 264  | 
|> map (fn ((x, 0), (_, t)) => (x, Logic.legacy_unvarify t))  | 
| 15837 | 265  | 
|> Symtab.make;  | 
| 23918 | 266  | 
val inst_morph = Element.inst_morphism thy (tinst', inst');  | 
267  | 
in SOME (attn, map (Element.morph_witness inst_morph) thms,  | 
|
268  | 
Termtab.map (Morphism.thm inst_morph) eqns)  | 
|
269  | 
end)  | 
|
| 15837 | 270  | 
end;  | 
271  | 
||
272  | 
(* add registration if not subsumed by ones already present,  | 
|
273  | 
additionally returns registrations that are strictly subsumed *)  | 
|
| 
20069
 
77a6b62418bb
Witness theorems of interpretations now transfered to current theory.
 
ballarin 
parents: 
20059 
diff
changeset
 | 
274  | 
fun insert thy (ts, attn) regs =  | 
| 15837 | 275  | 
let  | 
276  | 
val t = termify ts;  | 
|
| 
20069
 
77a6b62418bb
Witness theorems of interpretations now transfered to current theory.
 
ballarin 
parents: 
20059 
diff
changeset
 | 
277  | 
val subs = subsumers thy t regs ;  | 
| 15837 | 278  | 
in (case subs of  | 
279  | 
[] => let  | 
|
280  | 
val sups =  | 
|
| 
20069
 
77a6b62418bb
Witness theorems of interpretations now transfered to current theory.
 
ballarin 
parents: 
20059 
diff
changeset
 | 
281  | 
filter (fn (t', _) => Pattern.matches thy (t, t')) (dest_transfer thy regs);  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
282  | 
val sups' = map (apfst untermify) sups |> map (fn (ts, (n, w, d)) => (ts, (n, w)))  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
283  | 
in (Termtab.update (t, (attn, [], Termtab.empty)) regs, sups') end  | 
| 15837 | 284  | 
| _ => (regs, []))  | 
285  | 
end;  | 
|
286  | 
||
| 23918 | 287  | 
fun gen_add f ts thm regs =  | 
288  | 
let  | 
|
289  | 
val t = termify ts;  | 
|
290  | 
in  | 
|
291  | 
Termtab.update (t, f thm (the (Termtab.lookup regs t))) regs  | 
|
292  | 
end;  | 
|
293  | 
||
| 15837 | 294  | 
(* add witness theorem to registration,  | 
| 
16169
 
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
 
ballarin 
parents: 
16144 
diff
changeset
 | 
295  | 
only if instantiation is exact, otherwise exception Option raised *)  | 
| 15837 | 296  | 
fun add_witness ts thm regs =  | 
| 23918 | 297  | 
gen_add (fn thm => fn (x, thms, eqns) => (x, thm::thms, eqns)) ts thm regs;  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
298  | 
|
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
299  | 
(* add equation to registration, replaces previous equation with same lhs;  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
300  | 
only if instantiation is exact, otherwise exception Option raised;  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
301  | 
exception TERM raised if not a meta equality *)  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
302  | 
fun add_equation ts thm regs =  | 
| 23918 | 303  | 
gen_add (fn thm => fn (x, thms, eqns) =>  | 
304  | 
(x, thms, Termtab.update (thm |> prop_of |> Logic.dest_equals |> fst, thm) eqns))  | 
|
305  | 
ts thm regs;  | 
|
| 15837 | 306  | 
end;  | 
307  | 
||
| 
16736
 
1e792b32abef
Preparations for interpretation of locales in locales.
 
ballarin 
parents: 
16620 
diff
changeset
 | 
308  | 
|
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
309  | 
(** theory data : locales **)  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
310  | 
|
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
311  | 
structure LocalesData = TheoryDataFun  | 
| 22846 | 312  | 
(  | 
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
313  | 
type T = NameSpace.T * locale Symtab.table;  | 
| 15596 | 314  | 
(* 1st entry: locale namespace,  | 
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
315  | 
2nd entry: locales of the theory *)  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
316  | 
|
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
317  | 
val empty = (NameSpace.empty, Symtab.empty);  | 
| 12063 | 318  | 
val copy = I;  | 
| 16458 | 319  | 
val extend = I;  | 
| 12289 | 320  | 
|
| 
21665
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
321  | 
fun join_locales _  | 
| 22573 | 322  | 
    ({axiom, imports, elems, params, lparams, decls = (decls1, decls2), regs, intros}: locale,
 | 
| 
21665
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
323  | 
      {elems = elems', decls = (decls1', decls2'), regs = regs', ...}: locale) =
 | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
324  | 
     {axiom = axiom,
 | 
| 22573 | 325  | 
imports = imports,  | 
| 17496 | 326  | 
elems = gen_merge_lists (eq_snd (op =)) elems elems',  | 
| 
16736
 
1e792b32abef
Preparations for interpretation of locales in locales.
 
ballarin 
parents: 
16620 
diff
changeset
 | 
327  | 
params = params,  | 
| 19278 | 328  | 
lparams = lparams,  | 
| 
21499
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
329  | 
decls =  | 
| 
21665
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
330  | 
(Library.merge (eq_snd (op =)) (decls1, decls1'),  | 
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
331  | 
Library.merge (eq_snd (op =)) (decls2, decls2')),  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
332  | 
regs = merge_alists regs regs',  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
333  | 
intros = intros};  | 
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
334  | 
fun merge _ ((space1, locs1), (space2, locs2)) =  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
335  | 
(NameSpace.merge (space1, space2), Symtab.join join_locales (locs1, locs2));  | 
| 22846 | 336  | 
);  | 
| 15801 | 337  | 
|
338  | 
||
| 
15624
 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 
ballarin 
parents: 
15623 
diff
changeset
 | 
339  | 
|
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
340  | 
(** context data : registrations **)  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
341  | 
|
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
342  | 
structure RegistrationsData = GenericDataFun  | 
| 22846 | 343  | 
(  | 
344  | 
type T = Registrations.T Symtab.table; (*registrations, indexed by locale name*)  | 
|
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
345  | 
val empty = Symtab.empty;  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
346  | 
val extend = I;  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
347  | 
fun merge _ = Symtab.join (K Registrations.join);  | 
| 22846 | 348  | 
);  | 
| 12289 | 349  | 
|
| 12277 | 350  | 
|
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
351  | 
(** access locales **)  | 
| 12277 | 352  | 
|
| 22846 | 353  | 
fun print_locales thy =  | 
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
354  | 
let val (space, locs) = LocalesData.get thy in  | 
| 22846 | 355  | 
    Pretty.strs ("locales:" :: map #1 (NameSpace.extern_table (space, locs)))
 | 
356  | 
|> Pretty.writeln  | 
|
357  | 
end;  | 
|
| 
15624
 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 
ballarin 
parents: 
15623 
diff
changeset
 | 
358  | 
|
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
359  | 
val intern = NameSpace.intern o #1 o LocalesData.get;  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
360  | 
val extern = NameSpace.extern o #1 o LocalesData.get;  | 
| 
15624
 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 
ballarin 
parents: 
15623 
diff
changeset
 | 
361  | 
|
| 16144 | 362  | 
fun declare_locale name thy =  | 
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
363  | 
thy |> LocalesData.map (fn (space, locs) =>  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
364  | 
(Sign.declare_name thy name space, locs));  | 
| 11896 | 365  | 
|
| 15596 | 366  | 
fun put_locale name loc =  | 
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
367  | 
LocalesData.map (fn (space, locs) =>  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
368  | 
(space, Symtab.update (name, loc) locs));  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
369  | 
|
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
370  | 
fun get_locale thy name = Symtab.lookup (#2 (LocalesData.get thy)) name;  | 
| 11896 | 371  | 
|
| 12014 | 372  | 
fun the_locale thy name =  | 
373  | 
(case get_locale thy name of  | 
|
| 15531 | 374  | 
SOME loc => loc  | 
375  | 
  | NONE => error ("Unknown locale " ^ quote name));
 | 
|
| 11896 | 376  | 
|
| 18806 | 377  | 
fun change_locale name f thy =  | 
378  | 
let  | 
|
| 22573 | 379  | 
    val {axiom, imports, elems, params, lparams, decls, regs, intros} =
 | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
380  | 
the_locale thy name;  | 
| 22573 | 381  | 
val (axiom', imports', elems', params', lparams', decls', regs', intros') =  | 
382  | 
f (axiom, imports, elems, params, lparams, decls, regs, intros);  | 
|
| 18806 | 383  | 
in  | 
| 22573 | 384  | 
    put_locale name {axiom = axiom', imports = imports', elems = elems',
 | 
| 
21499
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
385  | 
params = params', lparams = lparams', decls = decls', regs = regs',  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
386  | 
intros = intros'} thy  | 
| 18806 | 387  | 
end;  | 
388  | 
||
| 12046 | 389  | 
|
| 15596 | 390  | 
(* access registrations *)  | 
391  | 
||
| 15696 | 392  | 
(* Ids of global registrations are varified,  | 
393  | 
Ids of local registrations are not.  | 
|
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
394  | 
Witness thms of registrations are never varified.  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
395  | 
Varification of eqns as varification of ids. *)  | 
| 15696 | 396  | 
|
| 
15624
 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 
ballarin 
parents: 
15623 
diff
changeset
 | 
397  | 
(* retrieve registration from theory or context *)  | 
| 
 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 
ballarin 
parents: 
15623 
diff
changeset
 | 
398  | 
|
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
399  | 
fun get_registrations ctxt name =  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
400  | 
case Symtab.lookup (RegistrationsData.get ctxt) name of  | 
| 15696 | 401  | 
NONE => []  | 
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
402  | 
| SOME reg => Registrations.dest (Context.theory_of ctxt) reg;  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
403  | 
|
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
404  | 
fun get_global_registrations thy = get_registrations (Context.Theory thy);  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
405  | 
fun get_local_registrations ctxt = get_registrations (Context.Proof ctxt);  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
406  | 
|
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
407  | 
|
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
408  | 
fun get_registration ctxt (name, ps) =  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
409  | 
case Symtab.lookup (RegistrationsData.get ctxt) name of  | 
| 
15624
 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 
ballarin 
parents: 
15623 
diff
changeset
 | 
410  | 
NONE => NONE  | 
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
411  | 
| SOME reg => Registrations.lookup (Context.theory_of ctxt) (reg, ps);  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
412  | 
|
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
413  | 
fun get_global_registration thy = get_registration (Context.Theory thy);  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
414  | 
fun get_local_registration ctxt = get_registration (Context.Proof ctxt);  | 
| 15596 | 415  | 
|
| 18343 | 416  | 
val test_global_registration = is_some oo get_global_registration;  | 
417  | 
val test_local_registration = is_some oo get_local_registration;  | 
|
| 
15624
 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 
ballarin 
parents: 
15623 
diff
changeset
 | 
418  | 
|
| 15837 | 419  | 
(* add registration to theory or context, ignored if subsumed *)  | 
| 
15624
 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 
ballarin 
parents: 
15623 
diff
changeset
 | 
420  | 
|
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
421  | 
fun put_registration (name, ps) attn ctxt =  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
422  | 
RegistrationsData.map (fn regs =>  | 
| 15837 | 423  | 
let  | 
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
424  | 
val thy = Context.theory_of ctxt;  | 
| 18343 | 425  | 
val reg = the_default Registrations.empty (Symtab.lookup regs name);  | 
| 16458 | 426  | 
val (reg', sups) = Registrations.insert thy (ps, attn) reg;  | 
| 15837 | 427  | 
val _ = if not (null sups) then warning  | 
428  | 
                ("Subsumed interpretation(s) of locale " ^
 | 
|
| 16458 | 429  | 
quote (extern thy name) ^  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
430  | 
"\nwith the following prefix(es):" ^  | 
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
431  | 
commas_quote (map (fn (_, (((_, s), _), _)) => s) sups))  | 
| 15837 | 432  | 
else ();  | 
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
433  | 
in Symtab.update (name, reg') regs end) ctxt;  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
434  | 
|
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
435  | 
fun put_global_registration id attn = Context.theory_map (put_registration id attn);  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
436  | 
fun put_local_registration id attn = Context.proof_map (put_registration id attn);  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
437  | 
|
| 15596 | 438  | 
|
| 18806 | 439  | 
fun put_registration_in_locale name id =  | 
| 22573 | 440  | 
change_locale name (fn (axiom, imports, elems, params, lparams, decls, regs, intros) =>  | 
441  | 
(axiom, imports, elems, params, lparams, decls, regs @ [(id, [])], intros));  | 
|
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
442  | 
|
| 
15624
 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 
ballarin 
parents: 
15623 
diff
changeset
 | 
443  | 
|
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
444  | 
(* add witness theorem to registration, ignored if registration not present *)  | 
| 15596 | 445  | 
|
| 
18123
 
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
 
wenzelm 
parents: 
18038 
diff
changeset
 | 
446  | 
fun add_witness (name, ps) thm =  | 
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
447  | 
RegistrationsData.map (Symtab.map_entry name (Registrations.add_witness ps thm));  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
448  | 
|
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
449  | 
fun add_global_witness id thm = Context.theory_map (add_witness id thm);  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
450  | 
fun add_local_witness id thm = Context.proof_map (add_witness id thm);  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
451  | 
|
| 15596 | 452  | 
|
| 18806 | 453  | 
fun add_witness_in_locale name id thm =  | 
| 22573 | 454  | 
change_locale name (fn (axiom, imports, elems, params, lparams, decls, regs, intros) =>  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
455  | 
let  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
456  | 
fun add (id', thms) =  | 
| 18806 | 457  | 
if id = id' then (id', thm :: thms) else (id', thms);  | 
| 22573 | 458  | 
in (axiom, imports, elems, params, lparams, decls, map add regs, intros) end);  | 
| 15596 | 459  | 
|
| 
14215
 
ebf291f3b449
Improvements to Isar/Locales:  premises generated by "includes" elements
 
ballarin 
parents: 
13629 
diff
changeset
 | 
460  | 
|
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
461  | 
(* add equation to registration, ignored if registration not present *)  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
462  | 
|
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
463  | 
fun add_equation (name, ps) thm =  | 
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
464  | 
RegistrationsData.map (Symtab.map_entry name (Registrations.add_equation ps thm));  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
465  | 
|
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
466  | 
fun add_global_equation id thm = Context.theory_map (add_equation id thm);  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
467  | 
fun add_local_equation id thm = Context.proof_map (add_equation id thm);  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
468  | 
|
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
469  | 
|
| 
15624
 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 
ballarin 
parents: 
15623 
diff
changeset
 | 
470  | 
(* printing of registrations *)  | 
| 15596 | 471  | 
|
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
472  | 
fun print_registrations show_wits loc ctxt =  | 
| 15596 | 473  | 
let  | 
| 15703 | 474  | 
val thy = ProofContext.theory_of ctxt;  | 
| 24920 | 475  | 
val prt_term = Pretty.quote o Syntax.pretty_term ctxt;  | 
| 23918 | 476  | 
val prt_thm = prt_term o prop_of;  | 
| 
17096
 
8327b71282ce
Improved generation of witnesses in interpretation.
 
ballarin 
parents: 
17033 
diff
changeset
 | 
477  | 
fun prt_inst ts =  | 
| 
 
8327b71282ce
Improved generation of witnesses in interpretation.
 
ballarin 
parents: 
17033 
diff
changeset
 | 
478  | 
        Pretty.enclose "(" ")" (Pretty.breaks (map prt_term ts));
 | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
479  | 
fun prt_attn ((false, prfx), atts) =  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
480  | 
Pretty.breaks (Pretty.str prfx :: Pretty.str "(optional)" ::  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
481  | 
Attrib.pretty_attribs ctxt atts)  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
482  | 
| prt_attn ((true, prfx), atts) =  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
483  | 
Pretty.breaks (Pretty.str prfx :: Attrib.pretty_attribs ctxt atts);  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
484  | 
fun prt_eqns [] = Pretty.str "no equations."  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
485  | 
| prt_eqns eqns = Pretty.block (Pretty.str "equations:" :: Pretty.brk 1 ::  | 
| 23918 | 486  | 
Pretty.breaks (map prt_thm eqns));  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
487  | 
fun prt_core ts eqns =  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
488  | 
[prt_inst ts, Pretty.fbrk, prt_eqns (Termtab.dest eqns |> map snd)];  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
489  | 
fun prt_witns [] = Pretty.str "no witnesses."  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
490  | 
| prt_witns witns = Pretty.block (Pretty.str "witnesses:" :: Pretty.brk 1 ::  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
491  | 
Pretty.breaks (map (Element.pretty_witness ctxt) witns))  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
492  | 
fun prt_reg (ts, (((_, ""), []), witns, eqns)) =  | 
| 17138 | 493  | 
if show_wits  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
494  | 
then Pretty.block (prt_core ts eqns @ [Pretty.fbrk, prt_witns witns])  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
495  | 
else Pretty.block (prt_core ts eqns)  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
496  | 
| prt_reg (ts, (attn, witns, eqns)) =  | 
| 17138 | 497  | 
if show_wits  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
498  | 
then Pretty.block ((prt_attn attn @ [Pretty.str ":", Pretty.brk 1] @  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
499  | 
prt_core ts eqns @ [Pretty.fbrk, prt_witns witns]))  | 
| 
17096
 
8327b71282ce
Improved generation of witnesses in interpretation.
 
ballarin 
parents: 
17033 
diff
changeset
 | 
500  | 
else Pretty.block ((prt_attn attn @  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
501  | 
[Pretty.str ":", Pretty.brk 1] @ prt_core ts eqns));  | 
| 15703 | 502  | 
|
| 16458 | 503  | 
val loc_int = intern thy loc;  | 
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
504  | 
val regs = RegistrationsData.get (Context.Proof ctxt);  | 
| 17412 | 505  | 
val loc_regs = Symtab.lookup regs loc_int;  | 
| 15596 | 506  | 
in  | 
507  | 
(case loc_regs of  | 
|
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
508  | 
        NONE => Pretty.str ("no interpretations")
 | 
| 
15763
 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 
ballarin 
parents: 
15749 
diff
changeset
 | 
509  | 
| SOME r => let  | 
| 
20069
 
77a6b62418bb
Witness theorems of interpretations now transfered to current theory.
 
ballarin 
parents: 
20059 
diff
changeset
 | 
510  | 
val r' = Registrations.dest thy r;  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
511  | 
val r'' = Library.sort_wrt (fn (_, (((_, prfx), _), _, _)) => prfx) r';  | 
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
512  | 
          in Pretty.big_list ("interpretations:") (map prt_reg r'') end)
 | 
| 15596 | 513  | 
|> Pretty.writeln  | 
514  | 
end;  | 
|
515  | 
||
516  | 
||
| 12277 | 517  | 
(* diagnostics *)  | 
| 12273 | 518  | 
|
| 12277 | 519  | 
fun err_in_locale ctxt msg ids =  | 
520  | 
let  | 
|
| 16458 | 521  | 
val thy = ProofContext.theory_of ctxt;  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
522  | 
fun prt_id (name, parms) =  | 
| 16458 | 523  | 
[Pretty.block (Pretty.breaks (map Pretty.str (extern thy name :: parms)))];  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19423 
diff
changeset
 | 
524  | 
val prt_ids = flat (separate [Pretty.str " +", Pretty.brk 1] (map prt_id ids));  | 
| 12502 | 525  | 
val err_msg =  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
526  | 
if forall (equal "" o #1) ids then msg  | 
| 12502 | 527  | 
else msg ^ "\n" ^ Pretty.string_of (Pretty.block  | 
528  | 
(Pretty.str "The error(s) above occurred in locale:" :: Pretty.brk 1 :: prt_ids));  | 
|
| 18678 | 529  | 
in error err_msg end;  | 
| 12063 | 530  | 
|
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
531  | 
fun err_in_locale' ctxt msg ids' = err_in_locale ctxt msg (map fst ids');  | 
| 12277 | 532  | 
|
533  | 
||
| 19783 | 534  | 
fun pretty_ren NONE = Pretty.str "_"  | 
535  | 
| pretty_ren (SOME (x, NONE)) = Pretty.str x  | 
|
536  | 
| pretty_ren (SOME (x, SOME syn)) =  | 
|
537  | 
Pretty.block [Pretty.str x, Pretty.brk 1, Syntax.pretty_mixfix syn];  | 
|
538  | 
||
539  | 
fun pretty_expr thy (Locale name) = Pretty.str (extern thy name)  | 
|
540  | 
| pretty_expr thy (Rename (expr, xs)) =  | 
|
541  | 
Pretty.block [pretty_expr thy expr, Pretty.brk 1, Pretty.block (map pretty_ren xs |> Pretty.breaks)]  | 
|
542  | 
| pretty_expr thy (Merge es) =  | 
|
543  | 
Pretty.separate "+" (map (pretty_expr thy) es) |> Pretty.block;  | 
|
544  | 
||
545  | 
fun err_in_expr _ msg (Merge []) = error msg  | 
|
546  | 
| err_in_expr ctxt msg expr =  | 
|
547  | 
error (msg ^ "\n" ^ Pretty.string_of (Pretty.block  | 
|
548  | 
[Pretty.str "The error(s) above occured in locale expression:", Pretty.brk 1,  | 
|
549  | 
pretty_expr (ProofContext.theory_of ctxt) expr]));  | 
|
550  | 
||
| 12046 | 551  | 
|
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
552  | 
(** structured contexts: rename + merge + implicit type instantiation **)  | 
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
553  | 
|
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
554  | 
(* parameter types *)  | 
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
555  | 
|
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
556  | 
fun frozen_tvars ctxt Ts =  | 
| 19914 | 557  | 
#1 (Variable.importT_inst (map Logic.mk_type Ts) ctxt)  | 
| 
19900
 
21a99d88d925
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19873 
diff
changeset
 | 
558  | 
|> map (fn ((xi, S), T) => (xi, (S, T)));  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
559  | 
|
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
560  | 
fun unify_frozen ctxt maxidx Ts Us =  | 
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
561  | 
let  | 
| 18343 | 562  | 
fun paramify NONE i = (NONE, i)  | 
563  | 
| paramify (SOME T) i = apfst SOME (TypeInfer.paramify_dummies T i);  | 
|
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
564  | 
|
| 18343 | 565  | 
val (Ts', maxidx') = fold_map paramify Ts maxidx;  | 
566  | 
val (Us', maxidx'') = fold_map paramify Us maxidx';  | 
|
| 16947 | 567  | 
val thy = ProofContext.theory_of ctxt;  | 
| 
14215
 
ebf291f3b449
Improvements to Isar/Locales:  premises generated by "includes" elements
 
ballarin 
parents: 
13629 
diff
changeset
 | 
568  | 
|
| 18190 | 569  | 
fun unify (SOME T, SOME U) env = (Sign.typ_unify thy (U, T) env  | 
570  | 
          handle Type.TUNIFY => raise TYPE ("unify_frozen: failed to unify types", [U, T], []))
 | 
|
571  | 
| unify _ env = env;  | 
|
572  | 
val (unifier, _) = fold unify (Ts' ~~ Us') (Vartab.empty, maxidx'');  | 
|
| 15570 | 573  | 
val Vs = map (Option.map (Envir.norm_type unifier)) Us';  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19423 
diff
changeset
 | 
574  | 
val unifier' = Vartab.extend (unifier, frozen_tvars ctxt (map_filter I Vs));  | 
| 15570 | 575  | 
in map (Option.map (Envir.norm_type unifier')) Vs end;  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
576  | 
|
| 
21665
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
577  | 
fun params_of elemss =  | 
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
578  | 
distinct (eq_fst (op = : string * string -> bool)) (maps (snd o fst) elemss);  | 
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
579  | 
|
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
580  | 
fun params_of' elemss =  | 
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
581  | 
distinct (eq_fst (op = : string * string -> bool)) (maps (snd o fst o fst) elemss);  | 
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
582  | 
|
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
583  | 
|
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
584  | 
fun params_prefix params = space_implode "_" params;  | 
| 
16102
 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 
ballarin 
parents: 
16028 
diff
changeset
 | 
585  | 
|
| 
14508
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
586  | 
|
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
587  | 
(* CB: param_types has the following type:  | 
| 15531 | 588  | 
  ('a * 'b option) list -> ('a * 'b) list *)
 | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19423 
diff
changeset
 | 
589  | 
fun param_types ps = map_filter (fn (_, NONE) => NONE | (x, SOME T) => SOME (x, T)) ps;  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
590  | 
|
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
591  | 
|
| 19932 | 592  | 
fun merge_syntax ctxt ids ss = Symtab.merge (op = : mixfix * mixfix -> bool) ss  | 
| 
23655
 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 
wenzelm 
parents: 
23591 
diff
changeset
 | 
593  | 
handle Symtab.DUP x => err_in_locale ctxt  | 
| 
 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 
wenzelm 
parents: 
23591 
diff
changeset
 | 
594  | 
    ("Conflicting syntax for parameter: " ^ quote x) (map fst ids);
 | 
| 
16102
 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 
ballarin 
parents: 
16028 
diff
changeset
 | 
595  | 
|
| 
 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 
ballarin 
parents: 
16028 
diff
changeset
 | 
596  | 
|
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
597  | 
(* Distinction of assumed vs. derived identifiers.  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
598  | 
The former may have axioms relating assumptions of the context to  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
599  | 
assumptions of the specification fragment (for locales with  | 
| 19780 | 600  | 
predicates). The latter have witnesses relating assumptions of the  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
601  | 
specification fragment to assumptions of other (assumed) specification  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
602  | 
fragments. *)  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
603  | 
|
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
604  | 
datatype 'a mode = Assumed of 'a | Derived of 'a;  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
605  | 
|
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
606  | 
fun map_mode f (Assumed x) = Assumed (f x)  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
607  | 
| map_mode f (Derived x) = Derived (f x);  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
608  | 
|
| 
18123
 
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
 
wenzelm 
parents: 
18038 
diff
changeset
 | 
609  | 
|
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
610  | 
(* flatten expressions *)  | 
| 11896 | 611  | 
|
| 12510 | 612  | 
local  | 
| 12502 | 613  | 
|
| 18137 | 614  | 
fun unify_parms ctxt fixed_parms raw_parmss =  | 
| 12502 | 615  | 
let  | 
| 16458 | 616  | 
val thy = ProofContext.theory_of ctxt;  | 
| 12502 | 617  | 
val maxidx = length raw_parmss;  | 
618  | 
val idx_parmss = (0 upto maxidx - 1) ~~ raw_parmss;  | 
|
619  | 
||
620  | 
fun varify i = Term.map_type_tfree (fn (a, S) => TVar ((a, i), S));  | 
|
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
621  | 
fun varify_parms (i, ps) = map (apsnd (varify i)) (param_types ps);  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19423 
diff
changeset
 | 
622  | 
val parms = fixed_parms @ maps varify_parms idx_parmss;  | 
| 12502 | 623  | 
|
| 18137 | 624  | 
fun unify T U envir = Sign.typ_unify thy (U, T) envir  | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
625  | 
handle Type.TUNIFY =>  | 
| 23418 | 626  | 
let  | 
| 
22339
 
0dc6b45e5662
more precise error message in parameter unification
 
schirmer 
parents: 
22300 
diff
changeset
 | 
627  | 
val T' = Envir.norm_type (fst envir) T;  | 
| 
 
0dc6b45e5662
more precise error message in parameter unification
 
schirmer 
parents: 
22300 
diff
changeset
 | 
628  | 
val U' = Envir.norm_type (fst envir) U;  | 
| 23418 | 629  | 
val prt = Sign.string_of_typ thy;  | 
| 
22339
 
0dc6b45e5662
more precise error message in parameter unification
 
schirmer 
parents: 
22300 
diff
changeset
 | 
630  | 
in  | 
| 18137 | 631  | 
          raise TYPE ("unify_parms: failed to unify types " ^
 | 
| 
22339
 
0dc6b45e5662
more precise error message in parameter unification
 
schirmer 
parents: 
22300 
diff
changeset
 | 
632  | 
prt U' ^ " and " ^ prt T', [U', T'], [])  | 
| 18137 | 633  | 
end;  | 
634  | 
fun unify_list (T :: Us) = fold (unify T) Us  | 
|
635  | 
| unify_list [] = I;  | 
|
| 18952 | 636  | 
val (unifier, _) = fold unify_list (map #2 (Symtab.dest (Symtab.make_list parms)))  | 
| 18137 | 637  | 
(Vartab.empty, maxidx);  | 
| 12502 | 638  | 
|
| 
19061
 
ffbbac0261c9
removed distinct, renamed gen_distinct to distinct;
 
wenzelm 
parents: 
19025 
diff
changeset
 | 
639  | 
val parms' = map (apsnd (Envir.norm_type unifier)) (distinct (eq_fst (op =)) parms);  | 
| 12502 | 640  | 
val unifier' = Vartab.extend (unifier, frozen_tvars ctxt (map #2 parms'));  | 
641  | 
||
642  | 
fun inst_parms (i, ps) =  | 
|
| 23178 | 643  | 
List.foldr Term.add_typ_tfrees [] (map_filter snd ps)  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19423 
diff
changeset
 | 
644  | 
|> map_filter (fn (a, S) =>  | 
| 12502 | 645  | 
let val T = Envir.norm_type unifier' (TVar ((a, i), S))  | 
| 18137 | 646  | 
in if T = TFree (a, S) then NONE else SOME (a, T) end)  | 
647  | 
|> Symtab.make;  | 
|
648  | 
in map inst_parms idx_parmss end;  | 
|
| 12502 | 649  | 
|
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
650  | 
in  | 
| 12502 | 651  | 
|
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
652  | 
fun unify_elemss _ _ [] = []  | 
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
653  | 
| unify_elemss _ [] [elems] = [elems]  | 
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
654  | 
| unify_elemss ctxt fixed_parms elemss =  | 
| 12502 | 655  | 
let  | 
| 18137 | 656  | 
val thy = ProofContext.theory_of ctxt;  | 
| 
21483
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
657  | 
val phis = unify_parms ctxt fixed_parms (map (snd o fst o fst) elemss)  | 
| 
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
658  | 
|> map (Element.instT_morphism thy);  | 
| 
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
659  | 
fun inst ((((name, ps), mode), elems), phi) =  | 
| 
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
660  | 
(((name, map (apsnd (Option.map (Morphism.typ phi))) ps),  | 
| 
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
661  | 
map_mode (map (Element.morph_witness phi)) mode),  | 
| 
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
662  | 
map (Element.morph_ctxt phi) elems);  | 
| 
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
663  | 
in map inst (elemss ~~ phis) end;  | 
| 12502 | 664  | 
|
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
665  | 
|
| 19810 | 666  | 
fun renaming xs parms = zip_options parms xs  | 
667  | 
handle Library.UnequalLengths =>  | 
|
668  | 
    error ("Too many arguments in renaming: " ^
 | 
|
669  | 
commas (map (fn NONE => "_" | SOME x => quote (fst x)) xs));  | 
|
670  | 
||
671  | 
||
| 19783 | 672  | 
(* params_of_expr:  | 
673  | 
Compute parameters (with types and syntax) of locale expression.  | 
|
674  | 
*)  | 
|
675  | 
||
676  | 
fun params_of_expr ctxt fixed_params expr (prev_parms, prev_types, prev_syn) =  | 
|
677  | 
let  | 
|
678  | 
val thy = ProofContext.theory_of ctxt;  | 
|
679  | 
||
680  | 
fun merge_tenvs fixed tenv1 tenv2 =  | 
|
681  | 
let  | 
|
682  | 
val [env1, env2] = unify_parms ctxt fixed  | 
|
683  | 
[tenv1 |> Symtab.dest |> map (apsnd SOME),  | 
|
684  | 
tenv2 |> Symtab.dest |> map (apsnd SOME)]  | 
|
685  | 
in  | 
|
686  | 
Symtab.merge (op =) (Symtab.map (Element.instT_type env1) tenv1,  | 
|
687  | 
Symtab.map (Element.instT_type env2) tenv2)  | 
|
688  | 
end;  | 
|
689  | 
||
690  | 
fun merge_syn expr syn1 syn2 =  | 
|
| 19932 | 691  | 
Symtab.merge (op = : mixfix * mixfix -> bool) (syn1, syn2)  | 
| 
23655
 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 
wenzelm 
parents: 
23591 
diff
changeset
 | 
692  | 
handle Symtab.DUP x => err_in_expr ctxt  | 
| 
 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 
wenzelm 
parents: 
23591 
diff
changeset
 | 
693  | 
          ("Conflicting syntax for parameter: " ^ quote x) expr;
 | 
| 
20366
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
694  | 
|
| 19783 | 695  | 
fun params_of (expr as Locale name) =  | 
696  | 
let  | 
|
| 22573 | 697  | 
            val {imports, params, ...} = the_locale thy name;
 | 
| 19783 | 698  | 
val parms = map (fst o fst) params;  | 
| 22573 | 699  | 
val (parms', types', syn') = params_of imports;  | 
| 19783 | 700  | 
val all_parms = merge_lists parms' parms;  | 
701  | 
val all_types = merge_tenvs [] types' (params |> map fst |> Symtab.make);  | 
|
702  | 
val all_syn = merge_syn expr syn' (params |> map (apfst fst) |> Symtab.make);  | 
|
703  | 
in (all_parms, all_types, all_syn) end  | 
|
704  | 
| params_of (expr as Rename (e, xs)) =  | 
|
705  | 
let  | 
|
706  | 
val (parms', types', syn') = params_of e;  | 
|
707  | 
val ren = renaming xs parms';  | 
|
708  | 
(* renaming may reduce number of parameters *)  | 
|
709  | 
val new_parms = map (Element.rename ren) parms' |> distinct (op =);  | 
|
710  | 
val ren_syn = syn' |> Symtab.dest |> map (Element.rename_var ren);  | 
|
711  | 
val new_syn = fold (Symtab.insert (op =)) ren_syn Symtab.empty  | 
|
712  | 
handle Symtab.DUP x =>  | 
|
713  | 
                  err_in_expr ctxt ("Conflicting syntax for parameter: " ^ quote x) expr;
 | 
|
| 22700 | 714  | 
val syn_types = map (apsnd (fn mx =>  | 
715  | 
SOME (Type.freeze_type (#1 (TypeInfer.paramify_dummies (Syntax.mixfixT mx) 0)))))  | 
|
716  | 
(Symtab.dest new_syn);  | 
|
| 19783 | 717  | 
val ren_types = types' |> Symtab.dest |> map (apfst (Element.rename ren));  | 
| 
20366
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
718  | 
val (env :: _) = unify_parms ctxt []  | 
| 19783 | 719  | 
((ren_types |> map (apsnd SOME)) :: map single syn_types);  | 
720  | 
val new_types = fold (Symtab.insert (op =))  | 
|
721  | 
(map (apsnd (Element.instT_type env)) ren_types) Symtab.empty;  | 
|
722  | 
in (new_parms, new_types, new_syn) end  | 
|
723  | 
| params_of (Merge es) =  | 
|
724  | 
fold (fn e => fn (parms, types, syn) =>  | 
|
725  | 
let  | 
|
726  | 
val (parms', types', syn') = params_of e  | 
|
727  | 
in  | 
|
728  | 
(merge_lists parms parms', merge_tenvs [] types types',  | 
|
729  | 
merge_syn e syn syn')  | 
|
730  | 
end)  | 
|
731  | 
es ([], Symtab.empty, Symtab.empty)  | 
|
732  | 
||
733  | 
val (parms, types, syn) = params_of expr;  | 
|
734  | 
in  | 
|
735  | 
(merge_lists prev_parms parms, merge_tenvs fixed_params prev_types types,  | 
|
736  | 
merge_syn expr prev_syn syn)  | 
|
737  | 
end;  | 
|
738  | 
||
739  | 
fun make_params_ids params = [(("", params), ([], Assumed []))];
 | 
|
740  | 
fun make_raw_params_elemss (params, tenv, syn) =  | 
|
741  | 
    [((("", map (fn p => (p, Symtab.lookup tenv p)) params), Assumed []),
 | 
|
742  | 
Int [Fixes (map (fn p =>  | 
|
743  | 
(p, Symtab.lookup tenv p, Symtab.lookup syn p |> the)) params)])];  | 
|
744  | 
||
745  | 
||
| 15596 | 746  | 
(* flatten_expr:  | 
747  | 
Extend list of identifiers by those new in locale expression expr.  | 
|
748  | 
Compute corresponding list of lists of locale elements (one entry per  | 
|
749  | 
identifier).  | 
|
750  | 
||
751  | 
Identifiers represent locale fragments and are in an extended form:  | 
|
752  | 
((name, ps), (ax_ps, axs))  | 
|
753  | 
(name, ps) is the locale name with all its parameters.  | 
|
754  | 
(ax_ps, axs) is the locale axioms with its parameters;  | 
|
755  | 
axs are always taken from the top level of the locale hierarchy,  | 
|
756  | 
hence axioms may contain additional parameters from later fragments:  | 
|
757  | 
ps subset of ax_ps. axs is either singleton or empty.  | 
|
758  | 
||
759  | 
Elements are enriched by identifier-like information:  | 
|
760  | 
(((name, ax_ps), axs), elems)  | 
|
761  | 
The parameters in ax_ps are the axiom parameters, but enriched by type  | 
|
762  | 
info: now each entry is a pair of string and typ option. Axioms are  | 
|
763  | 
type-instantiated.  | 
|
764  | 
||
765  | 
*)  | 
|
766  | 
||
| 
16102
 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 
ballarin 
parents: 
16028 
diff
changeset
 | 
767  | 
fun flatten_expr ctxt ((prev_idents, prev_syntax), expr) =  | 
| 12014 | 768  | 
let  | 
769  | 
val thy = ProofContext.theory_of ctxt;  | 
|
| 12263 | 770  | 
|
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
771  | 
fun rename_parms top ren ((name, ps), (parms, mode)) =  | 
| 19783 | 772  | 
((name, map (Element.rename ren) ps),  | 
773  | 
if top  | 
|
774  | 
then (map (Element.rename ren) parms,  | 
|
| 
21483
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
775  | 
map_mode (map (Element.morph_witness (Element.rename_morphism ren))) mode)  | 
| 19783 | 776  | 
else (parms, mode));  | 
| 12263 | 777  | 
|
| 
20167
 
fe5fd4fd4114
Strict dfs traversal of imported and registered identifiers.
 
ballarin 
parents: 
20092 
diff
changeset
 | 
778  | 
(* add (name, ps) and its registrations, recursively; adjust hyps of witnesses *)  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
779  | 
|
| 
20167
 
fe5fd4fd4114
Strict dfs traversal of imported and registered identifiers.
 
ballarin 
parents: 
20092 
diff
changeset
 | 
780  | 
fun add_with_regs ((name, pTs), mode) (wits, ids, visited) =  | 
| 
 
fe5fd4fd4114
Strict dfs traversal of imported and registered identifiers.
 
ballarin 
parents: 
20092 
diff
changeset
 | 
781  | 
if member (fn (a, (b, _)) => a = b) visited (name, map #1 pTs)  | 
| 
 
fe5fd4fd4114
Strict dfs traversal of imported and registered identifiers.
 
ballarin 
parents: 
20092 
diff
changeset
 | 
782  | 
then (wits, ids, visited)  | 
| 
 
fe5fd4fd4114
Strict dfs traversal of imported and registered identifiers.
 
ballarin 
parents: 
20092 
diff
changeset
 | 
783  | 
else  | 
| 
20366
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
784  | 
let  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
785  | 
            val {params, regs, ...} = the_locale thy name;
 | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
786  | 
val pTs' = map #1 params;  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
787  | 
val ren = map #1 pTs' ~~ map (fn (x, _) => (x, NONE)) pTs;  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
788  | 
(* dummy syntax, since required by rename *)  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
789  | 
val pTs'' = map (fn ((p, _), (_, T)) => (p, T)) (pTs ~~ pTs');  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
790  | 
val [env] = unify_parms ctxt pTs [map (apsnd SOME) pTs''];  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
791  | 
(* propagate parameter types, to keep them consistent *)  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
792  | 
val regs' = map (fn ((name, ps), wits) =>  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
793  | 
((name, map (Element.rename ren) ps),  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
794  | 
map (Element.transfer_witness thy) wits)) regs;  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
795  | 
val new_regs = regs';  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
796  | 
val new_ids = map fst new_regs;  | 
| 
21483
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
797  | 
val new_idTs =  | 
| 
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
798  | 
map (apsnd (map (fn p => (p, (the o AList.lookup (op =) pTs) p)))) new_ids;  | 
| 
17096
 
8327b71282ce
Improved generation of witnesses in interpretation.
 
ballarin 
parents: 
17033 
diff
changeset
 | 
799  | 
|
| 
20366
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
800  | 
val new_wits = new_regs |> map (#2 #> map  | 
| 
21483
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
801  | 
(Element.morph_witness  | 
| 
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
802  | 
(Element.instT_morphism thy env $>  | 
| 
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
803  | 
Element.rename_morphism ren $>  | 
| 
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
804  | 
Element.satisfy_morphism wits)));  | 
| 
20366
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
805  | 
val new_ids' = map (fn (id, wits) =>  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
806  | 
(id, ([], Derived wits))) (new_ids ~~ new_wits);  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
807  | 
val new_idTs' = map (fn ((n, pTs), (_, ([], mode))) =>  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
808  | 
((n, pTs), mode)) (new_idTs ~~ new_ids');  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
809  | 
val new_id = ((name, map #1 pTs), ([], mode));  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
810  | 
val (wits', ids', visited') = fold add_with_regs new_idTs'  | 
| 
20167
 
fe5fd4fd4114
Strict dfs traversal of imported and registered identifiers.
 
ballarin 
parents: 
20092 
diff
changeset
 | 
811  | 
(wits @ flat new_wits, ids, visited @ [new_id]);  | 
| 
20366
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
812  | 
in  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
813  | 
(wits', ids' @ [new_id], visited')  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
814  | 
end;  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
815  | 
|
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
816  | 
(* distribute top-level axioms over assumed ids *)  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
817  | 
|
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
818  | 
fun axiomify all_ps ((name, parms), (_, Assumed _)) axioms =  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
819  | 
let  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
820  | 
          val {elems, ...} = the_locale thy name;
 | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19423 
diff
changeset
 | 
821  | 
val ts = maps  | 
| 
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19423 
diff
changeset
 | 
822  | 
(fn (Assumes asms, _) => maps (map #1 o #2) asms  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
823  | 
| _ => [])  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19423 
diff
changeset
 | 
824  | 
elems;  | 
| 
19018
 
88b4979193d8
added abbreviations: activated by init, no expressions yet;
 
wenzelm 
parents: 
18964 
diff
changeset
 | 
825  | 
val (axs1, axs2) = chop (length ts) axioms;  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
826  | 
in (((name, parms), (all_ps, Assumed axs1)), axs2) end  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
827  | 
| axiomify all_ps (id, (_, Derived ths)) axioms =  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
828  | 
((id, (all_ps, Derived ths)), axioms);  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
829  | 
|
| 
17096
 
8327b71282ce
Improved generation of witnesses in interpretation.
 
ballarin 
parents: 
17033 
diff
changeset
 | 
830  | 
(* identifiers of an expression *)  | 
| 
 
8327b71282ce
Improved generation of witnesses in interpretation.
 
ballarin 
parents: 
17033 
diff
changeset
 | 
831  | 
|
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
832  | 
fun identify top (Locale name) =  | 
| 15596 | 833  | 
(* CB: ids_ax is a list of tuples of the form ((name, ps), axs),  | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
834  | 
where name is a locale name, ps a list of parameter names and axs  | 
| 
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
835  | 
a list of axioms relating to the identifier, axs is empty unless  | 
| 
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
836  | 
identify at top level (top = true);  | 
| 
14215
 
ebf291f3b449
Improvements to Isar/Locales:  premises generated by "includes" elements
 
ballarin 
parents: 
13629 
diff
changeset
 | 
837  | 
parms is accumulated list of parameters *)  | 
| 12289 | 838  | 
let  | 
| 22573 | 839  | 
            val {axiom, imports, params, ...} = the_locale thy name;
 | 
| 19278 | 840  | 
val ps = map (#1 o #1) params;  | 
| 22573 | 841  | 
val (ids', parms') = identify false imports;  | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
842  | 
(* acyclic import dependencies *)  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
843  | 
|
| 
20167
 
fe5fd4fd4114
Strict dfs traversal of imported and registered identifiers.
 
ballarin 
parents: 
20092 
diff
changeset
 | 
844  | 
val (_, ids'', _) = add_with_regs ((name, map #1 params), Assumed []) ([], ids', ids');  | 
| 
 
fe5fd4fd4114
Strict dfs traversal of imported and registered identifiers.
 
ballarin 
parents: 
20092 
diff
changeset
 | 
845  | 
val ids_ax = if top then fst (fold_map (axiomify ps) ids'' axiom) else ids'';  | 
| 
20035
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
846  | 
in (ids_ax, merge_lists parms' ps) end  | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
847  | 
| identify top (Rename (e, xs)) =  | 
| 12273 | 848  | 
let  | 
| 
20035
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
849  | 
val (ids', parms') = identify top e;  | 
| 12839 | 850  | 
val ren = renaming xs parms'  | 
| 18678 | 851  | 
handle ERROR msg => err_in_locale' ctxt msg ids';  | 
| 
17096
 
8327b71282ce
Improved generation of witnesses in interpretation.
 
ballarin 
parents: 
17033 
diff
changeset
 | 
852  | 
|
| 
19061
 
ffbbac0261c9
removed distinct, renamed gen_distinct to distinct;
 
wenzelm 
parents: 
19025 
diff
changeset
 | 
853  | 
val ids'' = distinct (eq_fst (op =)) (map (rename_parms top ren) ids');  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19423 
diff
changeset
 | 
854  | 
val parms'' = distinct (op =) (maps (#2 o #1) ids'');  | 
| 
20035
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
855  | 
in (ids'', parms'') end  | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
856  | 
| identify top (Merge es) =  | 
| 
20035
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
857  | 
fold (fn e => fn (ids, parms) =>  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
858  | 
let  | 
| 
20035
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
859  | 
val (ids', parms') = identify top e  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
860  | 
in  | 
| 
20035
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
861  | 
(merge_alists ids ids', merge_lists parms parms')  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
862  | 
end)  | 
| 
20035
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
863  | 
es ([], []);  | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
864  | 
|
| 
20035
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
865  | 
fun inst_wit all_params (t, th) = let  | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
866  | 
         val {hyps, prop, ...} = Thm.rep_thm th;
 | 
| 16861 | 867  | 
val ps = map (apsnd SOME) (fold Term.add_frees (prop :: hyps) []);  | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
868  | 
val [env] = unify_parms ctxt all_params [ps];  | 
| 18137 | 869  | 
val t' = Element.instT_term env t;  | 
870  | 
val th' = Element.instT_thm thy env th;  | 
|
| 
18123
 
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
 
wenzelm 
parents: 
18038 
diff
changeset
 | 
871  | 
in (t', th') end;  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
872  | 
|
| 
20035
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
873  | 
fun eval all_params tenv syn ((name, params), (locale_params, mode)) =  | 
| 
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
874  | 
let  | 
| 
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
875  | 
        val {params = ps_mx, elems = elems_stamped, ...} = the_locale thy name;
 | 
| 
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
876  | 
val elems = map fst elems_stamped;  | 
| 
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
877  | 
val ps = map fst ps_mx;  | 
| 
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
878  | 
fun lookup_syn x = (case Symtab.lookup syn x of SOME Structure => NONE | opt => opt);  | 
| 
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
879  | 
val locale_params' = map (fn p => (p, Symtab.lookup tenv p |> the)) locale_params;  | 
| 
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
880  | 
val mode' = map_mode (map (Element.map_witness (inst_wit all_params))) mode;  | 
| 
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
881  | 
val ren = map fst ps ~~ map (fn p => (p, lookup_syn p)) params;  | 
| 
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
882  | 
val [env] = unify_parms ctxt all_params [map (apfst (Element.rename ren) o apsnd SOME) ps];  | 
| 
21499
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
883  | 
val elem_morphism =  | 
| 
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
884  | 
Element.rename_morphism ren $>  | 
| 
21665
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
885  | 
Morphism.name_morphism (NameSpace.qualified (params_prefix params)) $>  | 
| 
21499
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
886  | 
Element.instT_morphism thy env;  | 
| 
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
887  | 
val elems' = map (Element.morph_ctxt elem_morphism) elems;  | 
| 
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
888  | 
in (((name, map (apsnd SOME) locale_params'), mode'), elems') end;  | 
| 
20035
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
889  | 
|
| 
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
890  | 
(* parameters, their types and syntax *)  | 
| 
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
891  | 
val (all_params', tenv, syn) = params_of_expr ctxt [] expr ([], Symtab.empty, Symtab.empty);  | 
| 
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
892  | 
val all_params = map (fn p => (p, Symtab.lookup tenv p |> the)) all_params';  | 
| 
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
893  | 
(* compute identifiers and syntax, merge with previous ones *)  | 
| 
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
894  | 
val (ids, _) = identify true expr;  | 
| 
20951
 
868120282837
gen_rem(s) abandoned in favour of remove / subtract
 
haftmann 
parents: 
20911 
diff
changeset
 | 
895  | 
val idents = subtract (eq_fst (op =)) prev_idents ids;  | 
| 
20035
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
896  | 
val syntax = merge_syntax ctxt ids (syn, prev_syntax);  | 
| 
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
897  | 
(* type-instantiate elements *)  | 
| 
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
898  | 
val final_elemss = map (eval all_params tenv syntax) idents;  | 
| 
16102
 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 
ballarin 
parents: 
16028 
diff
changeset
 | 
899  | 
in ((prev_idents @ idents, syntax), final_elemss) end;  | 
| 12046 | 900  | 
|
| 12510 | 901  | 
end;  | 
902  | 
||
| 12070 | 903  | 
|
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
904  | 
(* activate elements *)  | 
| 12273 | 905  | 
|
| 12510 | 906  | 
local  | 
907  | 
||
| 21686 | 908  | 
fun axioms_export axs _ As =  | 
909  | 
(Element.satisfy_thm axs #> Drule.implies_intr_list (Library.drop (length axs, As)), fn t => t);  | 
|
| 12263 | 910  | 
|
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
911  | 
|
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
912  | 
(* NB: derived ids contain only facts at this stage *)  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
913  | 
|
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
914  | 
fun activate_elem _ _ ((ctxt, mode), Fixes fixes) =  | 
| 18671 | 915  | 
((ctxt |> ProofContext.add_fixes_i fixes |> snd, mode), [])  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
916  | 
| activate_elem _ _ ((ctxt, mode), Constrains _) =  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
917  | 
((ctxt, mode), [])  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
918  | 
| activate_elem ax_in_ctxt _ ((ctxt, Assumed axs), Assumes asms) =  | 
| 
13399
 
c136276dc847
support locale ``views'' (for cumulative predicates);
 
wenzelm 
parents: 
13394 
diff
changeset
 | 
919  | 
let  | 
| 18728 | 920  | 
val asms' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) asms;  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19423 
diff
changeset
 | 
921  | 
val ts = maps (map #1 o #2) asms';  | 
| 
19018
 
88b4979193d8
added abbreviations: activated by init, no expressions yet;
 
wenzelm 
parents: 
18964 
diff
changeset
 | 
922  | 
val (ps, qs) = chop (length ts) axs;  | 
| 17856 | 923  | 
val (_, ctxt') =  | 
| 
21370
 
d9dd7b4e5e69
replaced Variable.fix_frees by Variable.auto_fixes (depends on body mode);
 
wenzelm 
parents: 
21361 
diff
changeset
 | 
924  | 
ctxt |> fold Variable.auto_fixes ts  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
925  | 
|> ProofContext.add_assms_i (axioms_export (if ax_in_ctxt then ps else [])) asms';  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
926  | 
in ((ctxt', Assumed qs), []) end  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
927  | 
| activate_elem _ _ ((ctxt, Derived ths), Assumes asms) =  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
928  | 
((ctxt, Derived ths), [])  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
929  | 
| activate_elem _ _ ((ctxt, Assumed axs), Defines defs) =  | 
| 15596 | 930  | 
let  | 
| 18728 | 931  | 
val defs' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) defs;  | 
| 19732 | 932  | 
val asms = defs' |> map (fn ((name, atts), (t, ps)) =>  | 
933  | 
let val ((c, _), t') = LocalDefs.cert_def ctxt t  | 
|
| 20872 | 934  | 
in (t', ((Thm.def_name_optional c name, atts), [(t', ps)])) end);  | 
| 17856 | 935  | 
val (_, ctxt') =  | 
| 
21370
 
d9dd7b4e5e69
replaced Variable.fix_frees by Variable.auto_fixes (depends on body mode);
 
wenzelm 
parents: 
21361 
diff
changeset
 | 
936  | 
ctxt |> fold (Variable.auto_fixes o #1) asms  | 
| 19732 | 937  | 
|> ProofContext.add_assms_i LocalDefs.def_export (map #2 asms);  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
938  | 
in ((ctxt', Assumed axs), []) end  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
939  | 
| activate_elem _ _ ((ctxt, Derived ths), Defines defs) =  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
940  | 
((ctxt, Derived ths), [])  | 
| 21441 | 941  | 
| activate_elem _ is_ext ((ctxt, mode), Notes (kind, facts)) =  | 
| 15596 | 942  | 
let  | 
| 18728 | 943  | 
val facts' = Attrib.map_facts (Attrib.attribute_i (ProofContext.theory_of ctxt)) facts;  | 
| 21441 | 944  | 
val (res, ctxt') = ctxt |> ProofContext.note_thmss_i kind facts';  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
945  | 
in ((ctxt', mode), if is_ext then res else []) end;  | 
| 12502 | 946  | 
|
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
947  | 
fun activate_elems ax_in_ctxt (((name, ps), mode), elems) ctxt =  | 
| 17033 | 948  | 
let  | 
| 
18123
 
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
 
wenzelm 
parents: 
18038 
diff
changeset
 | 
949  | 
val thy = ProofContext.theory_of ctxt;  | 
| 17033 | 950  | 
val ((ctxt', _), res) =  | 
| 21441 | 951  | 
foldl_map (activate_elem ax_in_ctxt (name = ""))  | 
952  | 
((ProofContext.qualified_names ctxt, mode), elems)  | 
|
953  | 
handle ERROR msg => err_in_locale ctxt msg [(name, map fst ps)];  | 
|
| 15696 | 954  | 
val ctxt'' = if name = "" then ctxt'  | 
955  | 
else let  | 
|
956  | 
val ps' = map (fn (n, SOME T) => Free (n, T)) ps;  | 
|
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
957  | 
in if test_local_registration ctxt' (name, ps') then ctxt'  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
958  | 
else let  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
959  | 
val ctxt'' = put_local_registration (name, ps') ((true, ""), []) ctxt'  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
960  | 
in case mode of  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
961  | 
Assumed axs =>  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
962  | 
fold (add_local_witness (name, ps') o  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
963  | 
Element.assume_witness thy o Element.witness_prop) axs ctxt''  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
964  | 
| Derived ths => fold (add_local_witness (name, ps')) ths ctxt''  | 
| 
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
965  | 
end  | 
| 15696 | 966  | 
end  | 
| 16144 | 967  | 
in (ProofContext.restore_naming ctxt ctxt'', res) end;  | 
| 
13399
 
c136276dc847
support locale ``views'' (for cumulative predicates);
 
wenzelm 
parents: 
13394 
diff
changeset
 | 
968  | 
|
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
969  | 
fun activate_elemss ax_in_ctxt prep_facts =  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
970  | 
fold_map (fn (((name, ps), mode), raw_elems) => fn ctxt =>  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
971  | 
let  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
972  | 
val elems = map (prep_facts ctxt) raw_elems;  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19423 
diff
changeset
 | 
973  | 
val (ctxt', res) = apsnd flat  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
974  | 
(activate_elems ax_in_ctxt (((name, ps), mode), elems) ctxt);  | 
| 21530 | 975  | 
val elems' = elems |> map (Element.map_ctxt_attrib Args.closure);  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
976  | 
in (((((name, ps), mode), elems'), res), ctxt') end);  | 
| 
12834
 
e5bec3268932
added locale_facts(_i) interface (useful for simple ML proof tools);
 
wenzelm 
parents: 
12806 
diff
changeset
 | 
977  | 
|
| 12546 | 978  | 
in  | 
979  | 
||
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
980  | 
(* CB: activate_facts prep_facts (ctxt, elemss),  | 
| 
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
981  | 
where elemss is a list of pairs consisting of identifiers and  | 
| 
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
982  | 
context elements, extends ctxt by the context elements yielding  | 
| 
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
983  | 
ctxt' and returns (ctxt', (elemss', facts)).  | 
| 
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
984  | 
Identifiers in the argument are of the form ((name, ps), axs) and  | 
| 
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
985  | 
assumptions use the axioms in the identifiers to set up exporters  | 
| 
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
986  | 
in ctxt'. elemss' does not contain identifiers and is obtained  | 
| 
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
987  | 
from elemss and the intermediate context with prep_facts.  | 
| 15703 | 988  | 
If read_facts or cert_facts is used for prep_facts, these also remove  | 
| 
14508
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
989  | 
the internal/external markers from elemss. *)  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
990  | 
|
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
991  | 
fun activate_facts ax_in_ctxt prep_facts (ctxt, args) =  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
992  | 
let val ((elemss, factss), ctxt') = activate_elemss ax_in_ctxt prep_facts args ctxt |>> split_list  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19423 
diff
changeset
 | 
993  | 
in (ctxt', (elemss, flat factss)) end;  | 
| 15703 | 994  | 
|
| 12510 | 995  | 
end;  | 
996  | 
||
| 12307 | 997  | 
|
| 15696 | 998  | 
|
| 18137 | 999  | 
(** prepare locale elements **)  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1000  | 
|
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1001  | 
(* expressions *)  | 
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1002  | 
|
| 16458 | 1003  | 
fun intern_expr thy (Locale xname) = Locale (intern thy xname)  | 
1004  | 
| intern_expr thy (Merge exprs) = Merge (map (intern_expr thy) exprs)  | 
|
1005  | 
| intern_expr thy (Rename (expr, xs)) = Rename (intern_expr thy expr, xs);  | 
|
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1006  | 
|
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1007  | 
|
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1008  | 
(* propositions and bindings *)  | 
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1009  | 
|
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1010  | 
(* flatten (ctxt, prep_expr) ((ids, syn), expr)  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1011  | 
normalises expr (which is either a locale  | 
| 
14508
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1012  | 
expression or a single context element) wrt.  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1013  | 
to the list ids of already accumulated identifiers.  | 
| 19783 | 1014  | 
It returns ((ids', syn'), elemss) where ids' is an extension of ids  | 
| 
14508
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1015  | 
with identifiers generated for expr, and elemss is the list of  | 
| 
16102
 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 
ballarin 
parents: 
16028 
diff
changeset
 | 
1016  | 
context elements generated from expr.  | 
| 
 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 
ballarin 
parents: 
16028 
diff
changeset
 | 
1017  | 
syn and syn' are symtabs mapping parameter names to their syntax. syn'  | 
| 
 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 
ballarin 
parents: 
16028 
diff
changeset
 | 
1018  | 
is an extension of syn.  | 
| 
 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 
ballarin 
parents: 
16028 
diff
changeset
 | 
1019  | 
For details, see flatten_expr.  | 
| 
 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 
ballarin 
parents: 
16028 
diff
changeset
 | 
1020  | 
|
| 15596 | 1021  | 
Additionally, for a locale expression, the elems are grouped into a single  | 
1022  | 
Int; individual context elements are marked Ext. In this case, the  | 
|
1023  | 
identifier-like information of the element is as follows:  | 
|
1024  | 
   - for Fixes: (("", ps), []) where the ps have type info NONE
 | 
|
1025  | 
   - for other elements: (("", []), []).
 | 
|
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1026  | 
The implementation of activate_facts relies on identifier names being  | 
| 
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1027  | 
empty strings for external elements.  | 
| 15596 | 1028  | 
*)  | 
| 
14508
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1029  | 
|
| 
16102
 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 
ballarin 
parents: 
16028 
diff
changeset
 | 
1030  | 
fun flatten (ctxt, _) ((ids, syn), Elem (Fixes fixes)) = let  | 
| 18137 | 1031  | 
        val ids' = ids @ [(("", map #1 fixes), ([], Assumed []))]
 | 
| 
16102
 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 
ballarin 
parents: 
16028 
diff
changeset
 | 
1032  | 
in  | 
| 18137 | 1033  | 
((ids',  | 
1034  | 
merge_syntax ctxt ids'  | 
|
1035  | 
(syn, Symtab.make (map (fn fx => (#1 fx, #3 fx)) fixes))  | 
|
| 
23655
 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 
wenzelm 
parents: 
23591 
diff
changeset
 | 
1036  | 
handle Symtab.DUP x => err_in_locale ctxt  | 
| 
 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 
wenzelm 
parents: 
23591 
diff
changeset
 | 
1037  | 
             ("Conflicting syntax for parameter: " ^ quote x)
 | 
| 
16102
 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 
ballarin 
parents: 
16028 
diff
changeset
 | 
1038  | 
(map #1 ids')),  | 
| 18137 | 1039  | 
         [((("", map (rpair NONE o #1) fixes), Assumed []), Ext (Fixes fixes))])
 | 
| 
16102
 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 
ballarin 
parents: 
16028 
diff
changeset
 | 
1040  | 
end  | 
| 
 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 
ballarin 
parents: 
16028 
diff
changeset
 | 
1041  | 
| flatten _ ((ids, syn), Elem elem) =  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1042  | 
      ((ids @ [(("", []), ([], Assumed []))], syn), [((("", []), Assumed []), Ext elem)])
 | 
| 
16102
 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 
ballarin 
parents: 
16028 
diff
changeset
 | 
1043  | 
| flatten (ctxt, prep_expr) ((ids, syn), Expr expr) =  | 
| 
 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 
ballarin 
parents: 
16028 
diff
changeset
 | 
1044  | 
apsnd (map (apsnd Int)) (flatten_expr ctxt ((ids, syn), prep_expr expr));  | 
| 
14508
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1045  | 
|
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1046  | 
local  | 
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1047  | 
|
| 12839 | 1048  | 
local  | 
1049  | 
||
| 12727 | 1050  | 
fun declare_int_elem (ctxt, Fixes fixes) =  | 
| 18671 | 1051  | 
(ctxt |> ProofContext.add_fixes_i (map (fn (x, T, mx) =>  | 
1052  | 
(x, Option.map (Term.map_type_tfree (TypeInfer.param 0)) T, mx)) fixes) |> snd, [])  | 
|
| 12727 | 1053  | 
| declare_int_elem (ctxt, _) = (ctxt, []);  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1054  | 
|
| 18671 | 1055  | 
fun declare_ext_elem prep_vars (ctxt, Fixes fixes) =  | 
1056  | 
let val (vars, _) = prep_vars fixes ctxt  | 
|
1057  | 
in (ctxt |> ProofContext.add_fixes_i vars |> snd, []) end  | 
|
1058  | 
| declare_ext_elem prep_vars (ctxt, Constrains csts) =  | 
|
1059  | 
let val (_, ctxt') = prep_vars (map (fn (x, T) => (x, SOME T, NoSyn)) csts) ctxt  | 
|
1060  | 
in (ctxt', []) end  | 
|
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1061  | 
| declare_ext_elem _ (ctxt, Assumes asms) = (ctxt, map #2 asms)  | 
| 19585 | 1062  | 
| declare_ext_elem _ (ctxt, Defines defs) = (ctxt, map (fn (_, (t, ps)) => [(t, ps)]) defs)  | 
| 21441 | 1063  | 
| declare_ext_elem _ (ctxt, Notes _) = (ctxt, []);  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1064  | 
|
| 18671 | 1065  | 
fun declare_elems prep_vars (ctxt, (((name, ps), Assumed _), elems)) =  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1066  | 
let val (ctxt', propps) =  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1067  | 
(case elems of  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1068  | 
Int es => foldl_map declare_int_elem (ctxt, es)  | 
| 18671 | 1069  | 
| Ext e => foldl_map (declare_ext_elem prep_vars) (ctxt, [e]))  | 
| 18678 | 1070  | 
handle ERROR msg => err_in_locale ctxt msg [(name, map fst ps)]  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1071  | 
in (ctxt', propps) end  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1072  | 
| declare_elems _ (ctxt, ((_, Derived _), elems)) = (ctxt, []);  | 
| 12727 | 1073  | 
|
| 12839 | 1074  | 
in  | 
1075  | 
||
| 18671 | 1076  | 
fun declare_elemss prep_vars fixed_params raw_elemss ctxt =  | 
| 12727 | 1077  | 
let  | 
| 
14215
 
ebf291f3b449
Improvements to Isar/Locales:  premises generated by "includes" elements
 
ballarin 
parents: 
13629 
diff
changeset
 | 
1078  | 
(* CB: fix of type bug of goal in target with context elements.  | 
| 
 
ebf291f3b449
Improvements to Isar/Locales:  premises generated by "includes" elements
 
ballarin 
parents: 
13629 
diff
changeset
 | 
1079  | 
Parameters new in context elements must receive types that are  | 
| 
 
ebf291f3b449
Improvements to Isar/Locales:  premises generated by "includes" elements
 
ballarin 
parents: 
13629 
diff
changeset
 | 
1080  | 
distinct from types of parameters in target (fixed_params). *)  | 
| 
 
ebf291f3b449
Improvements to Isar/Locales:  premises generated by "includes" elements
 
ballarin 
parents: 
13629 
diff
changeset
 | 
1081  | 
val ctxt_with_fixed =  | 
| 
19900
 
21a99d88d925
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19873 
diff
changeset
 | 
1082  | 
fold Variable.declare_term (map Free fixed_params) ctxt;  | 
| 12727 | 1083  | 
val int_elemss =  | 
1084  | 
raw_elemss  | 
|
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19423 
diff
changeset
 | 
1085  | 
|> map_filter (fn (id, Int es) => SOME (id, es) | _ => NONE)  | 
| 
14215
 
ebf291f3b449
Improvements to Isar/Locales:  premises generated by "includes" elements
 
ballarin 
parents: 
13629 
diff
changeset
 | 
1086  | 
|> unify_elemss ctxt_with_fixed fixed_params;  | 
| 12727 | 1087  | 
val (_, raw_elemss') =  | 
1088  | 
foldl_map (fn ((_, es) :: elemss, (id, Int _)) => (elemss, (id, Int es)) | x => x)  | 
|
1089  | 
(int_elemss, raw_elemss);  | 
|
| 18671 | 1090  | 
in foldl_map (declare_elems prep_vars) (ctxt, raw_elemss') end;  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1091  | 
|
| 12839 | 1092  | 
end;  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1093  | 
|
| 12839 | 1094  | 
local  | 
1095  | 
||
1096  | 
val norm_term = Envir.beta_norm oo Term.subst_atomic;  | 
|
1097  | 
||
| 16458 | 1098  | 
fun abstract_thm thy eq =  | 
1099  | 
Thm.assume (Thm.cterm_of thy eq) |> Drule.gen_all |> Drule.abs_def;  | 
|
| 12502 | 1100  | 
|
| 18190 | 1101  | 
fun bind_def ctxt (name, ps) eq (xs, env, ths) =  | 
| 12839 | 1102  | 
let  | 
| 18831 | 1103  | 
val ((y, T), b) = LocalDefs.abs_def eq;  | 
| 13308 | 1104  | 
val b' = norm_term env b;  | 
| 16458 | 1105  | 
val th = abstract_thm (ProofContext.theory_of ctxt) eq;  | 
| 13308 | 1106  | 
fun err msg = err_in_locale ctxt (msg ^ ": " ^ quote y) [(name, map fst ps)];  | 
| 12839 | 1107  | 
in  | 
| 21962 | 1108  | 
exists (equal y o #1) xs andalso  | 
1109  | 
err "Attempt to define previously specified variable";  | 
|
1110  | 
exists (fn (Free (y', _), _) => y = y' | _ => false) env andalso  | 
|
1111  | 
err "Attempt to redefine variable";  | 
|
| 16861 | 1112  | 
(Term.add_frees b' xs, (Free (y, T), b') :: env, th :: ths)  | 
| 12839 | 1113  | 
end;  | 
| 12575 | 1114  | 
|
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1115  | 
|
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1116  | 
(* CB: for finish_elems (Int and Ext),  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1117  | 
extracts specification, only of assumed elements *)  | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1118  | 
|
| 18190 | 1119  | 
fun eval_text _ _ _ (Fixes _) text = text  | 
1120  | 
| eval_text _ _ _ (Constrains _) text = text  | 
|
1121  | 
| eval_text _ (_, Assumed _) is_ext (Assumes asms)  | 
|
1122  | 
(((exts, exts'), (ints, ints')), (xs, env, defs)) =  | 
|
| 13394 | 1123  | 
let  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19423 
diff
changeset
 | 
1124  | 
val ts = maps (map #1 o #2) asms;  | 
| 13394 | 1125  | 
val ts' = map (norm_term env) ts;  | 
1126  | 
val spec' =  | 
|
1127  | 
if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints'))  | 
|
1128  | 
else ((exts, exts'), (ints @ ts, ints' @ ts'));  | 
|
| 16861 | 1129  | 
in (spec', (fold Term.add_frees ts' xs, env, defs)) end  | 
| 18190 | 1130  | 
| eval_text _ (_, Derived _) _ (Assumes _) text = text  | 
1131  | 
| eval_text ctxt (id, Assumed _) _ (Defines defs) (spec, binds) =  | 
|
1132  | 
(spec, fold (bind_def ctxt id o #1 o #2) defs binds)  | 
|
1133  | 
| eval_text _ (_, Derived _) _ (Defines _) text = text  | 
|
1134  | 
| eval_text _ _ _ (Notes _) text = text;  | 
|
| 13308 | 1135  | 
|
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1136  | 
|
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1137  | 
(* for finish_elems (Int),  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1138  | 
remove redundant elements of derived identifiers,  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1139  | 
turn assumptions and definitions into facts,  | 
| 
21483
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
1140  | 
satisfy hypotheses of facts *)  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1141  | 
|
| 
17096
 
8327b71282ce
Improved generation of witnesses in interpretation.
 
ballarin 
parents: 
17033 
diff
changeset
 | 
1142  | 
fun finish_derived _ _ (Assumed _) (Fixes fixes) = SOME (Fixes fixes)  | 
| 
 
8327b71282ce
Improved generation of witnesses in interpretation.
 
ballarin 
parents: 
17033 
diff
changeset
 | 
1143  | 
| finish_derived _ _ (Assumed _) (Constrains csts) = SOME (Constrains csts)  | 
| 
 
8327b71282ce
Improved generation of witnesses in interpretation.
 
ballarin 
parents: 
17033 
diff
changeset
 | 
1144  | 
| finish_derived _ _ (Assumed _) (Assumes asms) = SOME (Assumes asms)  | 
| 
 
8327b71282ce
Improved generation of witnesses in interpretation.
 
ballarin 
parents: 
17033 
diff
changeset
 | 
1145  | 
| finish_derived _ _ (Assumed _) (Defines defs) = SOME (Defines defs)  | 
| 
 
8327b71282ce
Improved generation of witnesses in interpretation.
 
ballarin 
parents: 
17033 
diff
changeset
 | 
1146  | 
|
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1147  | 
| finish_derived _ _ (Derived _) (Fixes _) = NONE  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1148  | 
| finish_derived _ _ (Derived _) (Constrains _) = NONE  | 
| 
21483
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
1149  | 
| finish_derived sign satisfy (Derived _) (Assumes asms) = asms  | 
| 
17096
 
8327b71282ce
Improved generation of witnesses in interpretation.
 
ballarin 
parents: 
17033 
diff
changeset
 | 
1150  | 
|> map (apsnd (map (fn (a, _) => ([Thm.assume (cterm_of sign a)], []))))  | 
| 21441 | 1151  | 
|> pair Thm.assumptionK |> Notes  | 
| 
21483
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
1152  | 
|> Element.morph_ctxt satisfy |> SOME  | 
| 
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
1153  | 
| finish_derived sign satisfy (Derived _) (Defines defs) = defs  | 
| 
17096
 
8327b71282ce
Improved generation of witnesses in interpretation.
 
ballarin 
parents: 
17033 
diff
changeset
 | 
1154  | 
|> map (apsnd (fn (d, _) => [([Thm.assume (cterm_of sign d)], [])]))  | 
| 21441 | 1155  | 
|> pair Thm.definitionK |> Notes  | 
| 
21483
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
1156  | 
|> Element.morph_ctxt satisfy |> SOME  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1157  | 
|
| 
21483
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
1158  | 
| finish_derived _ satisfy _ (Notes facts) = Notes facts  | 
| 
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
1159  | 
|> Element.morph_ctxt satisfy |> SOME;  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1160  | 
|
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1161  | 
(* CB: for finish_elems (Ext) *)  | 
| 
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1162  | 
|
| 13308 | 1163  | 
fun closeup _ false elem = elem  | 
1164  | 
| closeup ctxt true elem =  | 
|
| 12839 | 1165  | 
let  | 
| 13308 | 1166  | 
fun close_frees t =  | 
| 
19900
 
21a99d88d925
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19873 
diff
changeset
 | 
1167  | 
let val frees = rev (filter_out (Variable.is_fixed ctxt o #1) (Term.add_frees t []))  | 
| 13308 | 1168  | 
in Term.list_all_free (frees, t) end;  | 
1169  | 
||
1170  | 
fun no_binds [] = []  | 
|
| 18678 | 1171  | 
| no_binds _ = error "Illegal term bindings in locale element";  | 
| 13308 | 1172  | 
in  | 
1173  | 
(case elem of  | 
|
1174  | 
Assumes asms => Assumes (asms |> map (fn (a, propps) =>  | 
|
| 19585 | 1175  | 
(a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps)))  | 
| 13308 | 1176  | 
| Defines defs => Defines (defs |> map (fn (a, (t, ps)) =>  | 
| 18831 | 1177  | 
(a, (close_frees (#2 (LocalDefs.cert_def ctxt t)), no_binds ps))))  | 
| 13308 | 1178  | 
| e => e)  | 
1179  | 
end;  | 
|
| 12839 | 1180  | 
|
| 12502 | 1181  | 
|
| 12839 | 1182  | 
fun finish_ext_elem parms _ (Fixes fixes, _) = Fixes (map (fn (x, _, mx) =>  | 
| 17271 | 1183  | 
(x, AList.lookup (op =) parms x, mx)) fixes)  | 
| 
18899
 
a8e913c93578
theorem(_in_locale): Element.statement, Obtain.statement;
 
wenzelm 
parents: 
18890 
diff
changeset
 | 
1184  | 
| finish_ext_elem parms _ (Constrains _, _) = Constrains []  | 
| 12839 | 1185  | 
| finish_ext_elem _ close (Assumes asms, propp) =  | 
1186  | 
close (Assumes (map #1 asms ~~ propp))  | 
|
1187  | 
| finish_ext_elem _ close (Defines defs, propp) =  | 
|
| 19585 | 1188  | 
close (Defines (map #1 defs ~~ map (fn [(t, ps)] => (t, ps)) propp))  | 
| 12839 | 1189  | 
| finish_ext_elem _ _ (Notes facts, _) = Notes facts;  | 
1190  | 
||
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1191  | 
|
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1192  | 
(* CB: finish_parms introduces type info from parms to identifiers *)  | 
| 15531 | 1193  | 
(* CB: only needed for types that have been NONE so far???  | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1194  | 
If so, which are these??? *)  | 
| 
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1195  | 
|
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1196  | 
fun finish_parms parms (((name, ps), mode), elems) =  | 
| 19932 | 1197  | 
(((name, map (fn (x, _) => (x, AList.lookup (op = : string * string -> bool) parms x)) ps), mode), elems);  | 
| 12839 | 1198  | 
|
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1199  | 
fun finish_elems ctxt parms _ ((text, wits), ((id, Int e), _)) =  | 
| 12839 | 1200  | 
let  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1201  | 
val [(id' as (_, mode), es)] = unify_elemss ctxt parms [(id, e)];  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1202  | 
val wits' = case mode of Assumed _ => wits | Derived ths => wits @ ths;  | 
| 18190 | 1203  | 
val text' = fold (eval_text ctxt id' false) es text;  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19423 
diff
changeset
 | 
1204  | 
val es' = map_filter  | 
| 
21483
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
1205  | 
(finish_derived (ProofContext.theory_of ctxt) (Element.satisfy_morphism wits') mode) es;  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1206  | 
in ((text', wits'), (id', map Int es')) end  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1207  | 
| finish_elems ctxt parms do_close ((text, wits), ((id, Ext e), [propp])) =  | 
| 13308 | 1208  | 
let  | 
1209  | 
val e' = finish_ext_elem parms (closeup ctxt do_close) (e, propp);  | 
|
| 18190 | 1210  | 
val text' = eval_text ctxt id true e' text;  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1211  | 
in ((text', wits), (id, [Ext e'])) end  | 
| 12839 | 1212  | 
|
1213  | 
in  | 
|
| 12510 | 1214  | 
|
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1215  | 
(* CB: only called by prep_elemss *)  | 
| 
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1216  | 
|
| 13375 | 1217  | 
fun finish_elemss ctxt parms do_close =  | 
1218  | 
foldl_map (apsnd (finish_parms parms) o finish_elems ctxt parms do_close);  | 
|
| 12839 | 1219  | 
|
1220  | 
end;  | 
|
1221  | 
||
| 
16736
 
1e792b32abef
Preparations for interpretation of locales in locales.
 
ballarin 
parents: 
16620 
diff
changeset
 | 
1222  | 
|
| 
19942
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1223  | 
(* Remove duplicate Defines elements: temporary workaround to fix Afp/Category. *)  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1224  | 
|
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1225  | 
fun defs_ord (defs1, defs2) =  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1226  | 
list_ord (fn ((_, (d1, _)), (_, (d2, _))) =>  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1227  | 
Term.fast_term_ord (d1, d2)) (defs1, defs2);  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1228  | 
structure Defstab =  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1229  | 
TableFun(type key = ((string * Attrib.src list) * (term * term list)) list  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1230  | 
val ord = defs_ord);  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1231  | 
|
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1232  | 
fun rem_dup_defs es ds =  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1233  | 
fold_map (fn e as (Defines defs) => (fn ds =>  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1234  | 
if Defstab.defined ds defs  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1235  | 
then (Defines [], ds)  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1236  | 
else (e, Defstab.update (defs, ()) ds))  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1237  | 
| e => (fn ds => (e, ds))) es ds;  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1238  | 
fun rem_dup_elemss (Int es) ds = apfst Int (rem_dup_defs es ds)  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1239  | 
| rem_dup_elemss (Ext e) ds = (Ext e, ds);  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1240  | 
fun rem_dup_defines raw_elemss =  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1241  | 
fold_map (fn (id as (_, (Assumed _)), es) => (fn ds =>  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1242  | 
apfst (pair id) (rem_dup_elemss es ds))  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1243  | 
| (id as (_, (Derived _)), es) => (fn ds =>  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1244  | 
((id, es), ds))) raw_elemss Defstab.empty |> #1;  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1245  | 
|
| 
16736
 
1e792b32abef
Preparations for interpretation of locales in locales.
 
ballarin 
parents: 
16620 
diff
changeset
 | 
1246  | 
(* CB: type inference and consistency checks for locales.  | 
| 
 
1e792b32abef
Preparations for interpretation of locales in locales.
 
ballarin 
parents: 
16620 
diff
changeset
 | 
1247  | 
|
| 
 
1e792b32abef
Preparations for interpretation of locales in locales.
 
ballarin 
parents: 
16620 
diff
changeset
 | 
1248  | 
Works by building a context (through declare_elemss), extracting the  | 
| 
 
1e792b32abef
Preparations for interpretation of locales in locales.
 
ballarin 
parents: 
16620 
diff
changeset
 | 
1249  | 
required information and adjusting the context elements (finish_elemss).  | 
| 
 
1e792b32abef
Preparations for interpretation of locales in locales.
 
ballarin 
parents: 
16620 
diff
changeset
 | 
1250  | 
Can also universally close free vars in assms and defs. This is only  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1251  | 
needed for Ext elements and controlled by parameter do_close.  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1252  | 
|
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1253  | 
Only elements of assumed identifiers are considered.  | 
| 
16736
 
1e792b32abef
Preparations for interpretation of locales in locales.
 
ballarin 
parents: 
16620 
diff
changeset
 | 
1254  | 
*)  | 
| 15127 | 1255  | 
|
| 18671 | 1256  | 
fun prep_elemss prep_vars prepp do_close context fixed_params raw_elemss raw_concl =  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1257  | 
let  | 
| 15127 | 1258  | 
(* CB: contexts computed in the course of this function are discarded.  | 
1259  | 
They are used for type inference and consistency checks only. *)  | 
|
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1260  | 
(* CB: fixed_params are the parameters (with types) of the target locale,  | 
| 
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1261  | 
empty list if there is no target. *)  | 
| 
14508
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1262  | 
(* CB: raw_elemss are list of pairs consisting of identifiers and  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1263  | 
context elements, the latter marked as internal or external. *)  | 
| 
19942
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1264  | 
val raw_elemss = rem_dup_defines raw_elemss;  | 
| 18671 | 1265  | 
val (raw_ctxt, raw_proppss) = declare_elemss prep_vars fixed_params raw_elemss context;  | 
| 
14508
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1266  | 
(* CB: raw_ctxt is context with additional fixed variables derived from  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1267  | 
the fixes elements in raw_elemss,  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1268  | 
raw_proppss contains assumptions and definitions from the  | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1269  | 
external elements in raw_elemss. *)  | 
| 
18550
 
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
 
haftmann 
parents: 
18546 
diff
changeset
 | 
1270  | 
fun prep_prop raw_propp (raw_ctxt, raw_concl) =  | 
| 
18450
 
e57731ba01dd
discontinued unflat in favour of burrow and burrow_split
 
haftmann 
parents: 
18421 
diff
changeset
 | 
1271  | 
let  | 
| 
 
e57731ba01dd
discontinued unflat in favour of burrow and burrow_split
 
haftmann 
parents: 
18421 
diff
changeset
 | 
1272  | 
(* CB: add type information from fixed_params to context (declare_term) *)  | 
| 
 
e57731ba01dd
discontinued unflat in favour of burrow and burrow_split
 
haftmann 
parents: 
18421 
diff
changeset
 | 
1273  | 
(* CB: process patterns (conclusion and external elements only) *)  | 
| 
 
e57731ba01dd
discontinued unflat in favour of burrow and burrow_split
 
haftmann 
parents: 
18421 
diff
changeset
 | 
1274  | 
val (ctxt, all_propp) =  | 
| 
19900
 
21a99d88d925
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19873 
diff
changeset
 | 
1275  | 
prepp (fold Variable.declare_term (map Free fixed_params) raw_ctxt, raw_concl @ raw_propp);  | 
| 
18450
 
e57731ba01dd
discontinued unflat in favour of burrow and burrow_split
 
haftmann 
parents: 
18421 
diff
changeset
 | 
1276  | 
(* CB: add type information from conclusion and external elements to context *)  | 
| 
19900
 
21a99d88d925
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19873 
diff
changeset
 | 
1277  | 
val ctxt = fold Variable.declare_term (maps (map fst) all_propp) ctxt;  | 
| 
18450
 
e57731ba01dd
discontinued unflat in favour of burrow and burrow_split
 
haftmann 
parents: 
18421 
diff
changeset
 | 
1278  | 
(* CB: resolve schematic variables (patterns) in conclusion and external elements. *)  | 
| 
 
e57731ba01dd
discontinued unflat in favour of burrow and burrow_split
 
haftmann 
parents: 
18421 
diff
changeset
 | 
1279  | 
val all_propp' = map2 (curry (op ~~))  | 
| 
 
e57731ba01dd
discontinued unflat in favour of burrow and burrow_split
 
haftmann 
parents: 
18421 
diff
changeset
 | 
1280  | 
(#1 (#2 (ProofContext.bind_propp_schematic_i (ctxt, all_propp)))) (map (map snd) all_propp);  | 
| 
19018
 
88b4979193d8
added abbreviations: activated by init, no expressions yet;
 
wenzelm 
parents: 
18964 
diff
changeset
 | 
1281  | 
val (concl, propp) = chop (length raw_concl) all_propp';  | 
| 
18550
 
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
 
haftmann 
parents: 
18546 
diff
changeset
 | 
1282  | 
in (propp, (ctxt, concl)) end  | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1283  | 
|
| 
18550
 
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
 
haftmann 
parents: 
18546 
diff
changeset
 | 
1284  | 
val (proppss, (ctxt, concl)) =  | 
| 
 
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
 
haftmann 
parents: 
18546 
diff
changeset
 | 
1285  | 
(fold_burrow o fold_burrow) prep_prop raw_proppss (raw_ctxt, raw_concl);  | 
| 12502 | 1286  | 
|
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1287  | 
(* CB: obtain all parameters from identifier part of raw_elemss *)  | 
| 
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1288  | 
val xs = map #1 (params_of' raw_elemss);  | 
| 12727 | 1289  | 
val typing = unify_frozen ctxt 0  | 
| 
19900
 
21a99d88d925
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19873 
diff
changeset
 | 
1290  | 
(map (Variable.default_type raw_ctxt) xs)  | 
| 
 
21a99d88d925
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19873 
diff
changeset
 | 
1291  | 
(map (Variable.default_type ctxt) xs);  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1292  | 
val parms = param_types (xs ~~ typing);  | 
| 
14508
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1293  | 
(* CB: parms are the parameters from raw_elemss, with correct typing. *)  | 
| 12273 | 1294  | 
|
| 
14508
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1295  | 
(* CB: extract information from assumes and defines elements  | 
| 
16169
 
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
 
ballarin 
parents: 
16144 
diff
changeset
 | 
1296  | 
(fixes, constrains and notes in raw_elemss don't have an effect on  | 
| 
 
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
 
ballarin 
parents: 
16144 
diff
changeset
 | 
1297  | 
text and elemss), compute final form of context elements. *)  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1298  | 
val ((text, _), elemss) = finish_elemss ctxt parms do_close  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1299  | 
((((([], []), ([], [])), ([], [], [])), []), raw_elemss ~~ proppss);  | 
| 
14508
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1300  | 
(* CB: text has the following structure:  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1301  | 
(((exts, exts'), (ints, ints')), (xs, env, defs))  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1302  | 
where  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1303  | 
exts: external assumptions (terms in external assumes elements)  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1304  | 
exts': dito, normalised wrt. env  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1305  | 
ints: internal assumptions (terms in internal assumes elements)  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1306  | 
ints': dito, normalised wrt. env  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1307  | 
xs: the free variables in exts' and ints' and rhss of definitions,  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1308  | 
this includes parameters except defined parameters  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1309  | 
env: list of term pairs encoding substitutions, where the first term  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1310  | 
is a free variable; substitutions represent defines elements and  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1311  | 
the rhs is normalised wrt. the previous env  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1312  | 
defs: theorems representing the substitutions from defines elements  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1313  | 
(thms are normalised wrt. env).  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1314  | 
elemss is an updated version of raw_elemss:  | 
| 
16169
 
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
 
ballarin 
parents: 
16144 
diff
changeset
 | 
1315  | 
- type info added to Fixes and modified in Constrains  | 
| 
14508
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1316  | 
- axiom and definition statement replaced by corresponding one  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1317  | 
from proppss in Assumes and Defines  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1318  | 
- Facts unchanged  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1319  | 
*)  | 
| 13308 | 1320  | 
in ((parms, elemss, concl), text) end;  | 
| 12502 | 1321  | 
|
1322  | 
in  | 
|
1323  | 
||
| 18671 | 1324  | 
fun read_elemss x = prep_elemss ProofContext.read_vars ProofContext.read_propp_schematic x;  | 
1325  | 
fun cert_elemss x = prep_elemss ProofContext.cert_vars ProofContext.cert_propp_schematic x;  | 
|
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1326  | 
|
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1327  | 
end;  | 
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1328  | 
|
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1329  | 
|
| 15703 | 1330  | 
(* facts and attributes *)  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1331  | 
|
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1332  | 
local  | 
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1333  | 
|
| 20965 | 1334  | 
fun check_name name =  | 
| 18678 | 1335  | 
  if NameSpace.is_qualified name then error ("Illegal qualified name: " ^ quote name)
 | 
| 15703 | 1336  | 
else name;  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1337  | 
|
| 
21499
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
1338  | 
fun prep_facts _ _ _ ctxt (Int elem) = elem  | 
| 
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
1339  | 
|> Element.morph_ctxt (Morphism.thm_morphism (Thm.transfer (ProofContext.theory_of ctxt)))  | 
| 20965 | 1340  | 
| prep_facts prep_name get intern ctxt (Ext elem) = elem |> Element.map_ctxt  | 
| 15703 | 1341  | 
     {var = I, typ = I, term = I,
 | 
| 18678 | 1342  | 
name = prep_name,  | 
| 15703 | 1343  | 
fact = get ctxt,  | 
| 16458 | 1344  | 
attrib = Args.assignable o intern (ProofContext.theory_of ctxt)};  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1345  | 
|
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1346  | 
in  | 
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1347  | 
|
| 20965 | 1348  | 
fun read_facts x = prep_facts check_name ProofContext.get_thms Attrib.intern_src x;  | 
1349  | 
fun cert_facts x = prep_facts I (K I) (K I) x;  | 
|
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1350  | 
|
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1351  | 
end;  | 
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1352  | 
|
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1353  | 
|
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1354  | 
(* Get the specification of a locale *)  | 
| 
18795
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1355  | 
|
| 19780 | 1356  | 
(*The global specification is made from the parameters and global  | 
1357  | 
assumptions, the local specification from the parameters and the  | 
|
1358  | 
local assumptions.*)  | 
|
| 
18795
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1359  | 
|
| 
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1360  | 
local  | 
| 
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1361  | 
|
| 
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1362  | 
fun gen_asms_of get thy name =  | 
| 
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1363  | 
let  | 
| 
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1364  | 
val ctxt = ProofContext.init thy;  | 
| 
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1365  | 
val (_, raw_elemss) = flatten (ctxt, I) (([], Symtab.empty), Expr (Locale name));  | 
| 
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1366  | 
val ((_, elemss, _), _) = read_elemss false ctxt [] raw_elemss [];  | 
| 
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1367  | 
in  | 
| 18890 | 1368  | 
elemss |> get  | 
| 19780 | 1369  | 
|> maps (fn (_, es) => map (fn Int e => e) es)  | 
1370  | 
|> maps (fn Assumes asms => asms | _ => [])  | 
|
| 
18795
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1371  | 
|> map (apsnd (map fst))  | 
| 
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1372  | 
end;  | 
| 
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1373  | 
|
| 
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1374  | 
in  | 
| 
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1375  | 
|
| 
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1376  | 
fun parameters_of thy name =  | 
| 19278 | 1377  | 
the_locale thy name |> #params;  | 
| 
18795
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1378  | 
|
| 19276 | 1379  | 
fun parameters_of_expr thy expr =  | 
1380  | 
let  | 
|
1381  | 
val ctxt = ProofContext.init thy;  | 
|
| 19783 | 1382  | 
val pts = params_of_expr ctxt [] (intern_expr thy expr)  | 
1383  | 
([], Symtab.empty, Symtab.empty);  | 
|
1384  | 
val raw_params_elemss = make_raw_params_elemss pts;  | 
|
| 19276 | 1385  | 
val ((_, syn), raw_elemss) = flatten (ctxt, intern_expr thy)  | 
1386  | 
(([], Symtab.empty), Expr expr);  | 
|
| 19783 | 1387  | 
val ((parms, _, _), _) =  | 
1388  | 
read_elemss false ctxt [] (raw_params_elemss @ raw_elemss) [];  | 
|
| 19276 | 1389  | 
in map (fn p as (n, _) => (p, Symtab.lookup syn n |> the)) parms end;  | 
1390  | 
||
| 
18795
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1391  | 
fun local_asms_of thy name =  | 
| 18890 | 1392  | 
gen_asms_of (single o Library.last_elem) thy name;  | 
| 
18795
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1393  | 
|
| 
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1394  | 
fun global_asms_of thy name =  | 
| 18890 | 1395  | 
gen_asms_of I thy name;  | 
| 
18795
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1396  | 
|
| 19780 | 1397  | 
end;  | 
| 
18795
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1398  | 
|
| 22523 | 1399  | 
fun intros thy =  | 
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
1400  | 
#intros o the o Symtab.lookup (#2 (LocalesData.get thy));  | 
| 22523 | 1401  | 
(*returns introduction rule for delta predicate and locale predicate  | 
1402  | 
as a pair of singleton lists*)  | 
|
1403  | 
||
| 
18795
 
303793f49b0f
Interface to access the specification of a named locale.
 
ballarin 
parents: 
18782 
diff
changeset
 | 
1404  | 
|
| 22573 | 1405  | 
(* full context statements: imports + elements + conclusion *)  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1406  | 
|
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1407  | 
local  | 
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1408  | 
|
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1409  | 
fun prep_context_statement prep_expr prep_elemss prep_facts  | 
| 22573 | 1410  | 
do_close fixed_params imports elements raw_concl context =  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1411  | 
let  | 
| 16458 | 1412  | 
val thy = ProofContext.theory_of context;  | 
| 13375 | 1413  | 
|
| 19783 | 1414  | 
val (import_params, import_tenv, import_syn) =  | 
| 22573 | 1415  | 
params_of_expr context fixed_params (prep_expr thy imports)  | 
| 19783 | 1416  | 
([], Symtab.empty, Symtab.empty);  | 
1417  | 
val includes = map_filter (fn Expr e => SOME e | Elem _ => NONE) elements;  | 
|
1418  | 
val (incl_params, incl_tenv, incl_syn) = fold (params_of_expr context fixed_params)  | 
|
1419  | 
(map (prep_expr thy) includes) (import_params, import_tenv, import_syn);  | 
|
1420  | 
||
1421  | 
val ((import_ids, _), raw_import_elemss) =  | 
|
| 22573 | 1422  | 
flatten (context, prep_expr thy) (([], Symtab.empty), Expr imports);  | 
| 
14215
 
ebf291f3b449
Improvements to Isar/Locales:  premises generated by "includes" elements
 
ballarin 
parents: 
13629 
diff
changeset
 | 
1423  | 
(* CB: normalise "includes" among elements *)  | 
| 16458 | 1424  | 
val ((ids, syn), raw_elemsss) = foldl_map (flatten (context, prep_expr thy))  | 
| 19783 | 1425  | 
((import_ids, incl_syn), elements);  | 
| 15696 | 1426  | 
|
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19423 
diff
changeset
 | 
1427  | 
val raw_elemss = flat raw_elemsss;  | 
| 
14508
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1428  | 
(* CB: raw_import_elemss @ raw_elemss is the normalised list of  | 
| 
 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 
ballarin 
parents: 
14446 
diff
changeset
 | 
1429  | 
context elements obtained from import and elements. *)  | 
| 19783 | 1430  | 
(* Now additional elements for parameters are inserted. *)  | 
1431  | 
val import_params_ids = make_params_ids import_params;  | 
|
1432  | 
val incl_params_ids =  | 
|
1433  | 
make_params_ids (incl_params \\ import_params);  | 
|
1434  | 
val raw_import_params_elemss =  | 
|
1435  | 
make_raw_params_elemss (import_params, incl_tenv, incl_syn);  | 
|
1436  | 
val raw_incl_params_elemss =  | 
|
1437  | 
make_raw_params_elemss (incl_params \\ import_params, incl_tenv, incl_syn);  | 
|
| 13375 | 1438  | 
val ((parms, all_elemss, concl), (spec, (_, _, defs))) = prep_elemss do_close  | 
| 19783 | 1439  | 
context fixed_params  | 
1440  | 
(raw_import_params_elemss @ raw_import_elemss @ raw_incl_params_elemss @ raw_elemss) raw_concl;  | 
|
1441  | 
||
| 15696 | 1442  | 
(* replace extended ids (for axioms) by ids *)  | 
| 19783 | 1443  | 
val (import_ids', incl_ids) = chop (length import_ids) ids;  | 
| 
20035
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
1444  | 
val all_ids = import_params_ids @ import_ids' @ incl_params_ids @ incl_ids;  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1445  | 
val all_elemss' = map (fn (((_, ps), _), (((n, ps'), mode), elems)) =>  | 
| 17485 | 1446  | 
(((n, map (fn p => (p, (the o AList.lookup (op =) ps') p)) ps), mode), elems))  | 
| 
20035
 
80c79876d2de
Internal restructuring: identify no longer computes syntax.
 
ballarin 
parents: 
20032 
diff
changeset
 | 
1447  | 
(all_ids ~~ all_elemss);  | 
| 19783 | 1448  | 
(* CB: all_elemss and parms contain the correct parameter types *)  | 
| 15696 | 1449  | 
|
| 19783 | 1450  | 
val (ps, qs) = chop (length raw_import_params_elemss + length raw_import_elemss) all_elemss';  | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1451  | 
val (import_ctxt, (import_elemss, _)) =  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1452  | 
activate_facts false prep_facts (context, ps);  | 
| 
14215
 
ebf291f3b449
Improvements to Isar/Locales:  premises generated by "includes" elements
 
ballarin 
parents: 
13629 
diff
changeset
 | 
1453  | 
|
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1454  | 
val (ctxt, (elemss, _)) =  | 
| 21441 | 1455  | 
activate_facts false prep_facts (ProofContext.set_stmt true import_ctxt, qs);  | 
| 
12834
 
e5bec3268932
added locale_facts(_i) interface (useful for simple ML proof tools);
 
wenzelm 
parents: 
12806 
diff
changeset
 | 
1456  | 
in  | 
| 19783 | 1457  | 
((((import_ctxt, import_elemss), (ctxt, elemss, syn)),  | 
| 
19991
 
0c9650664d47
Locales no longer generate views.  The following functions have changed
 
ballarin 
parents: 
19984 
diff
changeset
 | 
1458  | 
(parms, spec, defs)), concl)  | 
| 
12834
 
e5bec3268932
added locale_facts(_i) interface (useful for simple ML proof tools);
 
wenzelm 
parents: 
12806 
diff
changeset
 | 
1459  | 
end;  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1460  | 
|
| 18806 | 1461  | 
fun prep_statement prep_locale prep_ctxt raw_locale elems concl ctxt =  | 
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1462  | 
let  | 
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1463  | 
val thy = ProofContext.theory_of ctxt;  | 
| 16458 | 1464  | 
val locale = Option.map (prep_locale thy) raw_locale;  | 
| 22573 | 1465  | 
val (fixed_params, imports) =  | 
| 18806 | 1466  | 
(case locale of  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1467  | 
NONE => ([], empty)  | 
| 15531 | 1468  | 
| SOME name =>  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1469  | 
          let val {params = ps, ...} = the_locale thy name
 | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1470  | 
in (map fst ps, Locale name) end);  | 
| 
19991
 
0c9650664d47
Locales no longer generate views.  The following functions have changed
 
ballarin 
parents: 
19984 
diff
changeset
 | 
1471  | 
val ((((locale_ctxt, _), (elems_ctxt, _, _)), _), concl') =  | 
| 22573 | 1472  | 
prep_ctxt false fixed_params imports elems concl ctxt;  | 
| 
19991
 
0c9650664d47
Locales no longer generate views.  The following functions have changed
 
ballarin 
parents: 
19984 
diff
changeset
 | 
1473  | 
in (locale, locale_ctxt, elems_ctxt, concl') end;  | 
| 
13399
 
c136276dc847
support locale ``views'' (for cumulative predicates);
 
wenzelm 
parents: 
13394 
diff
changeset
 | 
1474  | 
|
| 22573 | 1475  | 
fun prep_expr prep imports body ctxt =  | 
| 19780 | 1476  | 
let  | 
| 22573 | 1477  | 
val (((_, import_elemss), (ctxt', elemss, _)), _) = prep imports body ctxt;  | 
| 19780 | 1478  | 
val all_elems = maps snd (import_elemss @ elemss);  | 
1479  | 
in (all_elems, ctxt') end;  | 
|
1480  | 
||
| 
12529
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1481  | 
in  | 
| 
 
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
 
wenzelm 
parents: 
12514 
diff
changeset
 | 
1482  | 
|
| 18806 | 1483  | 
val read_ctxt = prep_context_statement intern_expr read_elemss read_facts;  | 
1484  | 
val cert_ctxt = prep_context_statement (K I) cert_elemss cert_facts;  | 
|
| 
14215
 
ebf291f3b449
Improvements to Isar/Locales:  premises generated by "includes" elements
 
ballarin 
parents: 
13629 
diff
changeset
 | 
1485  | 
|
| 22573 | 1486  | 
fun read_context imports body ctxt = #1 (read_ctxt true [] imports (map Elem body) [] ctxt);  | 
1487  | 
fun cert_context imports body ctxt = #1 (cert_ctxt true [] imports (map Elem body) [] ctxt);  | 
|
| 12502 | 1488  | 
|
| 19780 | 1489  | 
val read_expr = prep_expr read_context;  | 
1490  | 
val cert_expr = prep_expr cert_context;  | 
|
1491  | 
||
| 21035 | 1492  | 
fun read_context_statement loc = prep_statement intern read_ctxt loc;  | 
1493  | 
fun read_context_statement_i loc = prep_statement (K I) read_ctxt loc;  | 
|
1494  | 
fun cert_context_statement loc = prep_statement (K I) cert_ctxt loc;  | 
|
| 18806 | 1495  | 
|
| 12502 | 1496  | 
end;  | 
| 11896 | 1497  | 
|
1498  | 
||
| 
21665
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1499  | 
(* init *)  | 
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1500  | 
|
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1501  | 
fun init loc =  | 
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1502  | 
ProofContext.init  | 
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1503  | 
#> (#2 o cert_context_statement (SOME loc) [] []);  | 
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1504  | 
|
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1505  | 
|
| 13336 | 1506  | 
(* print locale *)  | 
| 12070 | 1507  | 
|
| 22573 | 1508  | 
fun print_locale thy show_facts imports body =  | 
1509  | 
let val (all_elems, ctxt) = read_expr imports body (ProofContext.init thy) in  | 
|
| 18137 | 1510  | 
Pretty.big_list "locale elements:" (all_elems  | 
| 17316 | 1511  | 
|> (if show_facts then I else filter (fn Notes _ => false | _ => true))  | 
| 21701 | 1512  | 
|> map (Element.pretty_ctxt ctxt) |> filter_out null  | 
1513  | 
|> map Pretty.chunks)  | 
|
| 13336 | 1514  | 
|> Pretty.writeln  | 
| 12277 | 1515  | 
end;  | 
| 12070 | 1516  | 
|
1517  | 
||
| 12706 | 1518  | 
|
| 16144 | 1519  | 
(** store results **)  | 
| 12702 | 1520  | 
|
| 
19018
 
88b4979193d8
added abbreviations: activated by init, no expressions yet;
 
wenzelm 
parents: 
18964 
diff
changeset
 | 
1521  | 
(* naming of interpreted theorems *)  | 
| 15696 | 1522  | 
|
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1523  | 
fun global_note_prefix_i kind (fully_qualified, prfx) args thy =  | 
| 16144 | 1524  | 
thy  | 
| 22796 | 1525  | 
|> Sign.qualified_names  | 
1526  | 
|> (if fully_qualified then Sign.sticky_prefix prfx else Sign.add_path prfx)  | 
|
| 16144 | 1527  | 
|> PureThy.note_thmss_i kind args  | 
| 22796 | 1528  | 
||> Sign.restore_naming thy;  | 
| 15696 | 1529  | 
|
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1530  | 
fun local_note_prefix_i kind (fully_qualified, prfx) args ctxt =  | 
| 16144 | 1531  | 
ctxt  | 
| 
19061
 
ffbbac0261c9
removed distinct, renamed gen_distinct to distinct;
 
wenzelm 
parents: 
19025 
diff
changeset
 | 
1532  | 
|> ProofContext.qualified_names  | 
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1533  | 
|> (if fully_qualified then ProofContext.sticky_prefix prfx else ProofContext.add_path prfx)  | 
| 21441 | 1534  | 
|> ProofContext.note_thmss_i kind args  | 
| 19780 | 1535  | 
||> ProofContext.restore_naming ctxt;  | 
| 16144 | 1536  | 
|
| 15696 | 1537  | 
|
| 23918 | 1538  | 
(* join equations of an id with already accumulated ones *)  | 
1539  | 
||
1540  | 
fun join_eqns get_reg prep_id ctxt id eqns =  | 
|
1541  | 
let  | 
|
1542  | 
val id' = prep_id id  | 
|
1543  | 
val eqns' = case get_reg id'  | 
|
1544  | 
of NONE => eqns  | 
|
1545  | 
| SOME (_, _, eqns') => Termtab.join (fn t => fn (_, e) => e) (eqns, eqns')  | 
|
1546  | 
handle Termtab.DUP t =>  | 
|
1547  | 
	      error ("Conflicting interpreting equations for term " ^
 | 
|
| 24920 | 1548  | 
quote (Syntax.string_of_term ctxt t))  | 
| 23918 | 1549  | 
in ((id', eqns'), eqns') end;  | 
1550  | 
||
1551  | 
||
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1552  | 
(* collect witnesses and equations up to a particular target for global  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1553  | 
registration; requires parameters and flattened list of identifiers  | 
| 17138 | 1554  | 
instead of recomputing it from the target *)  | 
1555  | 
||
1556  | 
fun collect_global_witnesses thy parms ids vts = let  | 
|
1557  | 
val ts = map Logic.unvarify vts;  | 
|
1558  | 
val (parms, parmTs) = split_list parms;  | 
|
| 19810 | 1559  | 
val parmvTs = map Logic.varifyT parmTs;  | 
| 17138 | 1560  | 
val vtinst = fold (Sign.typ_match thy) (parmvTs ~~ map Term.fastype_of ts) Vartab.empty;  | 
1561  | 
val tinst = Vartab.dest vtinst |> map (fn ((x, 0), (_, T)) => (x, T))  | 
|
| 18137 | 1562  | 
|> Symtab.make;  | 
| 17138 | 1563  | 
(* replace parameter names in ids by instantiations *)  | 
1564  | 
val vinst = Symtab.make (parms ~~ vts);  | 
|
| 17412 | 1565  | 
fun vinst_names ps = map (the o Symtab.lookup vinst) ps;  | 
| 17138 | 1566  | 
val inst = Symtab.make (parms ~~ ts);  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1567  | 
val inst_ids = map (apfst (apsnd vinst_names)) ids;  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1568  | 
val assumed_ids' = map_filter (fn (id, (_, Assumed _)) => SOME id | _ => NONE) inst_ids;  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1569  | 
val wits = maps (#2 o the o get_global_registration thy) assumed_ids';  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1570  | 
|
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1571  | 
val ids' = map fst inst_ids;  | 
| 23918 | 1572  | 
val eqns =  | 
1573  | 
fold_map (join_eqns (get_global_registration thy) I (ProofContext.init thy))  | 
|
1574  | 
ids' Termtab.empty |> snd |> Termtab.dest |> map snd;  | 
|
| 
24638
 
69cb317edf73
Morphisms applied in global interpretations behave correctly on types and terms.
 
ballarin 
parents: 
24525 
diff
changeset
 | 
1575  | 
|
| 
 
69cb317edf73
Morphisms applied in global interpretations behave correctly on types and terms.
 
ballarin 
parents: 
24525 
diff
changeset
 | 
1576  | 
val tinst' = fold (Sign.typ_match thy) (parmvTs ~~ map Term.fastype_of vts) Vartab.empty  | 
| 
 
69cb317edf73
Morphisms applied in global interpretations behave correctly on types and terms.
 
ballarin 
parents: 
24525 
diff
changeset
 | 
1577  | 
|> Vartab.dest |> map (fn ((x, 0), (_, T)) => (x, T)) |> Symtab.make;  | 
| 
 
69cb317edf73
Morphisms applied in global interpretations behave correctly on types and terms.
 
ballarin 
parents: 
24525 
diff
changeset
 | 
1578  | 
in ((tinst', vinst), (tinst, inst), wits, eqns) end;  | 
| 17138 | 1579  | 
|
1580  | 
||
| 15696 | 1581  | 
(* store instantiations of args for all registered interpretations  | 
1582  | 
of the theory *)  | 
|
1583  | 
||
| 21441 | 1584  | 
fun note_thmss_registrations target (kind, args) thy =  | 
| 15596 | 1585  | 
let  | 
| 19278 | 1586  | 
val parms = the_locale thy target |> #params |> map fst;  | 
| 16458 | 1587  | 
val ids = flatten (ProofContext.init thy, intern_expr thy)  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1588  | 
(([], Symtab.empty), Expr (Locale target)) |> fst |> fst;  | 
| 15696 | 1589  | 
|
1590  | 
val regs = get_global_registrations thy target;  | 
|
1591  | 
||
1592  | 
(* add args to thy for all registrations *)  | 
|
| 15596 | 1593  | 
|
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1594  | 
fun activate (vts, (((fully_qualified, prfx), atts2), _, _)) thy =  | 
| 15696 | 1595  | 
let  | 
| 
24638
 
69cb317edf73
Morphisms applied in global interpretations behave correctly on types and terms.
 
ballarin 
parents: 
24525 
diff
changeset
 | 
1596  | 
val (vinsts, insts, prems, eqns) = collect_global_witnesses thy parms ids vts;  | 
| 20911 | 1597  | 
val attrib = Attrib.attribute_i thy;  | 
| 
22241
 
5b2dc1b30e7a
interpretation: attempt to be more serious about name_morphism;
 
wenzelm 
parents: 
21962 
diff
changeset
 | 
1598  | 
val inst_atts = Args.morph_values  | 
| 
 
5b2dc1b30e7a
interpretation: attempt to be more serious about name_morphism;
 
wenzelm 
parents: 
21962 
diff
changeset
 | 
1599  | 
(Morphism.name_morphism (NameSpace.qualified prfx) $>  | 
| 
24638
 
69cb317edf73
Morphisms applied in global interpretations behave correctly on types and terms.
 
ballarin 
parents: 
24525 
diff
changeset
 | 
1600  | 
Element.inst_morphism' thy vinsts insts $>  | 
| 
22241
 
5b2dc1b30e7a
interpretation: attempt to be more serious about name_morphism;
 
wenzelm 
parents: 
21962 
diff
changeset
 | 
1601  | 
Element.satisfy_morphism prems $>  | 
| 23918 | 1602  | 
Morphism.term_morphism (MetaSimplifier.rewrite_term thy eqns []) $>  | 
1603  | 
Morphism.thm_morphism (MetaSimplifier.rewrite_rule eqns #> Drule.standard));  | 
|
| 20911 | 1604  | 
val inst_thm =  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1605  | 
Element.inst_thm thy insts #> Element.satisfy_thm prems #>  | 
| 23918 | 1606  | 
MetaSimplifier.rewrite_rule eqns #>  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1607  | 
Drule.standard;  | 
| 20911 | 1608  | 
val args' = args |> map (fn ((name, atts), bs) =>  | 
1609  | 
((name, map (attrib o inst_atts) atts),  | 
|
1610  | 
bs |> map (fn (ths, more_atts) =>  | 
|
1611  | 
(map inst_thm ths, map attrib (map inst_atts more_atts @ atts2)))));  | 
|
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1612  | 
in global_note_prefix_i kind (fully_qualified, prfx) args' thy |> snd end;  | 
| 18190 | 1613  | 
in fold activate regs thy end;  | 
| 15596 | 1614  | 
|
1615  | 
||
| 20911 | 1616  | 
(* locale results *)  | 
| 12958 | 1617  | 
|
| 21441 | 1618  | 
fun add_thmss loc kind args ctxt =  | 
| 12706 | 1619  | 
let  | 
| 21582 | 1620  | 
val (ctxt', ([(_, [Notes args'])], _)) =  | 
| 21441 | 1621  | 
activate_facts true cert_facts  | 
1622  | 
        (ctxt, [((("", []), Assumed []), [Ext (Notes (kind, args))])]);
 | 
|
| 20911 | 1623  | 
val ctxt'' = ctxt' |> ProofContext.theory  | 
1624  | 
(change_locale loc  | 
|
| 22573 | 1625  | 
(fn (axiom, imports, elems, params, lparams, decls, regs, intros) =>  | 
1626  | 
(axiom, imports, elems @ [(Notes args', stamp ())],  | 
|
| 
21499
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
1627  | 
params, lparams, decls, regs, intros))  | 
| 21441 | 1628  | 
#> note_thmss_registrations loc args');  | 
| 21582 | 1629  | 
in ctxt'' end;  | 
| 15703 | 1630  | 
|
| 11896 | 1631  | 
|
| 
21665
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1632  | 
(* declarations *)  | 
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1633  | 
|
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1634  | 
local  | 
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1635  | 
|
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1636  | 
fun decl_attrib decl phi = Thm.declaration_attribute (K (decl phi));  | 
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1637  | 
|
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1638  | 
fun add_decls add loc decl =  | 
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1639  | 
ProofContext.theory (change_locale loc  | 
| 22573 | 1640  | 
(fn (axiom, imports, elems, params, lparams, decls, regs, intros) =>  | 
1641  | 
(axiom, imports, elems, params, lparams, add (decl, stamp ()) decls, regs, intros))) #>  | 
|
| 24006 | 1642  | 
add_thmss loc Thm.internalK  | 
1643  | 
    [(("", [Attrib.internal (decl_attrib decl)]), [([Drule.dummy_thm], [])])];
 | 
|
| 
21665
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1644  | 
|
| 23418 | 1645  | 
in  | 
| 
21665
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1646  | 
|
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1647  | 
val add_type_syntax = add_decls (apfst o cons);  | 
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1648  | 
val add_term_syntax = add_decls (apsnd o cons);  | 
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1649  | 
val add_declaration = add_decls (K I);  | 
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1650  | 
|
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1651  | 
end;  | 
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1652  | 
|
| 
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1653  | 
|
| 18137 | 1654  | 
|
1655  | 
(** define locales **)  | 
|
1656  | 
||
| 13336 | 1657  | 
(* predicate text *)  | 
| 15596 | 1658  | 
(* CB: generate locale predicates and delta predicates *)  | 
| 13336 | 1659  | 
|
| 13375 | 1660  | 
local  | 
1661  | 
||
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1662  | 
(* introN: name of theorems for introduction rules of locale and  | 
| 
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1663  | 
delta predicates;  | 
| 
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1664  | 
axiomsN: name of theorem set with destruct rules for locale predicates,  | 
| 
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1665  | 
also name suffix of delta predicates. *)  | 
| 
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1666  | 
|
| 13375 | 1667  | 
val introN = "intro";  | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1668  | 
val axiomsN = "axioms";  | 
| 13375 | 1669  | 
|
| 16458 | 1670  | 
fun atomize_spec thy ts =  | 
| 13375 | 1671  | 
let  | 
| 23418 | 1672  | 
val t = Logic.mk_conjunction_balanced ts;  | 
| 16458 | 1673  | 
val body = ObjectLogic.atomize_term thy t;  | 
| 13375 | 1674  | 
val bodyT = Term.fastype_of body;  | 
1675  | 
in  | 
|
| 16458 | 1676  | 
if bodyT = propT then (t, propT, Thm.reflexive (Thm.cterm_of thy t))  | 
| 23591 | 1677  | 
else (body, bodyT, ObjectLogic.atomize (Thm.cterm_of thy t))  | 
| 13375 | 1678  | 
end;  | 
1679  | 
||
| 13394 | 1680  | 
fun aprop_tr' n c = (c, fn args =>  | 
1681  | 
if length args = n then Syntax.const "_aprop" $ Term.list_comb (Syntax.free c, args)  | 
|
1682  | 
else raise Match);  | 
|
| 13336 | 1683  | 
|
| 15104 | 1684  | 
(* CB: define one predicate including its intro rule and axioms  | 
1685  | 
- bname: predicate name  | 
|
1686  | 
- parms: locale parameters  | 
|
1687  | 
- defs: thms representing substitutions from defines elements  | 
|
1688  | 
- ts: terms representing locale assumptions (not normalised wrt. defs)  | 
|
1689  | 
- norm_ts: terms representing locale assumptions (normalised wrt. defs)  | 
|
1690  | 
- thy: the theory  | 
|
1691  | 
*)  | 
|
1692  | 
||
| 13420 | 1693  | 
fun def_pred bname parms defs ts norm_ts thy =  | 
| 13375 | 1694  | 
let  | 
| 16458 | 1695  | 
val name = Sign.full_name thy bname;  | 
| 13375 | 1696  | 
|
| 16458 | 1697  | 
val (body, bodyT, body_eq) = atomize_spec thy norm_ts;  | 
| 13394 | 1698  | 
val env = Term.add_term_free_names (body, []);  | 
| 20664 | 1699  | 
val xs = filter (member (op =) env o #1) parms;  | 
| 13394 | 1700  | 
val Ts = map #2 xs;  | 
| 23178 | 1701  | 
val extraTs = (Term.term_tfrees body \\ List.foldr Term.add_typ_tfrees [] Ts)  | 
| 13394 | 1702  | 
|> Library.sort_wrt #1 |> map TFree;  | 
| 
13399
 
c136276dc847
support locale ``views'' (for cumulative predicates);
 
wenzelm 
parents: 
13394 
diff
changeset
 | 
1703  | 
val predT = map Term.itselfT extraTs ---> Ts ---> bodyT;  | 
| 13336 | 1704  | 
|
| 13394 | 1705  | 
val args = map Logic.mk_type extraTs @ map Free xs;  | 
1706  | 
val head = Term.list_comb (Const (name, predT), args);  | 
|
| 
18123
 
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
 
wenzelm 
parents: 
18038 
diff
changeset
 | 
1707  | 
val statement = ObjectLogic.ensure_propT thy head;  | 
| 13375 | 1708  | 
|
| 18358 | 1709  | 
val ([pred_def], defs_thy) =  | 
| 13375 | 1710  | 
thy  | 
| 13394 | 1711  | 
|> (if bodyT <> propT then I else  | 
| 22796 | 1712  | 
Sign.add_trfuns ([], [], map (aprop_tr' (length args)) (NameSpace.accesses' name), []))  | 
1713  | 
|> Sign.add_consts_i [(bname, predT, NoSyn)]  | 
|
| 13375 | 1714  | 
|> PureThy.add_defs_i false [((Thm.def_name bname, Logic.mk_equals (head, body)), [])];  | 
| 20059 | 1715  | 
val defs_ctxt = ProofContext.init defs_thy |> Variable.declare_term head;  | 
| 13394 | 1716  | 
|
| 16458 | 1717  | 
val cert = Thm.cterm_of defs_thy;  | 
| 13375 | 1718  | 
|
| 20059 | 1719  | 
val intro = Goal.prove_global defs_thy [] norm_ts statement (fn _ =>  | 
| 21708 | 1720  | 
MetaSimplifier.rewrite_goals_tac [pred_def] THEN  | 
| 13375 | 1721  | 
Tactic.compose_tac (false, body_eq RS Drule.equal_elim_rule1, 1) 1 THEN  | 
| 23418 | 1722  | 
Tactic.compose_tac (false,  | 
1723  | 
Conjunction.intr_balanced (map (Thm.assume o cert) norm_ts), 0) 1);  | 
|
| 13375 | 1724  | 
|
1725  | 
val conjuncts =  | 
|
| 19423 | 1726  | 
(Drule.equal_elim_rule2 OF [body_eq,  | 
| 21708 | 1727  | 
MetaSimplifier.rewrite_rule [pred_def] (Thm.assume (cert statement))])  | 
| 23418 | 1728  | 
|> Conjunction.elim_balanced (length ts);  | 
| 
17257
 
0ab67cb765da
introduced binding priority 1 for linear combinators etc.
 
haftmann 
parents: 
17228 
diff
changeset
 | 
1729  | 
val axioms = ts ~~ conjuncts |> map (fn (t, ax) =>  | 
| 20059 | 1730  | 
Element.prove_witness defs_ctxt t  | 
| 21708 | 1731  | 
(MetaSimplifier.rewrite_goals_tac defs THEN  | 
| 13375 | 1732  | 
Tactic.compose_tac (false, ax, 0) 1));  | 
| 
18550
 
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
 
haftmann 
parents: 
18546 
diff
changeset
 | 
1733  | 
in ((statement, intro, axioms), defs_thy) end;  | 
| 13375 | 1734  | 
|
| 
18550
 
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
 
haftmann 
parents: 
18546 
diff
changeset
 | 
1735  | 
fun assumes_to_notes (Assumes asms) axms =  | 
| 21441 | 1736  | 
fold_map (fn (a, spec) => fn axs =>  | 
1737  | 
let val (ps, qs) = chop (length spec) axs  | 
|
1738  | 
in ((a, [(ps, [])]), qs) end) asms axms  | 
|
1739  | 
|> apfst (curry Notes Thm.assumptionK)  | 
|
| 
18550
 
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
 
haftmann 
parents: 
18546 
diff
changeset
 | 
1740  | 
| assumes_to_notes e axms = (e, axms);  | 
| 13394 | 1741  | 
|
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1742  | 
(* CB: the following two change only "new" elems, these have identifier ("", _). *)
 | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1743  | 
|
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1744  | 
(* turn Assumes into Notes elements *)  | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1745  | 
|
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1746  | 
fun change_assumes_elemss axioms elemss =  | 
| 
18550
 
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
 
haftmann 
parents: 
18546 
diff
changeset
 | 
1747  | 
let  | 
| 
21483
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
1748  | 
val satisfy = Element.morph_ctxt (Element.satisfy_morphism axioms);  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1749  | 
    fun change (id as ("", _), es) =
 | 
| 
21483
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
1750  | 
fold_map assumes_to_notes (map satisfy es)  | 
| 
18550
 
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
 
haftmann 
parents: 
18546 
diff
changeset
 | 
1751  | 
#-> (fn es' => pair (id, es'))  | 
| 
 
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
 
haftmann 
parents: 
18546 
diff
changeset
 | 
1752  | 
| change e = pair e;  | 
| 
 
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
 
haftmann 
parents: 
18546 
diff
changeset
 | 
1753  | 
in  | 
| 19780 | 1754  | 
fst (fold_map change elemss (map Element.conclude_witness axioms))  | 
| 
18550
 
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
 
haftmann 
parents: 
18546 
diff
changeset
 | 
1755  | 
end;  | 
| 13394 | 1756  | 
|
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1757  | 
(* adjust hyps of Notes elements *)  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1758  | 
|
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1759  | 
fun change_elemss_hyps axioms elemss =  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1760  | 
let  | 
| 
21483
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
1761  | 
val satisfy = Element.morph_ctxt (Element.satisfy_morphism axioms);  | 
| 
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
1762  | 
    fun change (id as ("", _), es) = (id, map (fn e as Notes _ => satisfy e | e => e) es)
 | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1763  | 
| change e = e;  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1764  | 
in map change elemss end;  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1765  | 
|
| 13394 | 1766  | 
in  | 
| 13375 | 1767  | 
|
| 15104 | 1768  | 
(* CB: main predicate definition function *)  | 
1769  | 
||
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1770  | 
fun define_preds pname (parms, ((exts, exts'), (ints, ints')), defs) elemss thy =  | 
| 13394 | 1771  | 
let  | 
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1772  | 
val ((elemss', more_ts), a_elem, a_intro, thy'') =  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1773  | 
if null exts then ((elemss, []), [], [], thy)  | 
| 13394 | 1774  | 
else  | 
1775  | 
let  | 
|
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1776  | 
val aname = if null ints then pname else pname ^ "_" ^ axiomsN;  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1777  | 
val ((statement, intro, axioms), thy') =  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1778  | 
thy  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1779  | 
|> def_pred aname parms defs exts exts';  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1780  | 
val elemss' = change_assumes_elemss axioms elemss;  | 
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1781  | 
          val a_elem = [(("", []), [Assumes [((pname ^ "_" ^ axiomsN, []), [(statement, [])])]])];
 | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1782  | 
val (_, thy'') =  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1783  | 
thy'  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1784  | 
|> PureThy.note_thmss_qualified Thm.internalK aname [((introN, []), [([intro], [])])];  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1785  | 
in ((elemss', [statement]), a_elem, [intro], thy'') end;  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1786  | 
val (predicate, stmt', elemss'', b_intro, thy'''') =  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1787  | 
if null ints then (([], []), more_ts, elemss' @ a_elem, [], thy'')  | 
| 13394 | 1788  | 
else  | 
1789  | 
let  | 
|
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1790  | 
val ((statement, intro, axioms), thy''') =  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1791  | 
thy''  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1792  | 
|> def_pred pname parms defs (ints @ more_ts) (ints' @ more_ts);  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1793  | 
val cstatement = Thm.cterm_of thy''' statement;  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1794  | 
val elemss'' = change_elemss_hyps axioms elemss';  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1795  | 
          val b_elem = [(("", []),
 | 
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1796  | 
[Assumes [((pname ^ "_" ^ axiomsN, []), [(statement, [])])]])];  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1797  | 
val (_, thy'''') =  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1798  | 
thy'''  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1799  | 
|> PureThy.note_thmss_qualified Thm.internalK pname  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1800  | 
[((introN, []), [([intro], [])]),  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1801  | 
((axiomsN, []), [(map (Drule.standard o Element.conclude_witness) axioms, [])])];  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1802  | 
in (([cstatement], axioms), [statement], elemss'' @ b_elem, [intro], thy'''') end;  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1803  | 
in (((elemss'', predicate, stmt'), (a_intro, b_intro)), thy'''') end;  | 
| 13375 | 1804  | 
|
1805  | 
end;  | 
|
| 13336 | 1806  | 
|
1807  | 
||
| 13297 | 1808  | 
(* add_locale(_i) *)  | 
1809  | 
||
1810  | 
local  | 
|
1811  | 
||
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1812  | 
(* turn Defines into Notes elements, accumulate definition terms *)  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1813  | 
|
| 
19942
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1814  | 
fun defines_to_notes is_ext thy (Defines defs) defns =  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1815  | 
let  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1816  | 
      val defs' = map (fn (_, (def, _)) => (("", []), (def, []))) defs
 | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1817  | 
val notes = map (fn (a, (def, _)) =>  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1818  | 
(a, [([assume (cterm_of thy def)], [])])) defs  | 
| 21441 | 1819  | 
in  | 
1820  | 
(if is_ext then SOME (Notes (Thm.definitionK, notes)) else NONE, defns @ [Defines defs'])  | 
|
1821  | 
end  | 
|
| 
19942
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1822  | 
| defines_to_notes _ _ e defns = (SOME e, defns);  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1823  | 
|
| 
19942
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1824  | 
fun change_defines_elemss thy elemss defns =  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1825  | 
let  | 
| 
19942
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1826  | 
fun change (id as (n, _), es) defns =  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1827  | 
let  | 
| 
19942
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1828  | 
val (es', defns') = fold_map (defines_to_notes (n="") thy) es defns  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1829  | 
in ((id, map_filter I es'), defns') end  | 
| 
 
dc92e3aebc71
Improved handling of defines imported in duplicate.
 
ballarin 
parents: 
19932 
diff
changeset
 | 
1830  | 
in fold_map change elemss defns end;  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1831  | 
|
| 18343 | 1832  | 
fun gen_add_locale prep_ctxt prep_expr  | 
| 22573 | 1833  | 
predicate_name bname raw_imports raw_body thy =  | 
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1834  | 
(* predicate_name: NONE - open locale without predicate  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1835  | 
SOME "" - locale with predicate named as locale  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1836  | 
SOME "name" - locale with predicate named "name" *)  | 
| 13297 | 1837  | 
let  | 
| 16458 | 1838  | 
val name = Sign.full_name thy bname;  | 
| 21962 | 1839  | 
val _ = is_some (get_locale thy name) andalso  | 
1840  | 
      error ("Duplicate definition of locale " ^ quote name);
 | 
|
| 13297 | 1841  | 
|
1842  | 
val thy_ctxt = ProofContext.init thy;  | 
|
| 17228 | 1843  | 
val (((import_ctxt, import_elemss), (body_ctxt, body_elemss, syn)),  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1844  | 
text as (parms, ((_, exts'), _), defs)) =  | 
| 22573 | 1845  | 
prep_ctxt raw_imports raw_body thy_ctxt;  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1846  | 
val elemss = import_elemss @ body_elemss |>  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1847  | 
map_filter (fn ((id, Assumed axs), elems) => SOME (id, elems) | _ => NONE);  | 
| 22573 | 1848  | 
val imports = prep_expr thy raw_imports;  | 
| 13297 | 1849  | 
|
| 23178 | 1850  | 
val extraTs = List.foldr Term.add_term_tfrees [] exts' \\  | 
1851  | 
List.foldr Term.add_typ_tfrees [] (map snd parms);  | 
|
| 17228 | 1852  | 
val _ = if null extraTs then ()  | 
| 17437 | 1853  | 
      else warning ("Additional type variable(s) in locale specification " ^ quote bname);
 | 
| 17228 | 1854  | 
|
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1855  | 
val ((((elemss', predicate as (predicate_statement, predicate_axioms), stmt'), intros),  | 
| 22573 | 1856  | 
pred_thy), (imports, regs)) =  | 
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1857  | 
case predicate_name  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1858  | 
of SOME predicate_name =>  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1859  | 
let  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1860  | 
val predicate_name' = case predicate_name of "" => bname | _ => predicate_name;  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1861  | 
val (elemss', defns) = change_defines_elemss thy elemss [];  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1862  | 
              val elemss'' = elemss' @ [(("", []), defns)];
 | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1863  | 
val (((elemss''', predicate as (_, axioms), stmt'), intros), thy') =  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1864  | 
define_preds predicate_name' text elemss'' thy;  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1865  | 
fun mk_regs elemss wits =  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1866  | 
fold_map (fn (id, elems) => fn wts => let  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1867  | 
val ts = List.concat (List.mapPartial (fn (Assumes asms) =>  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1868  | 
SOME (List.concat (map (map #1 o #2) asms)) | _ => NONE) elems);  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1869  | 
val (wts1, wts2) = chop (length ts) wts  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1870  | 
in ((apsnd (map fst) id, wts1), wts2) end) elemss wits |> fst;  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1871  | 
val regs = mk_regs elemss''' axioms |>  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1872  | 
                    map_filter (fn (("", _), _) => NONE | e => SOME e);
 | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1873  | 
in ((((elemss''', predicate, stmt'), intros), thy'), (empty, regs)) end  | 
| 22573 | 1874  | 
| NONE => ((((elemss, ([], []), []), ([], [])), thy), (imports, []));  | 
| 13420 | 1875  | 
|
| 18137 | 1876  | 
fun axiomify axioms elemss =  | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1877  | 
(axioms, elemss) |> foldl_map (fn (axs, (id, elems)) => let  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19423 
diff
changeset
 | 
1878  | 
val ts = flat (map_filter (fn (Assumes asms) =>  | 
| 
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19423 
diff
changeset
 | 
1879  | 
SOME (maps (map #1 o #2) asms) | _ => NONE) elems);  | 
| 
19018
 
88b4979193d8
added abbreviations: activated by init, no expressions yet;
 
wenzelm 
parents: 
18964 
diff
changeset
 | 
1880  | 
val (axs1, axs2) = chop (length ts) axs;  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
1881  | 
in (axs2, ((id, Assumed axs1), elems)) end)  | 
| 
15206
 
09d78ec709c7
Modified locales: improved implementation of "includes".
 
ballarin 
parents: 
15127 
diff
changeset
 | 
1882  | 
|> snd;  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1883  | 
val (ctxt, (_, facts)) = activate_facts true (K I)  | 
| 20965 | 1884  | 
(ProofContext.init pred_thy, axiomify predicate_axioms elemss');  | 
| 
21393
 
c648e24fd7ee
add_locale: re-init result context (avoids subtle modifications after introducing predicate views internally);
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
1885  | 
val view_ctxt = Assumption.add_view thy_ctxt predicate_statement ctxt;  | 
| 21602 | 1886  | 
val export = Goal.close_result o Goal.norm_result o  | 
1887  | 
singleton (ProofContext.export view_ctxt thy_ctxt);  | 
|
| 13420 | 1888  | 
val facts' = facts |> map (fn (a, ths) => ((a, []), [(map export ths, [])]));  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19423 
diff
changeset
 | 
1889  | 
val elems' = maps #2 (filter (equal "" o #1 o #1) elemss');  | 
| 19783 | 1890  | 
val elems'' = map_filter (fn (Fixes _) => NONE | e => SOME e) elems';  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1891  | 
val axs' = map (Element.assume_witness pred_thy) stmt';  | 
| 
21393
 
c648e24fd7ee
add_locale: re-init result context (avoids subtle modifications after introducing predicate views internally);
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
1892  | 
val loc_ctxt = pred_thy  | 
| 21441 | 1893  | 
|> PureThy.note_thmss_qualified Thm.assumptionK bname facts' |> snd  | 
| 
19061
 
ffbbac0261c9
removed distinct, renamed gen_distinct to distinct;
 
wenzelm 
parents: 
19025 
diff
changeset
 | 
1894  | 
|> declare_locale name  | 
| 
 
ffbbac0261c9
removed distinct, renamed gen_distinct to distinct;
 
wenzelm 
parents: 
19025 
diff
changeset
 | 
1895  | 
|> put_locale name  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1896  | 
       {axiom = axs',
 | 
| 22573 | 1897  | 
imports = imports,  | 
| 19783 | 1898  | 
elems = map (fn e => (e, stamp ())) elems'',  | 
| 19278 | 1899  | 
params = params_of elemss' |> map (fn (x, SOME T) => ((x, T), the (Symtab.lookup syn x))),  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1900  | 
lparams = map #1 (params_of' body_elemss),  | 
| 
21665
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
1901  | 
decls = ([], []),  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1902  | 
regs = regs,  | 
| 
21393
 
c648e24fd7ee
add_locale: re-init result context (avoids subtle modifications after introducing predicate views internally);
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
1903  | 
intros = intros}  | 
| 
 
c648e24fd7ee
add_locale: re-init result context (avoids subtle modifications after introducing predicate views internally);
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
1904  | 
|> init name;  | 
| 
 
c648e24fd7ee
add_locale: re-init result context (avoids subtle modifications after introducing predicate views internally);
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
1905  | 
in (name, loc_ctxt) end;  | 
| 13297 | 1906  | 
|
1907  | 
in  | 
|
1908  | 
||
| 18917 | 1909  | 
val add_locale = gen_add_locale read_context intern_expr;  | 
1910  | 
val add_locale_i = gen_add_locale cert_context (K I);  | 
|
| 13297 | 1911  | 
|
1912  | 
end;  | 
|
1913  | 
||
| 
19018
 
88b4979193d8
added abbreviations: activated by init, no expressions yet;
 
wenzelm 
parents: 
18964 
diff
changeset
 | 
1914  | 
val _ = Context.add_setup  | 
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1915  | 
(add_locale_i (SOME "") "var" empty [Fixes [(Name.internal "x", NONE, NoSyn)]] #>  | 
| 20965 | 1916  | 
snd #> ProofContext.theory_of #>  | 
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
1917  | 
add_locale_i (SOME "") "struct" empty [Fixes [(Name.internal "S", NONE, Structure)]] #>  | 
| 20965 | 1918  | 
snd #> ProofContext.theory_of);  | 
| 15801 | 1919  | 
|
| 13297 | 1920  | 
|
| 12730 | 1921  | 
|
| 17355 | 1922  | 
|
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1923  | 
(** Normalisation of locale statements ---  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1924  | 
discharges goals implied by interpretations **)  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1925  | 
|
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1926  | 
local  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1927  | 
|
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1928  | 
fun locale_assm_intros thy =  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1929  | 
  Symtab.fold (fn (_, {intros = (a, _), ...}) => fn intros => (a @ intros))
 | 
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
1930  | 
(#2 (LocalesData.get thy)) [];  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1931  | 
fun locale_base_intros thy =  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1932  | 
  Symtab.fold (fn (_, {intros = (_, b), ...}) => fn intros => (b @ intros))
 | 
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
1933  | 
(#2 (LocalesData.get thy)) [];  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1934  | 
|
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1935  | 
fun all_witnesses ctxt =  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1936  | 
let  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1937  | 
val thy = ProofContext.theory_of ctxt;  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1938  | 
fun get registrations = Symtab.fold (fn (_, regs) => fn thms =>  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1939  | 
(Registrations.dest thy regs |> map (fn (_, (_, wits, _)) =>  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1940  | 
map Element.conclude_witness wits) |> flat) @ thms)  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1941  | 
registrations [];  | 
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
1942  | 
in get (RegistrationsData.get (Context.Proof ctxt)) end;  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1943  | 
(* FIXME: proper varification *)  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1944  | 
|
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1945  | 
in  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1946  | 
|
| 
19984
 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 
ballarin 
parents: 
19942 
diff
changeset
 | 
1947  | 
fun intro_locales_tac eager ctxt facts st =  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1948  | 
let  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1949  | 
val wits = all_witnesses ctxt |> map Thm.varifyT;  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1950  | 
val thy = ProofContext.theory_of ctxt;  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1951  | 
val intros = locale_base_intros thy @ (if eager then locale_assm_intros thy else []);  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1952  | 
in  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1953  | 
(ALLGOALS (Method.insert_tac facts THEN'  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1954  | 
REPEAT_ALL_NEW (resolve_tac (wits @ intros)))  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1955  | 
THEN Tactic.distinct_subgoals_tac) st  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1956  | 
end;  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1957  | 
|
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1958  | 
val _ = Context.add_setup (Method.add_methods  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1959  | 
  [("intro_locales",
 | 
| 
19984
 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 
ballarin 
parents: 
19942 
diff
changeset
 | 
1960  | 
Method.ctxt_args (fn ctxt => Method.METHOD (intro_locales_tac false ctxt)),  | 
| 
 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 
ballarin 
parents: 
19942 
diff
changeset
 | 
1961  | 
"back-chain introduction rules of locales without unfolding predicates"),  | 
| 
 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 
ballarin 
parents: 
19942 
diff
changeset
 | 
1962  | 
   ("unfold_locales",
 | 
| 
 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 
ballarin 
parents: 
19942 
diff
changeset
 | 
1963  | 
Method.ctxt_args (fn ctxt => Method.METHOD (intro_locales_tac true ctxt)),  | 
| 
 
29bb4659f80a
Method intro_locales replaced by intro_locales and unfold_locales.
 
ballarin 
parents: 
19942 
diff
changeset
 | 
1964  | 
"back-chain all introduction rules of locales")]);  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1965  | 
|
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1966  | 
end;  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
1967  | 
|
| 19780 | 1968  | 
|
| 
15598
 
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
 
ballarin 
parents: 
15596 
diff
changeset
 | 
1969  | 
(** Interpretation commands **)  | 
| 15596 | 1970  | 
|
1971  | 
local  | 
|
1972  | 
||
1973  | 
(* extract proof obligations (assms and defs) from elements *)  | 
|
1974  | 
||
| 19780 | 1975  | 
fun extract_asms_elems ((id, Assumed _), elems) = (id, maps Element.prems_of elems)  | 
| 17138 | 1976  | 
| extract_asms_elems ((id, Derived _), _) = (id, []);  | 
| 15596 | 1977  | 
|
1978  | 
||
| 
15624
 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 
ballarin 
parents: 
15623 
diff
changeset
 | 
1979  | 
(* activate instantiated facts in theory or context *)  | 
| 15596 | 1980  | 
|
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1981  | 
structure Idtab =  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1982  | 
TableFun(type key = string * term list  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1983  | 
val ord = prod_ord string_ord (list_ord Term.fast_term_ord));  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1984  | 
|
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1985  | 
fun gen_activate_facts_elemss mk_ctxt get_reg note attrib std put_reg add_wit add_eqn  | 
| 
18123
 
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
 
wenzelm 
parents: 
18038 
diff
changeset
 | 
1986  | 
attn all_elemss new_elemss propss thmss thy_ctxt =  | 
| 
21499
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
1987  | 
let  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1988  | 
val ctxt = mk_ctxt thy_ctxt;  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1989  | 
val (propss, eq_props) = chop (length new_elemss) propss;  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1990  | 
val (thmss, eq_thms) = chop (length new_elemss) thmss;  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1991  | 
|
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1992  | 
fun activate_elem eqns disch ((fully_qualified, prfx), atts) (Notes (kind, facts)) thy_ctxt =  | 
| 
21499
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
1993  | 
let  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
1994  | 
val ctxt = mk_ctxt thy_ctxt;  | 
| 
22241
 
5b2dc1b30e7a
interpretation: attempt to be more serious about name_morphism;
 
wenzelm 
parents: 
21962 
diff
changeset
 | 
1995  | 
val fact_morphism =  | 
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
1996  | 
Morphism.name_morphism (NameSpace.qualified prfx)  | 
| 22756 | 1997  | 
$> Morphism.term_morphism (MetaSimplifier.rewrite_term  | 
1998  | 
(ProofContext.theory_of ctxt) eqns [])  | 
|
1999  | 
$> Morphism.thm_morphism (disch #> MetaSimplifier.rewrite_rule eqns);  | 
|
| 
21499
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2000  | 
val facts' = facts  | 
| 
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2001  | 
(* discharge hyps in attributes *)  | 
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2002  | 
|> Attrib.map_facts  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2003  | 
(attrib thy_ctxt o Args.morph_values fact_morphism)  | 
| 
21499
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2004  | 
(* append interpretation attributes *)  | 
| 
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2005  | 
|> map (apfst (apsnd (fn a => a @ map (attrib thy_ctxt) atts)))  | 
| 
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2006  | 
(* discharge hyps *)  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2007  | 
|> map (apsnd (map (apfst (map disch))))  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2008  | 
(* unfold eqns *)  | 
| 22756 | 2009  | 
|> map (apsnd (map (apfst (map (MetaSimplifier.rewrite_rule eqns)))))  | 
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
2010  | 
in snd (note kind (fully_qualified, prfx) facts' thy_ctxt) end  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2011  | 
| activate_elem _ _ _ _ thy_ctxt = thy_ctxt;  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2012  | 
|
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2013  | 
fun activate_elems eqns disch ((id, _), elems) thy_ctxt =  | 
| 
21499
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2014  | 
let  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2015  | 
val (prfx_atts, _, _) = case get_reg thy_ctxt id  | 
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
2016  | 
of SOME x => x  | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
2017  | 
          | NONE => sys_error ("Unknown registration of " ^ quote (fst id)
 | 
| 
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
2018  | 
^ " while activating facts.");  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2019  | 
in fold (activate_elem (the (Idtab.lookup eqns id)) disch prfx_atts) elems thy_ctxt end;  | 
| 17033 | 2020  | 
|
| 
21499
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2021  | 
val thy_ctxt' = thy_ctxt  | 
| 
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2022  | 
(* add registrations *)  | 
| 
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2023  | 
|> fold (fn ((id, _), _) => put_reg id attn) new_elemss  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2024  | 
(* add witnesses of Assumed elements (only those generate proof obligations) *)  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2025  | 
|> fold (fn (id, thms) => fold (add_wit id) thms) (map fst propss ~~ thmss)  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2026  | 
(* add equations *)  | 
| 24693 | 2027  | 
|> fold (fn (id, thms) => fold (add_eqn id) thms) (map fst eq_props ~~  | 
| 
24952
 
f336c36f41a0
removed redundant strip_vars/abs_eqn, use improved Drule.abs_def instead;
 
wenzelm 
parents: 
24941 
diff
changeset
 | 
2028  | 
(map o map) (Drule.abs_def o LocalDefs.meta_rewrite_rule ctxt o  | 
| 23918 | 2029  | 
Element.conclude_witness) eq_thms);  | 
| 15596 | 2030  | 
|
| 
21499
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2031  | 
val prems = flat (map_filter  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2032  | 
(fn ((id, Assumed _), _) => Option.map #2 (get_reg thy_ctxt' id)  | 
| 
21499
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2033  | 
| ((_, Derived _), _) => NONE) all_elemss);  | 
| 
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2034  | 
val satisfy = Element.satisfy_morphism prems;  | 
| 
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2035  | 
val thy_ctxt'' = thy_ctxt'  | 
| 
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2036  | 
(* add witnesses of Derived elements *)  | 
| 
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2037  | 
|> fold (fn (id, thms) => fold (add_wit id o Element.morph_witness satisfy) thms)  | 
| 
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2038  | 
(map_filter (fn ((_, Assumed _), _) => NONE  | 
| 
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2039  | 
| ((id, Derived thms), _) => SOME (id, thms)) new_elemss)  | 
| 19780 | 2040  | 
|
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2041  | 
(* Accumulate equations *)  | 
| 23918 | 2042  | 
val all_eqns =  | 
2043  | 
fold_map (join_eqns (get_reg thy_ctxt'') (fst o fst) ctxt) all_elemss Termtab.empty  | 
|
2044  | 
|> fst  | 
|
2045  | 
|> map (apsnd (map snd o Termtab.dest))  | 
|
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2046  | 
|> (fn xs => fold Idtab.update xs Idtab.empty)  | 
| 23918 | 2047  | 
(* Idtab.make wouldn't work here: can have conflicting duplicates,  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2048  | 
because instantiation may equate ids and equations are accumulated! *)  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2049  | 
|
| 
21499
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2050  | 
val disch' = std o Morphism.thm satisfy; (* FIXME *)  | 
| 
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2051  | 
in  | 
| 
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2052  | 
thy_ctxt''  | 
| 
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2053  | 
(* add facts to theory or context *)  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2054  | 
|> fold (activate_elems all_eqns disch') new_elemss  | 
| 
21499
 
fbd6422a847a
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 
wenzelm 
parents: 
21483 
diff
changeset
 | 
2055  | 
end;  | 
| 15596 | 2056  | 
|
| 17355 | 2057  | 
fun global_activate_facts_elemss x = gen_activate_facts_elemss  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2058  | 
ProofContext.init  | 
| 15696 | 2059  | 
(fn thy => fn (name, ps) =>  | 
2060  | 
get_global_registration thy (name, map Logic.varify ps))  | 
|
| 21441 | 2061  | 
global_note_prefix_i  | 
| 18728 | 2062  | 
Attrib.attribute_i Drule.standard  | 
| 17033 | 2063  | 
(fn (name, ps) => put_global_registration (name, map Logic.varify ps))  | 
| 19780 | 2064  | 
(fn (n, ps) => add_global_witness (n, map Logic.varify ps) o  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
2065  | 
Element.map_witness (fn (t, th) => (Logic.legacy_unvarify t, legacy_unvarify th))  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2066  | 
(* FIXME *))  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2067  | 
(fn (n, ps) => add_global_equation (n, map Logic.varify ps))  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2068  | 
x;  | 
| 17355 | 2069  | 
|
2070  | 
fun local_activate_facts_elemss x = gen_activate_facts_elemss  | 
|
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2071  | 
I  | 
| 15696 | 2072  | 
get_local_registration  | 
| 
19018
 
88b4979193d8
added abbreviations: activated by init, no expressions yet;
 
wenzelm 
parents: 
18964 
diff
changeset
 | 
2073  | 
local_note_prefix_i  | 
| 18728 | 2074  | 
(Attrib.attribute_i o ProofContext.theory_of) I  | 
| 17033 | 2075  | 
put_local_registration  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2076  | 
add_local_witness  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2077  | 
add_local_equation  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2078  | 
x;  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2079  | 
|
| 
24941
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2080  | 
fun prep_instantiations parse_term parse_prop ctxt parms (insts, eqns) =  | 
| 15596 | 2081  | 
let  | 
| 
24941
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2082  | 
(* parameters *)  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2083  | 
val (parm_names, parm_types) = parms |> split_list  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2084  | 
||> map (TypeInfer.paramify_vars o Logic.varifyT);  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2085  | 
val type_parms = fold Term.add_tvarsT parm_types [] |> map TVar;  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2086  | 
val type_parm_names = fold Term.add_tfreesT (map snd parms) [] |> map fst;  | 
| 
22772
 
e0788ff2e811
read_instantiations: proper type-inference with fixed variables, infer parameter types as well;
 
wenzelm 
parents: 
22756 
diff
changeset
 | 
2087  | 
|
| 
 
e0788ff2e811
read_instantiations: proper type-inference with fixed variables, infer parameter types as well;
 
wenzelm 
parents: 
22756 
diff
changeset
 | 
2088  | 
(* parameter instantiations *)  | 
| 
 
e0788ff2e811
read_instantiations: proper type-inference with fixed variables, infer parameter types as well;
 
wenzelm 
parents: 
22756 
diff
changeset
 | 
2089  | 
val d = length parms - length insts;  | 
| 
 
e0788ff2e811
read_instantiations: proper type-inference with fixed variables, infer parameter types as well;
 
wenzelm 
parents: 
22756 
diff
changeset
 | 
2090  | 
val insts =  | 
| 
 
e0788ff2e811
read_instantiations: proper type-inference with fixed variables, infer parameter types as well;
 
wenzelm 
parents: 
22756 
diff
changeset
 | 
2091  | 
if d < 0 then error "More arguments than parameters in instantiation."  | 
| 
 
e0788ff2e811
read_instantiations: proper type-inference with fixed variables, infer parameter types as well;
 
wenzelm 
parents: 
22756 
diff
changeset
 | 
2092  | 
else insts @ replicate d NONE;  | 
| 
24941
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2093  | 
val (given_ps, given_insts) =  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2094  | 
((parm_names ~~ parm_types) ~~ insts) |> map_filter  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2095  | 
(fn (_, NONE) => NONE  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2096  | 
| ((n, T), SOME inst) => SOME ((n, T), inst))  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2097  | 
|> split_list;  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2098  | 
val (given_parm_names, given_parm_types) = given_ps |> split_list;  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2099  | 
|
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2100  | 
(* prepare insts / eqns *)  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2101  | 
val given_insts' = map (parse_term ctxt) given_insts;  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2102  | 
val eqns' = map (parse_prop ctxt) eqns;  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2103  | 
|
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2104  | 
(* infer types *)  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2105  | 
val res = Syntax.check_terms ctxt  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2106  | 
(map Logic.mk_type type_parms @  | 
| 25002 | 2107  | 
map2 TypeInfer.constrain given_parm_types given_insts' @  | 
| 
24941
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2108  | 
eqns');  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2109  | 
val ctxt' = ctxt |> fold Variable.auto_fixes res;  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2110  | 
|
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2111  | 
(* results *)  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2112  | 
val (type_parms'', res') = chop (length type_parms) res;  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2113  | 
val (given_insts'', eqns'') = chop (length given_insts) res';  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2114  | 
val instT = Symtab.make (type_parm_names ~~ map Logic.dest_type type_parms'');  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2115  | 
val inst = Symtab.make (given_parm_names ~~ given_insts'');  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2116  | 
val standard = Variable.export_morphism ctxt' ctxt;  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2117  | 
in ((instT, inst), eqns'') end;  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2118  | 
|
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2119  | 
|
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2120  | 
val read_instantiations = prep_instantiations Syntax.parse_term Syntax.parse_prop;  | 
| 
24952
 
f336c36f41a0
removed redundant strip_vars/abs_eqn, use improved Drule.abs_def instead;
 
wenzelm 
parents: 
24941 
diff
changeset
 | 
2121  | 
val check_instantiations = prep_instantiations (K I) (K I);  | 
| 
15598
 
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
 
ballarin 
parents: 
15596 
diff
changeset
 | 
2122  | 
|
| 
22772
 
e0788ff2e811
read_instantiations: proper type-inference with fixed variables, infer parameter types as well;
 
wenzelm 
parents: 
22756 
diff
changeset
 | 
2123  | 
fun gen_prep_registration mk_ctxt test_reg activate  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2124  | 
prep_attr prep_expr prep_insts  | 
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2125  | 
thy_ctxt raw_attn raw_expr raw_insts =  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2126  | 
let  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2127  | 
val ctxt = mk_ctxt thy_ctxt;  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2128  | 
val thy = ProofContext.theory_of ctxt;  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2129  | 
val ctxt' = ProofContext.init thy;  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2130  | 
|
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2131  | 
val attn = (apsnd o map)  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2132  | 
(Attrib.crude_closure ctxt o Args.assignable o prep_attr thy) raw_attn;  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2133  | 
val expr = prep_expr thy raw_expr;  | 
| 15596 | 2134  | 
|
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2135  | 
val pts = params_of_expr ctxt' [] expr ([], Symtab.empty, Symtab.empty);  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2136  | 
val params_ids = make_params_ids (#1 pts);  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2137  | 
val raw_params_elemss = make_raw_params_elemss pts;  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2138  | 
val ((ids, _), raw_elemss) = flatten (ctxt', I) (([], Symtab.empty), Expr expr);  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2139  | 
val ((parms, all_elemss, _), (_, (_, defs, _))) =  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2140  | 
read_elemss false ctxt' [] (raw_params_elemss @ raw_elemss) [];  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2141  | 
|
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2142  | 
(** compute instantiation **)  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2143  | 
|
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2144  | 
(* consistency check: equations need to be stored in a particular locale,  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2145  | 
therefore if equations are present locale expression must be a name *)  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2146  | 
|
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2147  | 
val _ = case (expr, snd raw_insts) of  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2148  | 
(Locale _, _) => () | (_, []) => ()  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2149  | 
| (_, _) => error "Interpretations with `where' only permitted if locale expression is a name.";  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2150  | 
|
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2151  | 
(* read or certify instantiation *)  | 
| 23918 | 2152  | 
val ((instT, inst1), eqns) = prep_insts ctxt parms raw_insts;  | 
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2153  | 
|
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2154  | 
(* defined params without given instantiation *)  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2155  | 
val not_given = filter_out (Symtab.defined inst1 o fst) parms;  | 
| 18137 | 2156  | 
fun add_def (p, pT) inst =  | 
| 15596 | 2157  | 
let  | 
2158  | 
val (t, T) = case find_first (fn (Free (a, _), _) => a = p) defs of  | 
|
2159  | 
               NONE => error ("Instance missing for parameter " ^ quote p)
 | 
|
2160  | 
| SOME (Free (_, T), t) => (t, T);  | 
|
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2161  | 
val d = Element.inst_term (instT, inst) t;  | 
| 18137 | 2162  | 
in Symtab.update_new (p, d) inst end;  | 
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2163  | 
val inst2 = fold add_def not_given inst1;  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2164  | 
val inst_morphism = Element.inst_morphism thy (instT, inst2);  | 
| 18137 | 2165  | 
(* Note: insts contain no vars. *)  | 
| 15596 | 2166  | 
|
2167  | 
(** compute proof obligations **)  | 
|
2168  | 
||
| 
15598
 
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
 
ballarin 
parents: 
15596 
diff
changeset
 | 
2169  | 
(* restore "small" ids *)  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
2170  | 
val ids' = map (fn ((n, ps), (_, mode)) =>  | 
| 19783 | 2171  | 
((n, map (fn p => Free (p, (the o AList.lookup (op =) parms) p)) ps), mode))  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
2172  | 
ids;  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
2173  | 
val (_, all_elemss') = chop (length raw_params_elemss) all_elemss  | 
| 15596 | 2174  | 
(* instantiate ids and elements *)  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
19914 
diff
changeset
 | 
2175  | 
val inst_elemss = (ids' ~~ all_elemss') |> map (fn (((n, ps), _), ((_, mode), elems)) =>  | 
| 
21483
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
2176  | 
((n, map (Morphism.term inst_morphism) ps),  | 
| 
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
2177  | 
map (fn Int e => Element.morph_ctxt inst_morphism e) elems)  | 
| 
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
2178  | 
|> apfst (fn id => (id, map_mode (map (Element.morph_witness inst_morphism)) mode)));  | 
| 15596 | 2179  | 
|
| 
15624
 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 
ballarin 
parents: 
15623 
diff
changeset
 | 
2180  | 
(* remove fragments already registered with theory or context *)  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2181  | 
val new_inst_elemss = filter_out (fn ((id, _), _) =>  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2182  | 
test_reg thy_ctxt id) inst_elemss;  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2183  | 
(* val new_ids = map #1 new_inst_elemss; *)  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2184  | 
|
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2185  | 
(* equations *)  | 
| 23918 | 2186  | 
val eqn_elems = if null eqns then []  | 
2187  | 
else [(Library.last_elem inst_elemss |> fst |> fst, eqns)];  | 
|
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2188  | 
|
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2189  | 
val propss = map extract_asms_elems new_inst_elemss @ eqn_elems;  | 
| 15596 | 2190  | 
|
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2191  | 
in (propss, activate attn inst_elemss new_inst_elemss propss) end;  | 
| 15596 | 2192  | 
|
| 22756 | 2193  | 
fun gen_prep_global_registration mk_ctxt = gen_prep_registration ProofContext.init  | 
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2194  | 
(fn thy => fn (name, ps) => test_global_registration thy (name, map Logic.varify ps))  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2195  | 
global_activate_facts_elemss mk_ctxt;  | 
| 
15624
 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 
ballarin 
parents: 
15623 
diff
changeset
 | 
2196  | 
|
| 22756 | 2197  | 
fun gen_prep_local_registration mk_ctxt = gen_prep_registration I  | 
| 
24787
 
df56433cc059
Theory/context data restructured; simplified interface for printing of interpretations.
 
ballarin 
parents: 
24751 
diff
changeset
 | 
2198  | 
test_local_registration  | 
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2199  | 
local_activate_facts_elemss mk_ctxt;  | 
| 
15624
 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 
ballarin 
parents: 
15623 
diff
changeset
 | 
2200  | 
|
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2201  | 
val prep_global_registration = gen_prep_global_registration  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2202  | 
Attrib.intern_src intern_expr read_instantiations;  | 
| 
24941
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2203  | 
(* FIXME: NEW  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2204  | 
val prep_global_registration_i = gen_prep_global_registration  | 
| 
24952
 
f336c36f41a0
removed redundant strip_vars/abs_eqn, use improved Drule.abs_def instead;
 
wenzelm 
parents: 
24941 
diff
changeset
 | 
2205  | 
(K I) (K I) check_instantiations;  | 
| 
24941
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2206  | 
*)  | 
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2207  | 
val prep_global_registration_i = gen_prep_global_registration  | 
| 22756 | 2208  | 
(K I) (K I) ((K o K) I);  | 
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2209  | 
|
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2210  | 
val prep_local_registration = gen_prep_local_registration  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2211  | 
Attrib.intern_src intern_expr read_instantiations;  | 
| 
24941
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2212  | 
(* FIXME: NEW  | 
| 
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2213  | 
val prep_local_registration_i = gen_prep_local_registration  | 
| 
24952
 
f336c36f41a0
removed redundant strip_vars/abs_eqn, use improved Drule.abs_def instead;
 
wenzelm 
parents: 
24941 
diff
changeset
 | 
2214  | 
(K I) (K I) check_instantiations;  | 
| 
24941
 
9ab032df81c8
Prepare proper interface of interpretation_i, interpret_i.
 
ballarin 
parents: 
24920 
diff
changeset
 | 
2215  | 
*)  | 
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2216  | 
val prep_local_registration_i = gen_prep_local_registration  | 
| 22756 | 2217  | 
(K I) (K I) ((K o K) I);  | 
| 15596 | 2218  | 
|
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
2219  | 
fun prep_registration_in_locale target expr thy =  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
2220  | 
(* target already in internal form *)  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
2221  | 
let  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
2222  | 
val ctxt = ProofContext.init thy;  | 
| 17138 | 2223  | 
val ((raw_target_ids, target_syn), _) = flatten (ctxt, I)  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
2224  | 
(([], Symtab.empty), Expr (Locale target));  | 
| 19278 | 2225  | 
val fixed = the_locale thy target |> #params |> map #1;  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
2226  | 
val ((all_ids, syn), raw_elemss) = flatten (ctxt, intern_expr thy)  | 
| 17138 | 2227  | 
((raw_target_ids, target_syn), Expr expr);  | 
| 
19018
 
88b4979193d8
added abbreviations: activated by init, no expressions yet;
 
wenzelm 
parents: 
18964 
diff
changeset
 | 
2228  | 
val (target_ids, ids) = chop (length raw_target_ids) all_ids;  | 
| 17138 | 2229  | 
val ((parms, elemss, _), _) = read_elemss false ctxt fixed raw_elemss [];  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
2230  | 
|
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
2231  | 
(** compute proof obligations **)  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
2232  | 
|
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
2233  | 
(* restore "small" ids, with mode *)  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
2234  | 
val ids' = map (apsnd snd) ids;  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
2235  | 
(* remove Int markers *)  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
2236  | 
val elemss' = map (fn (_, es) =>  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
2237  | 
map (fn Int e => e) es) elemss  | 
| 
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
2238  | 
(* extract assumptions and defs *)  | 
| 17138 | 2239  | 
val ids_elemss = ids' ~~ elemss';  | 
| 19780 | 2240  | 
val propss = map extract_asms_elems ids_elemss;  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
2241  | 
|
| 17138 | 2242  | 
(** activation function:  | 
2243  | 
- add registrations to the target locale  | 
|
2244  | 
- add induced registrations for all global registrations of  | 
|
2245  | 
the target, unless already present  | 
|
2246  | 
- add facts of induced registrations to theory **)  | 
|
2247  | 
||
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2248  | 
(* val t_ids = map_filter  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2249  | 
(fn (id, (_, Assumed _)) => SOME id | _ => NONE) target_ids; *)  | 
| 17138 | 2250  | 
|
| 
18123
 
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
 
wenzelm 
parents: 
18038 
diff
changeset
 | 
2251  | 
fun activate thmss thy = let  | 
| 19780 | 2252  | 
val satisfy = Element.satisfy_thm (flat thmss);  | 
| 
18123
 
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
 
wenzelm 
parents: 
18038 
diff
changeset
 | 
2253  | 
val ids_elemss_thmss = ids_elemss ~~ thmss;  | 
| 17138 | 2254  | 
val regs = get_global_registrations thy target;  | 
2255  | 
||
2256  | 
fun activate_id (((id, Assumed _), _), thms) thy =  | 
|
| 17033 | 2257  | 
thy |> put_registration_in_locale target id  | 
| 17138 | 2258  | 
|> fold (add_witness_in_locale target id) thms  | 
2259  | 
| activate_id _ thy = thy;  | 
|
2260  | 
||
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22573 
diff
changeset
 | 
2261  | 
fun activate_reg (vts, (((fully_qualified, prfx), atts2), _, _)) thy =  | 
| 
21483
 
e4be91feca50
replaced Args.map_values/Element.map_ctxt_values by general morphism application;
 
wenzelm 
parents: 
21441 
diff
changeset
 | 
2262  | 
let  | 
| 
24638
 
69cb317edf73
Morphisms applied in global interpretations behave correctly on types and terms.
 
ballarin 
parents: 
24525 
diff
changeset
 | 
2263  | 
val (vinsts, insts, wits, _) = collect_global_witnesses thy fixed target_ids vts;  | 
| 18137 | 2264  | 
fun inst_parms ps = map  | 
| 17485 | 2265  | 
(the o AList.lookup (op =) (map #1 fixed ~~ vts)) ps;  | 
| 19780 | 2266  | 
val disch = Element.satisfy_thm wits;  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19423 
diff
changeset
 | 
2267  | 
val new_elemss = filter (fn (((name, ps), _), _) =>  | 
| 17138 | 2268  | 
not (test_global_registration thy (name, inst_parms ps))) (ids_elemss);  | 
2269  | 
fun activate_assumed_id (((_, Derived _), _), _) thy = thy  | 
|
2270  | 
| activate_assumed_id ((((name, ps), Assumed _), _), thms) thy = let  | 
|
2271  | 
val ps' = inst_parms ps;  | 
|
2272  | 
in  | 
|
2273  | 
if test_global_registration thy (name, ps')  | 
|
2274  | 
then thy  | 
|
2275  | 
else thy  | 
|
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
2276  | 
|> put_global_registration (name, ps') ((fully_qualified, prfx), atts2)  | 
| 19780 | 2277  | 
|> fold (fn witn => fn thy => add_global_witness (name, ps')  | 
| 
24638
 
69cb317edf73
Morphisms applied in global interpretations behave correctly on types and terms.
 
ballarin 
parents: 
24525 
diff
changeset
 | 
2278  | 
(Element.morph_witness (Element.inst_morphism' thy vinsts insts) witn) thy) thms  | 
| 17138 | 2279  | 
end;  | 
2280  | 
||
2281  | 
fun activate_derived_id ((_, Assumed _), _) thy = thy  | 
|
2282  | 
| activate_derived_id (((name, ps), Derived ths), _) thy = let  | 
|
2283  | 
val ps' = inst_parms ps;  | 
|
2284  | 
in  | 
|
2285  | 
if test_global_registration thy (name, ps')  | 
|
2286  | 
then thy  | 
|
2287  | 
else thy  | 
|
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
2288  | 
|> put_global_registration (name, ps') ((fully_qualified, prfx), atts2)  | 
| 19780 | 2289  | 
|> fold (fn witn => fn thy => add_global_witness (name, ps')  | 
2290  | 
(witn |> Element.map_witness (fn (t, th) => (* FIXME *)  | 
|
| 18137 | 2291  | 
(Element.inst_term insts t,  | 
| 19780 | 2292  | 
disch (Element.inst_thm thy insts (satisfy th))))) thy) ths  | 
| 17138 | 2293  | 
end;  | 
2294  | 
||
| 21441 | 2295  | 
fun activate_elem (Notes (kind, facts)) thy =  | 
| 17138 | 2296  | 
let  | 
| 21523 | 2297  | 
val att_morphism =  | 
| 
22241
 
5b2dc1b30e7a
interpretation: attempt to be more serious about name_morphism;
 
wenzelm 
parents: 
21962 
diff
changeset
 | 
2298  | 
Morphism.name_morphism (NameSpace.qualified prfx) $>  | 
| 21523 | 2299  | 
Morphism.thm_morphism satisfy $>  | 
2300  | 
Element.inst_morphism thy insts $>  | 
|
2301  | 
Morphism.thm_morphism disch;  | 
|
| 17138 | 2302  | 
val facts' = facts  | 
| 21523 | 2303  | 
|> Attrib.map_facts (Attrib.attribute_i thy o Args.morph_values att_morphism)  | 
2304  | 
|> map (apfst (apsnd (fn a => a @ map (Attrib.attribute thy) atts2)))  | 
|
2305  | 
|> map (apsnd (map (apfst (map (disch o Element.inst_thm thy insts o satisfy)))))  | 
|
| 17138 | 2306  | 
in  | 
| 18377 | 2307  | 
thy  | 
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
2308  | 
|> global_note_prefix_i kind (fully_qualified, prfx) facts'  | 
| 18377 | 2309  | 
|> snd  | 
| 17138 | 2310  | 
end  | 
2311  | 
| activate_elem _ thy = thy;  | 
|
2312  | 
||
2313  | 
fun activate_elems (_, elems) thy = fold activate_elem elems thy;  | 
|
2314  | 
||
2315  | 
in thy |> fold activate_assumed_id ids_elemss_thmss  | 
|
2316  | 
|> fold activate_derived_id ids_elemss  | 
|
2317  | 
|> fold activate_elems new_elemss end;  | 
|
| 17033 | 2318  | 
in  | 
| 17138 | 2319  | 
thy |> fold activate_id ids_elemss_thmss  | 
2320  | 
|> fold activate_reg regs  | 
|
| 17033 | 2321  | 
end;  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
2322  | 
|
| 17033 | 2323  | 
in (propss, activate) end;  | 
| 
17000
 
552df70f52c2
First version of interpretation in locales.  Not yet fully functional.
 
ballarin 
parents: 
16947 
diff
changeset
 | 
2324  | 
|
| 21005 | 2325  | 
fun prep_propp propss = propss |> map (fn (_, props) =>  | 
| 21361 | 2326  | 
map (rpair [] o Element.mark_witness) props);  | 
| 
18123
 
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
 
wenzelm 
parents: 
18038 
diff
changeset
 | 
2327  | 
|
| 
 
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
 
wenzelm 
parents: 
18038 
diff
changeset
 | 
2328  | 
fun prep_result propps thmss =  | 
| 19780 | 2329  | 
ListPair.map (fn ((_, props), thms) => map2 Element.make_witness props thms) (propps, thmss);  | 
| 17437 | 2330  | 
|
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2331  | 
fun gen_interpretation prep_registration after_qed (prfx, raw_atts) raw_expr raw_insts thy =  | 
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
2332  | 
(* prfx = (flag indicating full qualification, name prefix) *)  | 
| 17355 | 2333  | 
let  | 
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2334  | 
val (propss, activate) = prep_registration thy (prfx, raw_atts) raw_expr raw_insts;  | 
| 
20366
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
2335  | 
fun after_qed' results =  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
2336  | 
ProofContext.theory (activate (prep_result propss results))  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
2337  | 
#> after_qed;  | 
| 
18123
 
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
 
wenzelm 
parents: 
18038 
diff
changeset
 | 
2338  | 
in  | 
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2339  | 
thy  | 
| 23418 | 2340  | 
|> ProofContext.init  | 
| 21441 | 2341  | 
|> Proof.theorem_i NONE after_qed' (prep_propp propss)  | 
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2342  | 
|> Element.refine_witness  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2343  | 
|> Seq.hd  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2344  | 
end;  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2345  | 
|
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2346  | 
fun gen_interpret prep_registration after_qed (prfx, atts) expr insts int state =  | 
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
2347  | 
(* prfx = (flag indicating full qualification, name prefix) *)  | 
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2348  | 
let  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2349  | 
val _ = Proof.assert_forward_or_chain state;  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2350  | 
val ctxt = Proof.context_of state;  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2351  | 
val (propss, activate) = prep_registration ctxt (prfx, atts) expr insts;  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2352  | 
fun after_qed' results =  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2353  | 
Proof.map_context (K (ctxt |> activate (prep_result propss results)))  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2354  | 
#> Proof.put_facts NONE  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2355  | 
#> after_qed;  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2356  | 
in  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2357  | 
state  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2358  | 
|> Proof.local_goal (ProofDisplay.print_results int) (K I) ProofContext.bind_propp_i  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2359  | 
      "interpret" NONE after_qed' (map (pair ("", [])) (prep_propp propss))
 | 
| 19810 | 2360  | 
|> Element.refine_witness |> Seq.hd  | 
| 
18123
 
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
 
wenzelm 
parents: 
18038 
diff
changeset
 | 
2361  | 
end;  | 
| 17355 | 2362  | 
|
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2363  | 
in  | 
| 
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2364  | 
|
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
2365  | 
val interpretation_i = gen_interpretation prep_global_registration_i;  | 
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2366  | 
val interpretation = gen_interpretation prep_global_registration;  | 
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
2367  | 
|
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2368  | 
|
| 
20366
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
2369  | 
fun interpretation_in_locale after_qed (raw_target, expr) thy =  | 
| 17355 | 2370  | 
let  | 
2371  | 
val target = intern thy raw_target;  | 
|
2372  | 
val (propss, activate) = prep_registration_in_locale target expr thy;  | 
|
| 21361 | 2373  | 
val raw_propp = prep_propp propss;  | 
| 21005 | 2374  | 
|
2375  | 
val (_, _, goal_ctxt, propp) = thy  | 
|
| 
21665
 
ba07e24dc941
encode declarations as notes (attributes), which enables proper inheritance, interpretation etc.;
 
wenzelm 
parents: 
21602 
diff
changeset
 | 
2376  | 
|> ProofContext.init  | 
| 21361 | 2377  | 
|> cert_context_statement (SOME target) [] raw_propp;  | 
| 21005 | 2378  | 
|
2379  | 
fun after_qed' results =  | 
|
| 
20366
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
2380  | 
ProofContext.theory (activate (prep_result propss results))  | 
| 
 
867696dc64fc
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 
wenzelm 
parents: 
20317 
diff
changeset
 | 
2381  | 
#> after_qed;  | 
| 
18123
 
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
 
wenzelm 
parents: 
18038 
diff
changeset
 | 
2382  | 
in  | 
| 21005 | 2383  | 
goal_ctxt  | 
| 21441 | 2384  | 
|> Proof.theorem_i NONE after_qed' propp  | 
| 19810 | 2385  | 
|> Element.refine_witness |> Seq.hd  | 
| 
18123
 
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
 
wenzelm 
parents: 
18038 
diff
changeset
 | 
2386  | 
end;  | 
| 
17449
 
429ca1e21289
interpretation: use goal commands without target -- no storing of results;
 
wenzelm 
parents: 
17438 
diff
changeset
 | 
2387  | 
|
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22339 
diff
changeset
 | 
2388  | 
val interpret_i = gen_interpret prep_local_registration_i;  | 
| 
22300
 
596f49b70c70
internal interfaces for interpretation and interpret
 
haftmann 
parents: 
22241 
diff
changeset
 | 
2389  | 
val interpret = gen_interpret prep_local_registration;  | 
| 15596 | 2390  | 
|
| 11896 | 2391  | 
end;  | 
| 17355 | 2392  | 
|
2393  | 
end;  |