author | ballarin |
Fri, 07 Jul 2006 09:28:39 +0200 | |
changeset 20035 | 80c79876d2de |
parent 20032 | 2087e5634598 |
child 20037 | d4102c7cf051 |
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 |
17138 | 35 |
- var vs. fixes in locale to be interpreted (interpretation in locale) |
36 |
(implicit locale expressions generated by multiple registrations) |
|
16169
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16144
diff
changeset
|
37 |
*) |
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16144
diff
changeset
|
38 |
|
11896 | 39 |
signature LOCALE = |
40 |
sig |
|
12273 | 41 |
datatype expr = |
42 |
Locale of string | |
|
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
43 |
Rename of expr * (string * mixfix option) option list | |
12273 | 44 |
Merge of expr list |
45 |
val empty: expr |
|
18137 | 46 |
datatype 'a element = Elem of 'a | Expr of expr |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
47 |
|
16458 | 48 |
val intern: theory -> xstring -> string |
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 |
|
303793f49b0f
Interface to access the specification of a named locale.
ballarin
parents:
18782
diff
changeset
|
54 |
val parameters_of: theory -> string -> |
18917 | 55 |
((string * typ) * mixfix) list |
19276 | 56 |
val parameters_of_expr: theory -> expr -> |
57 |
((string * typ) * mixfix) list |
|
18795
303793f49b0f
Interface to access the specification of a named locale.
ballarin
parents:
18782
diff
changeset
|
58 |
val local_asms_of: theory -> string -> |
303793f49b0f
Interface to access the specification of a named locale.
ballarin
parents:
18782
diff
changeset
|
59 |
((string * Attrib.src list) * term list) list |
303793f49b0f
Interface to access the specification of a named locale.
ballarin
parents:
18782
diff
changeset
|
60 |
val global_asms_of: theory -> string -> |
303793f49b0f
Interface to access the specification of a named locale.
ballarin
parents:
18782
diff
changeset
|
61 |
((string * Attrib.src list) * term list) list |
303793f49b0f
Interface to access the specification of a named locale.
ballarin
parents:
18782
diff
changeset
|
62 |
|
18899
a8e913c93578
theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents:
18890
diff
changeset
|
63 |
(* Processing of locale statements *) |
18137 | 64 |
val read_context_statement: xstring option -> Element.context element list -> |
19585 | 65 |
(string * string list) list list -> Proof.context -> |
19991
0c9650664d47
Locales no longer generate views. The following functions have changed
ballarin
parents:
19984
diff
changeset
|
66 |
string option * Proof.context * Proof.context * (term * term list) list list |
18137 | 67 |
val cert_context_statement: string option -> Element.context_i element list -> |
19585 | 68 |
(term * term list) list list -> Proof.context -> |
19991
0c9650664d47
Locales no longer generate views. The following functions have changed
ballarin
parents:
19984
diff
changeset
|
69 |
string option * Proof.context * Proof.context * (term * term list) list list |
19780 | 70 |
val read_expr: expr -> Element.context list -> Proof.context -> |
71 |
Element.context_i list * Proof.context |
|
72 |
val cert_expr: expr -> Element.context_i list -> Proof.context -> |
|
73 |
Element.context_i list * Proof.context |
|
15596 | 74 |
|
75 |
(* Diagnostic functions *) |
|
12758 | 76 |
val print_locales: theory -> unit |
18137 | 77 |
val print_locale: theory -> bool -> expr -> Element.context list -> unit |
17138 | 78 |
val print_global_registrations: bool -> string -> theory -> unit |
18617 | 79 |
val print_local_registrations': bool -> string -> Proof.context -> unit |
80 |
val print_local_registrations: bool -> string -> Proof.context -> unit |
|
18137 | 81 |
|
18917 | 82 |
val add_locale: bool -> bstring -> expr -> Element.context list -> theory |
19293 | 83 |
-> (string * Proof.context (*body context!*)) * theory |
18917 | 84 |
val add_locale_i: bool -> bstring -> expr -> Element.context_i list -> theory |
19293 | 85 |
-> (string * Proof.context (*body context!*)) * theory |
15596 | 86 |
|
15696 | 87 |
(* Storing results *) |
15703 | 88 |
val note_thmss: string -> xstring -> |
18806 | 89 |
((string * Attrib.src list) * (thmref * Attrib.src list) list) list -> |
90 |
theory -> ((string * thm list) list * (string * thm list) list) * (Proof.context * theory) |
|
15703 | 91 |
val note_thmss_i: string -> string -> |
18806 | 92 |
((string * Attrib.src list) * (thm list * Attrib.src list) list) list -> |
93 |
theory -> ((string * thm list) list * (string * thm list) list) * (Proof.context * theory) |
|
94 |
val add_thmss: string -> string -> |
|
95 |
((string * Attrib.src list) * (thm list * Attrib.src list) list) list -> |
|
19991
0c9650664d47
Locales no longer generate views. The following functions have changed
ballarin
parents:
19984
diff
changeset
|
96 |
Proof.context -> |
18806 | 97 |
((string * thm list) list * (string * thm list) list) * Proof.context |
20035
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
98 |
val add_term_syntax: string -> (Proof.context -> Proof.context) -> |
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
99 |
cterm list * Proof.context -> Proof.context |
18137 | 100 |
|
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
101 |
val theorem: string -> Method.text option -> (thm list list -> theory -> theory) -> |
18899
a8e913c93578
theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents:
18890
diff
changeset
|
102 |
string * Attrib.src list -> Element.context element list -> Element.statement -> |
17355 | 103 |
theory -> Proof.state |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
104 |
val theorem_i: string -> Method.text option -> (thm list list -> theory -> theory) -> |
18907 | 105 |
string * Attrib.src list -> Element.context_i element list -> Element.statement_i -> |
17355 | 106 |
theory -> Proof.state |
17856 | 107 |
val theorem_in_locale: string -> Method.text option -> |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
108 |
(thm list list -> thm list list -> theory -> theory) -> |
18899
a8e913c93578
theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents:
18890
diff
changeset
|
109 |
xstring -> string * Attrib.src list -> Element.context element list -> Element.statement -> |
17355 | 110 |
theory -> Proof.state |
17856 | 111 |
val theorem_in_locale_i: string -> Method.text option -> |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
112 |
(thm list list -> thm list list -> theory -> theory) -> |
18137 | 113 |
string -> string * Attrib.src list -> Element.context_i element list -> |
18907 | 114 |
Element.statement_i -> theory -> Proof.state |
17355 | 115 |
val smart_theorem: string -> xstring option -> |
18899
a8e913c93578
theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents:
18890
diff
changeset
|
116 |
string * Attrib.src list -> Element.context element list -> Element.statement -> |
17355 | 117 |
theory -> Proof.state |
118 |
val interpretation: string * Attrib.src list -> expr -> string option list -> |
|
119 |
theory -> Proof.state |
|
120 |
val interpretation_in_locale: xstring * expr -> theory -> Proof.state |
|
121 |
val interpret: string * Attrib.src list -> expr -> string option list -> |
|
122 |
bool -> Proof.state -> Proof.state |
|
11896 | 123 |
end; |
12839 | 124 |
|
12289 | 125 |
structure Locale: LOCALE = |
11896 | 126 |
struct |
127 |
||
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
128 |
fun legacy_unvarifyT thm = |
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
129 |
let |
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
130 |
val cT = Thm.ctyp_of (Thm.theory_of_thm thm); |
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
131 |
val tfrees = map (fn ((x, _), S) => SOME (cT (TFree (x, S)))) (Drule.tvars_of thm); |
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
132 |
in Drule.instantiate' tfrees [] thm end; |
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
133 |
|
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
134 |
fun legacy_unvarify raw_thm = |
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
135 |
let |
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
136 |
val thm = legacy_unvarifyT raw_thm; |
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
137 |
val ct = Thm.cterm_of (Thm.theory_of_thm thm); |
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
138 |
val frees = map (fn ((x, _), T) => SOME (ct (Free (x, T)))) (Drule.vars_of thm); |
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
139 |
in Drule.instantiate' [] frees thm end; |
19780 | 140 |
|
12273 | 141 |
(** locale elements and expressions **) |
11896 | 142 |
|
18137 | 143 |
datatype ctxt = datatype Element.ctxt; |
17355 | 144 |
|
12273 | 145 |
datatype expr = |
146 |
Locale of string | |
|
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
147 |
Rename of expr * (string * mixfix option) option list | |
12273 | 148 |
Merge of expr list; |
11896 | 149 |
|
12273 | 150 |
val empty = Merge []; |
151 |
||
18137 | 152 |
datatype 'a element = |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
153 |
Elem of 'a | Expr of expr; |
12273 | 154 |
|
18137 | 155 |
fun map_elem f (Elem e) = Elem (f e) |
156 |
| map_elem _ (Expr e) = Expr e; |
|
157 |
||
12070 | 158 |
type locale = |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
159 |
{axiom: Element.witness list, |
19942
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
160 |
(* For locales that define predicates this is [A [A]], where A is the locale |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
161 |
specification. Otherwise []. *) |
19018
88b4979193d8
added abbreviations: activated by init, no expressions yet;
wenzelm
parents:
18964
diff
changeset
|
162 |
import: expr, (*dynamic import*) |
19783 | 163 |
elems: (Element.context_i * stamp) list, |
164 |
(* Static content, neither Fixes nor Constrains elements *) |
|
19278 | 165 |
params: ((string * typ) * mixfix) list, (*all params*) |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
166 |
lparams: string list, (*local params*) |
19662
2f5d076fde32
replaced abbrevs by term_syntax, which is both simpler and more general;
wenzelm
parents:
19585
diff
changeset
|
167 |
term_syntax: ((Proof.context -> Proof.context) * stamp) list, (* FIXME depend on morphism *) |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
168 |
regs: ((string * string list) * Element.witness list) list, |
19780 | 169 |
(* 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
|
170 |
intros: thm list * thm list} |
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
171 |
(* Introduction rules: of delta predicate and locale predicate. *) |
11896 | 172 |
|
15703 | 173 |
(* CB: an internal (Int) locale element was either imported or included, |
174 |
an external (Ext) element appears directly in the locale text. *) |
|
175 |
||
176 |
datatype ('a, 'b) int_ext = Int of 'a | Ext of 'b; |
|
177 |
||
178 |
||
179 |
||
16736
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
180 |
(** management of registrations in theories and proof contexts **) |
11896 | 181 |
|
15837 | 182 |
structure Registrations : |
183 |
sig |
|
184 |
type T |
|
185 |
val empty: T |
|
186 |
val join: T * T -> T |
|
19780 | 187 |
val dest: T -> (term list * ((string * Attrib.src list) * Element.witness list)) list |
16458 | 188 |
val lookup: theory -> T * term list -> |
19780 | 189 |
((string * Attrib.src list) * Element.witness list) option |
16458 | 190 |
val insert: theory -> term list * (string * Attrib.src list) -> T -> |
19780 | 191 |
T * (term list * ((string * Attrib.src list) * Element.witness list)) list |
192 |
val add_witness: term list -> Element.witness -> T -> T |
|
15837 | 193 |
end = |
194 |
struct |
|
195 |
(* a registration consists of theorems instantiating locale assumptions |
|
196 |
and prefix and attributes, indexed by parameter instantiation *) |
|
19780 | 197 |
type T = ((string * Attrib.src list) * Element.witness list) Termtab.table; |
15837 | 198 |
|
199 |
val empty = Termtab.empty; |
|
200 |
||
201 |
(* term list represented as single term, for simultaneous matching *) |
|
202 |
fun termify ts = |
|
18343 | 203 |
Term.list_comb (Const ("", map fastype_of ts ---> propT), ts); |
15837 | 204 |
fun untermify t = |
205 |
let fun ut (Const _) ts = ts |
|
206 |
| ut (s $ t) ts = ut s (t::ts) |
|
207 |
in ut t [] end; |
|
208 |
||
16620
2a7f46324218
Proper treatment of beta-redexes in witness theorems.
ballarin
parents:
16458
diff
changeset
|
209 |
(* joining of registrations: prefix and attributes of left theory, |
15837 | 210 |
thms are equal, no attempt to subsumption testing *) |
19025 | 211 |
fun join (r1, r2) = Termtab.join (fn _ => fn (reg, _) => reg) (r1, r2); |
15837 | 212 |
|
213 |
fun dest regs = map (apfst untermify) (Termtab.dest regs); |
|
214 |
||
215 |
(* registrations that subsume t *) |
|
17203 | 216 |
fun subsumers thy t regs = |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19423
diff
changeset
|
217 |
filter (fn (t', _) => Pattern.matches thy (t', t)) (Termtab.dest regs); |
15837 | 218 |
|
219 |
(* look up registration, pick one that subsumes the query *) |
|
19780 | 220 |
fun lookup thy (regs, ts) = |
15837 | 221 |
let |
222 |
val t = termify ts; |
|
19780 | 223 |
val subs = subsumers thy t regs; |
15837 | 224 |
in (case subs of |
225 |
[] => NONE |
|
226 |
| ((t', (attn, thms)) :: _) => let |
|
19780 | 227 |
val (tinst, inst) = Pattern.match thy (t', t) (Vartab.empty, Vartab.empty); |
15837 | 228 |
(* thms contain Frees, not Vars *) |
229 |
val tinst' = tinst |> Vartab.dest |
|
19810 | 230 |
|> map (fn ((x, 0), (_, T)) => (x, Logic.legacy_unvarifyT T)) |
15837 | 231 |
|> Symtab.make; |
232 |
val inst' = inst |> Vartab.dest |
|
19810 | 233 |
|> map (fn ((x, 0), (_, t)) => (x, Logic.legacy_unvarify t)) |
15837 | 234 |
|> Symtab.make; |
235 |
in |
|
19780 | 236 |
SOME (attn, map (Element.inst_witness thy (tinst', inst')) thms) |
15837 | 237 |
end) |
238 |
end; |
|
239 |
||
240 |
(* add registration if not subsumed by ones already present, |
|
241 |
additionally returns registrations that are strictly subsumed *) |
|
242 |
fun insert sign (ts, attn) regs = |
|
243 |
let |
|
244 |
val t = termify ts; |
|
17203 | 245 |
val subs = subsumers sign t regs ; |
15837 | 246 |
in (case subs of |
247 |
[] => let |
|
248 |
val sups = |
|
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19423
diff
changeset
|
249 |
filter (fn (t', _) => Pattern.matches sign (t, t')) (Termtab.dest regs); |
15837 | 250 |
val sups' = map (apfst untermify) sups |
17412 | 251 |
in (Termtab.update (t, (attn, [])) regs, sups') end |
15837 | 252 |
| _ => (regs, [])) |
253 |
end; |
|
254 |
||
255 |
(* add witness theorem to registration, |
|
16169
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16144
diff
changeset
|
256 |
only if instantiation is exact, otherwise exception Option raised *) |
15837 | 257 |
fun add_witness ts thm regs = |
258 |
let |
|
259 |
val t = termify ts; |
|
18343 | 260 |
val (x, thms) = the (Termtab.lookup regs t); |
15837 | 261 |
in |
17412 | 262 |
Termtab.update (t, (x, thm::thms)) regs |
15837 | 263 |
end; |
264 |
end; |
|
265 |
||
16736
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
266 |
|
15837 | 267 |
(** theory data **) |
15596 | 268 |
|
16458 | 269 |
structure GlobalLocalesData = TheoryDataFun |
270 |
(struct |
|
12014 | 271 |
val name = "Isar/locales"; |
15837 | 272 |
type T = NameSpace.T * locale Symtab.table * Registrations.T Symtab.table; |
15596 | 273 |
(* 1st entry: locale namespace, |
274 |
2nd entry: locales of the theory, |
|
15837 | 275 |
3rd entry: registrations, indexed by locale name *) |
11896 | 276 |
|
15596 | 277 |
val empty = (NameSpace.empty, Symtab.empty, Symtab.empty); |
12063 | 278 |
val copy = I; |
16458 | 279 |
val extend = I; |
12289 | 280 |
|
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
281 |
fun join_locs _ ({axiom, import, elems, params, lparams, term_syntax, regs, intros}: locale, |
19662
2f5d076fde32
replaced abbrevs by term_syntax, which is both simpler and more general;
wenzelm
parents:
19585
diff
changeset
|
282 |
{elems = elems', term_syntax = term_syntax', regs = regs', ...}: locale) = |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
283 |
{axiom = axiom, |
19018
88b4979193d8
added abbreviations: activated by init, no expressions yet;
wenzelm
parents:
18964
diff
changeset
|
284 |
import = import, |
17496 | 285 |
elems = gen_merge_lists (eq_snd (op =)) elems elems', |
16736
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
286 |
params = params, |
19278 | 287 |
lparams = lparams, |
19662
2f5d076fde32
replaced abbrevs by term_syntax, which is both simpler and more general;
wenzelm
parents:
19585
diff
changeset
|
288 |
term_syntax = Library.merge (eq_snd (op =)) (term_syntax, term_syntax'), |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
289 |
regs = merge_alists regs regs', |
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
290 |
intros = intros}; |
16458 | 291 |
fun merge _ ((space1, locs1, regs1), (space2, locs2, regs2)) = |
15596 | 292 |
(NameSpace.merge (space1, space2), Symtab.join join_locs (locs1, locs2), |
19025 | 293 |
Symtab.join (K Registrations.join) (regs1, regs2)); |
12289 | 294 |
|
15596 | 295 |
fun print _ (space, locs, _) = |
16346 | 296 |
Pretty.strs ("locales:" :: map #1 (NameSpace.extern_table (space, locs))) |
12014 | 297 |
|> Pretty.writeln; |
16458 | 298 |
end); |
11896 | 299 |
|
18708 | 300 |
val _ = Context.add_setup GlobalLocalesData.init; |
15801 | 301 |
|
302 |
||
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
303 |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
304 |
(** context data **) |
11896 | 305 |
|
16458 | 306 |
structure LocalLocalesData = ProofDataFun |
307 |
(struct |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
308 |
val name = "Isar/locales"; |
15837 | 309 |
type T = Registrations.T Symtab.table; |
310 |
(* registrations, indexed by locale name *) |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
311 |
fun init _ = Symtab.empty; |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
312 |
fun print _ _ = (); |
16458 | 313 |
end); |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
314 |
|
18708 | 315 |
val _ = Context.add_setup LocalLocalesData.init; |
12289 | 316 |
|
12277 | 317 |
|
318 |
(* access locales *) |
|
319 |
||
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
320 |
val print_locales = GlobalLocalesData.print; |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
321 |
|
16458 | 322 |
val intern = NameSpace.intern o #1 o GlobalLocalesData.get; |
323 |
val extern = NameSpace.extern o #1 o GlobalLocalesData.get; |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
324 |
|
16144 | 325 |
fun declare_locale name thy = |
326 |
thy |> GlobalLocalesData.map (fn (space, locs, regs) => |
|
16458 | 327 |
(Sign.declare_name thy name space, locs, regs)); |
11896 | 328 |
|
15596 | 329 |
fun put_locale name loc = |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
330 |
GlobalLocalesData.map (fn (space, locs, regs) => |
17412 | 331 |
(space, Symtab.update (name, loc) locs, regs)); |
332 |
||
333 |
fun get_locale thy name = Symtab.lookup (#2 (GlobalLocalesData.get thy)) name; |
|
11896 | 334 |
|
12014 | 335 |
fun the_locale thy name = |
336 |
(case get_locale thy name of |
|
15531 | 337 |
SOME loc => loc |
338 |
| NONE => error ("Unknown locale " ^ quote name)); |
|
11896 | 339 |
|
18806 | 340 |
fun change_locale name f thy = |
341 |
let |
|
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
342 |
val {axiom, import, elems, params, lparams, term_syntax, regs, intros} = |
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
343 |
the_locale thy name; |
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
344 |
val (axiom', import', elems', params', lparams', term_syntax', regs', intros') = |
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
345 |
f (axiom, import, elems, params, lparams, term_syntax, regs, intros); |
18806 | 346 |
in |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
347 |
put_locale name {axiom = axiom', import = import', elems = elems', |
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
348 |
params = params', lparams = lparams', term_syntax = term_syntax', regs = regs', |
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
349 |
intros = intros'} thy |
18806 | 350 |
end; |
351 |
||
12046 | 352 |
|
15596 | 353 |
(* access registrations *) |
354 |
||
15696 | 355 |
(* Ids of global registrations are varified, |
356 |
Ids of local registrations are not. |
|
357 |
Thms of registrations are never varified. *) |
|
358 |
||
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
359 |
(* retrieve registration from theory or context *) |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
360 |
|
15696 | 361 |
fun gen_get_registrations get thy_ctxt name = |
17412 | 362 |
case Symtab.lookup (get thy_ctxt) name of |
15696 | 363 |
NONE => [] |
15837 | 364 |
| SOME reg => Registrations.dest reg; |
15696 | 365 |
|
366 |
val get_global_registrations = |
|
367 |
gen_get_registrations (#3 o GlobalLocalesData.get); |
|
368 |
val get_local_registrations = |
|
369 |
gen_get_registrations LocalLocalesData.get; |
|
370 |
||
16458 | 371 |
fun gen_get_registration get thy_of thy_ctxt (name, ps) = |
17412 | 372 |
case Symtab.lookup (get thy_ctxt) name of |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
373 |
NONE => NONE |
16458 | 374 |
| SOME reg => Registrations.lookup (thy_of thy_ctxt) (reg, ps); |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
375 |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
376 |
val get_global_registration = |
16458 | 377 |
gen_get_registration (#3 o GlobalLocalesData.get) I; |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
378 |
val get_local_registration = |
16458 | 379 |
gen_get_registration LocalLocalesData.get ProofContext.theory_of; |
15596 | 380 |
|
18343 | 381 |
val test_global_registration = is_some oo get_global_registration; |
382 |
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
|
383 |
fun smart_test_registration ctxt id = |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
384 |
let |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
385 |
val thy = ProofContext.theory_of ctxt; |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
386 |
in |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
387 |
test_global_registration thy id orelse test_local_registration ctxt id |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
388 |
end; |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
389 |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
390 |
|
15837 | 391 |
(* 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
|
392 |
|
16458 | 393 |
fun gen_put_registration map_data thy_of (name, ps) attn thy_ctxt = |
15837 | 394 |
map_data (fn regs => |
395 |
let |
|
16458 | 396 |
val thy = thy_of thy_ctxt; |
18343 | 397 |
val reg = the_default Registrations.empty (Symtab.lookup regs name); |
16458 | 398 |
val (reg', sups) = Registrations.insert thy (ps, attn) reg; |
15837 | 399 |
val _ = if not (null sups) then warning |
400 |
("Subsumed interpretation(s) of locale " ^ |
|
16458 | 401 |
quote (extern thy name) ^ |
15837 | 402 |
"\nby interpretation(s) with the following prefix(es):\n" ^ |
403 |
commas_quote (map (fn (_, ((s, _), _)) => s) sups)) |
|
404 |
else (); |
|
17412 | 405 |
in Symtab.update (name, reg') regs end) thy_ctxt; |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
406 |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
407 |
val put_global_registration = |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
408 |
gen_put_registration (fn f => |
16458 | 409 |
GlobalLocalesData.map (fn (space, locs, regs) => (space, locs, f regs))) I; |
15837 | 410 |
val put_local_registration = |
16458 | 411 |
gen_put_registration LocalLocalesData.map ProofContext.theory_of; |
15596 | 412 |
|
18806 | 413 |
fun put_registration_in_locale name id = |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
414 |
change_locale name (fn (axiom, import, elems, params, lparams, term_syntax, regs, intros) => |
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
415 |
(axiom, import, elems, params, lparams, term_syntax, regs @ [(id, [])], intros)); |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
416 |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
417 |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
418 |
(* add witness theorem to registration in theory or context, |
15596 | 419 |
ignored if registration not present *) |
420 |
||
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
421 |
fun add_witness (name, ps) thm = |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
422 |
Symtab.map_entry name (Registrations.add_witness ps thm); |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
423 |
|
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
424 |
fun add_global_witness id thm = |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
425 |
GlobalLocalesData.map (fn (space, locs, regs) => (space, locs, add_witness id thm regs)); |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
426 |
|
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
427 |
val add_local_witness = LocalLocalesData.map oo add_witness; |
15596 | 428 |
|
18806 | 429 |
fun add_witness_in_locale name id thm = |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
430 |
change_locale name (fn (axiom, import, elems, params, lparams, term_syntax, regs, intros) => |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
431 |
let |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
432 |
fun add (id', thms) = |
18806 | 433 |
if id = id' then (id', thm :: thms) else (id', thms); |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
434 |
in (axiom, import, elems, params, lparams, term_syntax, map add regs, intros) end); |
15596 | 435 |
|
14215
ebf291f3b449
Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents:
13629
diff
changeset
|
436 |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
437 |
(* printing of registrations *) |
15596 | 438 |
|
17138 | 439 |
fun gen_print_registrations get_regs mk_ctxt msg show_wits loc thy_ctxt = |
15596 | 440 |
let |
15703 | 441 |
val ctxt = mk_ctxt thy_ctxt; |
442 |
val thy = ProofContext.theory_of ctxt; |
|
443 |
||
444 |
val prt_term = Pretty.quote o ProofContext.pretty_term ctxt; |
|
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
445 |
fun prt_inst ts = |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
446 |
Pretty.enclose "(" ")" (Pretty.breaks (map prt_term ts)); |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
447 |
fun prt_attn (prfx, atts) = |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
448 |
Pretty.breaks (Pretty.str prfx :: Args.pretty_attribs ctxt atts); |
19780 | 449 |
fun prt_witns witns = Pretty.enclose "[" "]" |
450 |
(Pretty.breaks (map (prt_term o Element.witness_prop) witns)); |
|
451 |
fun prt_reg (ts, (("", []), witns)) = |
|
17138 | 452 |
if show_wits |
19780 | 453 |
then Pretty.block [prt_inst ts, Pretty.fbrk, prt_witns witns] |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
454 |
else prt_inst ts |
19780 | 455 |
| prt_reg (ts, (attn, witns)) = |
17138 | 456 |
if show_wits |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
457 |
then Pretty.block ((prt_attn attn @ |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
458 |
[Pretty.str ":", Pretty.brk 1, prt_inst ts, Pretty.fbrk, |
19780 | 459 |
prt_witns witns])) |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
460 |
else Pretty.block ((prt_attn attn @ |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
461 |
[Pretty.str ":", Pretty.brk 1, prt_inst ts])); |
15703 | 462 |
|
16458 | 463 |
val loc_int = intern thy loc; |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
464 |
val regs = get_regs thy_ctxt; |
17412 | 465 |
val loc_regs = Symtab.lookup regs loc_int; |
15596 | 466 |
in |
467 |
(case loc_regs of |
|
17355 | 468 |
NONE => Pretty.str ("no interpretations in " ^ msg) |
15763
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
ballarin
parents:
15749
diff
changeset
|
469 |
| SOME r => let |
15837 | 470 |
val r' = Registrations.dest r; |
15763
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
ballarin
parents:
15749
diff
changeset
|
471 |
val r'' = Library.sort_wrt (fn (_, ((prfx, _), _)) => prfx) r'; |
17355 | 472 |
in Pretty.big_list ("interpretations in " ^ msg ^ ":") |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
473 |
(map prt_reg r'') |
15763
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
ballarin
parents:
15749
diff
changeset
|
474 |
end) |
15596 | 475 |
|> Pretty.writeln |
476 |
end; |
|
477 |
||
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
478 |
val print_global_registrations = |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
479 |
gen_print_registrations (#3 o GlobalLocalesData.get) |
15703 | 480 |
ProofContext.init "theory"; |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
481 |
val print_local_registrations' = |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
482 |
gen_print_registrations LocalLocalesData.get |
15703 | 483 |
I "context"; |
17138 | 484 |
fun print_local_registrations show_wits loc ctxt = |
485 |
(print_global_registrations show_wits loc (ProofContext.theory_of ctxt); |
|
486 |
print_local_registrations' show_wits loc ctxt); |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
487 |
|
15596 | 488 |
|
12277 | 489 |
(* diagnostics *) |
12273 | 490 |
|
12277 | 491 |
fun err_in_locale ctxt msg ids = |
492 |
let |
|
16458 | 493 |
val thy = ProofContext.theory_of ctxt; |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
494 |
fun prt_id (name, parms) = |
16458 | 495 |
[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
|
496 |
val prt_ids = flat (separate [Pretty.str " +", Pretty.brk 1] (map prt_id ids)); |
12502 | 497 |
val err_msg = |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
498 |
if forall (equal "" o #1) ids then msg |
12502 | 499 |
else msg ^ "\n" ^ Pretty.string_of (Pretty.block |
500 |
(Pretty.str "The error(s) above occurred in locale:" :: Pretty.brk 1 :: prt_ids)); |
|
18678 | 501 |
in error err_msg end; |
12063 | 502 |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
503 |
fun err_in_locale' ctxt msg ids' = err_in_locale ctxt msg (map fst ids'); |
12277 | 504 |
|
505 |
||
19783 | 506 |
fun pretty_ren NONE = Pretty.str "_" |
507 |
| pretty_ren (SOME (x, NONE)) = Pretty.str x |
|
508 |
| pretty_ren (SOME (x, SOME syn)) = |
|
509 |
Pretty.block [Pretty.str x, Pretty.brk 1, Syntax.pretty_mixfix syn]; |
|
510 |
||
511 |
fun pretty_expr thy (Locale name) = Pretty.str (extern thy name) |
|
512 |
| pretty_expr thy (Rename (expr, xs)) = |
|
513 |
Pretty.block [pretty_expr thy expr, Pretty.brk 1, Pretty.block (map pretty_ren xs |> Pretty.breaks)] |
|
514 |
| pretty_expr thy (Merge es) = |
|
515 |
Pretty.separate "+" (map (pretty_expr thy) es) |> Pretty.block; |
|
516 |
||
517 |
fun err_in_expr _ msg (Merge []) = error msg |
|
518 |
| err_in_expr ctxt msg expr = |
|
519 |
error (msg ^ "\n" ^ Pretty.string_of (Pretty.block |
|
520 |
[Pretty.str "The error(s) above occured in locale expression:", Pretty.brk 1, |
|
521 |
pretty_expr (ProofContext.theory_of ctxt) expr])); |
|
522 |
||
12046 | 523 |
|
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
524 |
(** structured contexts: rename + merge + implicit type instantiation **) |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
525 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
526 |
(* parameter types *) |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
527 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
528 |
fun frozen_tvars ctxt Ts = |
19914 | 529 |
#1 (Variable.importT_inst (map Logic.mk_type Ts) ctxt) |
19900
21a99d88d925
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19873
diff
changeset
|
530 |
|> map (fn ((xi, S), T) => (xi, (S, T))); |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
531 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
532 |
fun unify_frozen ctxt maxidx Ts Us = |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
533 |
let |
18343 | 534 |
fun paramify NONE i = (NONE, i) |
535 |
| 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
|
536 |
|
18343 | 537 |
val (Ts', maxidx') = fold_map paramify Ts maxidx; |
538 |
val (Us', maxidx'') = fold_map paramify Us maxidx'; |
|
16947 | 539 |
val thy = ProofContext.theory_of ctxt; |
14215
ebf291f3b449
Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents:
13629
diff
changeset
|
540 |
|
18190 | 541 |
fun unify (SOME T, SOME U) env = (Sign.typ_unify thy (U, T) env |
542 |
handle Type.TUNIFY => raise TYPE ("unify_frozen: failed to unify types", [U, T], [])) |
|
543 |
| unify _ env = env; |
|
544 |
val (unifier, _) = fold unify (Ts' ~~ Us') (Vartab.empty, maxidx''); |
|
15570 | 545 |
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
|
546 |
val unifier' = Vartab.extend (unifier, frozen_tvars ctxt (map_filter I Vs)); |
15570 | 547 |
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
|
548 |
|
19932 | 549 |
fun params_of elemss = distinct (eq_fst (op = : string * string -> bool)) (maps (snd o fst) elemss); |
550 |
fun params_of' elemss = distinct (eq_fst (op = : string * string -> bool)) (maps (snd o fst o fst) elemss); |
|
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
551 |
|
14508
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
552 |
|
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
553 |
(* CB: param_types has the following type: |
15531 | 554 |
('a * 'b option) list -> ('a * 'b) list *) |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19423
diff
changeset
|
555 |
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
|
556 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
557 |
|
19932 | 558 |
fun merge_syntax ctxt ids ss = Symtab.merge (op = : mixfix * mixfix -> bool) ss |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
559 |
handle Symtab.DUPS xs => err_in_locale ctxt |
16105 | 560 |
("Conflicting syntax for parameter(s): " ^ commas_quote xs) (map fst ids); |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
561 |
|
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
562 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
563 |
(* Distinction of assumed vs. derived identifiers. |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
564 |
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
|
565 |
assumptions of the specification fragment (for locales with |
19780 | 566 |
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
|
567 |
specification fragment to assumptions of other (assumed) specification |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
568 |
fragments. *) |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
569 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
570 |
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
|
571 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
572 |
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
|
573 |
| map_mode f (Derived x) = Derived (f x); |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
574 |
|
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
575 |
|
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
576 |
(* flatten expressions *) |
11896 | 577 |
|
12510 | 578 |
local |
12502 | 579 |
|
18137 | 580 |
fun unify_parms ctxt fixed_parms raw_parmss = |
12502 | 581 |
let |
16458 | 582 |
val thy = ProofContext.theory_of ctxt; |
12502 | 583 |
val maxidx = length raw_parmss; |
584 |
val idx_parmss = (0 upto maxidx - 1) ~~ raw_parmss; |
|
585 |
||
586 |
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
|
587 |
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
|
588 |
val parms = fixed_parms @ maps varify_parms idx_parmss; |
12502 | 589 |
|
18137 | 590 |
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
|
591 |
handle Type.TUNIFY => |
18137 | 592 |
let val prt = Sign.string_of_typ thy in |
593 |
raise TYPE ("unify_parms: failed to unify types " ^ |
|
594 |
prt U ^ " and " ^ prt T, [U, T], []) |
|
595 |
end; |
|
596 |
fun unify_list (T :: Us) = fold (unify T) Us |
|
597 |
| unify_list [] = I; |
|
18952 | 598 |
val (unifier, _) = fold unify_list (map #2 (Symtab.dest (Symtab.make_list parms))) |
18137 | 599 |
(Vartab.empty, maxidx); |
12502 | 600 |
|
19061
ffbbac0261c9
removed distinct, renamed gen_distinct to distinct;
wenzelm
parents:
19025
diff
changeset
|
601 |
val parms' = map (apsnd (Envir.norm_type unifier)) (distinct (eq_fst (op =)) parms); |
12502 | 602 |
val unifier' = Vartab.extend (unifier, frozen_tvars ctxt (map #2 parms')); |
603 |
||
604 |
fun inst_parms (i, ps) = |
|
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19423
diff
changeset
|
605 |
foldr Term.add_typ_tfrees [] (map_filter snd ps) |
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19423
diff
changeset
|
606 |
|> map_filter (fn (a, S) => |
12502 | 607 |
let val T = Envir.norm_type unifier' (TVar ((a, i), S)) |
18137 | 608 |
in if T = TFree (a, S) then NONE else SOME (a, T) end) |
609 |
|> Symtab.make; |
|
610 |
in map inst_parms idx_parmss end; |
|
12502 | 611 |
|
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
612 |
in |
12502 | 613 |
|
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
614 |
fun unify_elemss _ _ [] = [] |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
615 |
| unify_elemss _ [] [elems] = [elems] |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
616 |
| unify_elemss ctxt fixed_parms elemss = |
12502 | 617 |
let |
18137 | 618 |
val thy = ProofContext.theory_of ctxt; |
17756 | 619 |
val envs = unify_parms ctxt fixed_parms (map (snd o fst o fst) elemss); |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
620 |
fun inst ((((name, ps), mode), elems), env) = |
18137 | 621 |
(((name, map (apsnd (Option.map (Element.instT_type env))) ps), |
19780 | 622 |
map_mode (map (Element.instT_witness thy env)) mode), |
18137 | 623 |
map (Element.instT_ctxt thy env) elems); |
12839 | 624 |
in map inst (elemss ~~ envs) end; |
12502 | 625 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
626 |
|
19810 | 627 |
fun renaming xs parms = zip_options parms xs |
628 |
handle Library.UnequalLengths => |
|
629 |
error ("Too many arguments in renaming: " ^ |
|
630 |
commas (map (fn NONE => "_" | SOME x => quote (fst x)) xs)); |
|
631 |
||
632 |
||
19783 | 633 |
(* params_of_expr: |
634 |
Compute parameters (with types and syntax) of locale expression. |
|
635 |
*) |
|
636 |
||
637 |
fun params_of_expr ctxt fixed_params expr (prev_parms, prev_types, prev_syn) = |
|
638 |
let |
|
639 |
val thy = ProofContext.theory_of ctxt; |
|
640 |
||
641 |
fun merge_tenvs fixed tenv1 tenv2 = |
|
642 |
let |
|
643 |
val [env1, env2] = unify_parms ctxt fixed |
|
644 |
[tenv1 |> Symtab.dest |> map (apsnd SOME), |
|
645 |
tenv2 |> Symtab.dest |> map (apsnd SOME)] |
|
646 |
in |
|
647 |
Symtab.merge (op =) (Symtab.map (Element.instT_type env1) tenv1, |
|
648 |
Symtab.map (Element.instT_type env2) tenv2) |
|
649 |
end; |
|
650 |
||
651 |
fun merge_syn expr syn1 syn2 = |
|
19932 | 652 |
Symtab.merge (op = : mixfix * mixfix -> bool) (syn1, syn2) |
19783 | 653 |
handle Symtab.DUPS xs => err_in_expr ctxt |
654 |
("Conflicting syntax for parameter(s): " ^ commas_quote xs) expr; |
|
655 |
||
656 |
fun params_of (expr as Locale name) = |
|
657 |
let |
|
658 |
val {import, params, ...} = the_locale thy name; |
|
659 |
val parms = map (fst o fst) params; |
|
660 |
val (parms', types', syn') = params_of import; |
|
661 |
val all_parms = merge_lists parms' parms; |
|
662 |
val all_types = merge_tenvs [] types' (params |> map fst |> Symtab.make); |
|
663 |
val all_syn = merge_syn expr syn' (params |> map (apfst fst) |> Symtab.make); |
|
664 |
in (all_parms, all_types, all_syn) end |
|
665 |
| params_of (expr as Rename (e, xs)) = |
|
666 |
let |
|
667 |
val (parms', types', syn') = params_of e; |
|
668 |
val ren = renaming xs parms'; |
|
669 |
(* renaming may reduce number of parameters *) |
|
670 |
val new_parms = map (Element.rename ren) parms' |> distinct (op =); |
|
671 |
val ren_syn = syn' |> Symtab.dest |> map (Element.rename_var ren); |
|
672 |
val new_syn = fold (Symtab.insert (op =)) ren_syn Symtab.empty |
|
673 |
handle Symtab.DUP x => |
|
674 |
err_in_expr ctxt ("Conflicting syntax for parameter: " ^ quote x) expr; |
|
675 |
val syn_types = map (apsnd (fn mx => SOME (Type.freeze_type (#1 (TypeInfer.paramify_dummies (TypeInfer.mixfixT mx) 0))))) (Symtab.dest new_syn); |
|
676 |
val ren_types = types' |> Symtab.dest |> map (apfst (Element.rename ren)); |
|
677 |
val (env :: _) = unify_parms ctxt [] |
|
678 |
((ren_types |> map (apsnd SOME)) :: map single syn_types); |
|
679 |
val new_types = fold (Symtab.insert (op =)) |
|
680 |
(map (apsnd (Element.instT_type env)) ren_types) Symtab.empty; |
|
681 |
in (new_parms, new_types, new_syn) end |
|
682 |
| params_of (Merge es) = |
|
683 |
fold (fn e => fn (parms, types, syn) => |
|
684 |
let |
|
685 |
val (parms', types', syn') = params_of e |
|
686 |
in |
|
687 |
(merge_lists parms parms', merge_tenvs [] types types', |
|
688 |
merge_syn e syn syn') |
|
689 |
end) |
|
690 |
es ([], Symtab.empty, Symtab.empty) |
|
691 |
||
692 |
val (parms, types, syn) = params_of expr; |
|
693 |
in |
|
694 |
(merge_lists prev_parms parms, merge_tenvs fixed_params prev_types types, |
|
695 |
merge_syn expr prev_syn syn) |
|
696 |
end; |
|
697 |
||
698 |
fun make_params_ids params = [(("", params), ([], Assumed []))]; |
|
699 |
fun make_raw_params_elemss (params, tenv, syn) = |
|
700 |
[((("", map (fn p => (p, Symtab.lookup tenv p)) params), Assumed []), |
|
701 |
Int [Fixes (map (fn p => |
|
702 |
(p, Symtab.lookup tenv p, Symtab.lookup syn p |> the)) params)])]; |
|
703 |
||
704 |
||
15596 | 705 |
(* flatten_expr: |
706 |
Extend list of identifiers by those new in locale expression expr. |
|
707 |
Compute corresponding list of lists of locale elements (one entry per |
|
708 |
identifier). |
|
709 |
||
710 |
Identifiers represent locale fragments and are in an extended form: |
|
711 |
((name, ps), (ax_ps, axs)) |
|
712 |
(name, ps) is the locale name with all its parameters. |
|
713 |
(ax_ps, axs) is the locale axioms with its parameters; |
|
714 |
axs are always taken from the top level of the locale hierarchy, |
|
715 |
hence axioms may contain additional parameters from later fragments: |
|
716 |
ps subset of ax_ps. axs is either singleton or empty. |
|
717 |
||
718 |
Elements are enriched by identifier-like information: |
|
719 |
(((name, ax_ps), axs), elems) |
|
720 |
The parameters in ax_ps are the axiom parameters, but enriched by type |
|
721 |
info: now each entry is a pair of string and typ option. Axioms are |
|
722 |
type-instantiated. |
|
723 |
||
724 |
*) |
|
725 |
||
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
726 |
fun flatten_expr ctxt ((prev_idents, prev_syntax), expr) = |
12014 | 727 |
let |
728 |
val thy = ProofContext.theory_of ctxt; |
|
12263 | 729 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
730 |
fun rename_parms top ren ((name, ps), (parms, mode)) = |
19783 | 731 |
((name, map (Element.rename ren) ps), |
732 |
if top |
|
733 |
then (map (Element.rename ren) parms, |
|
734 |
map_mode (map (Element.rename_witness ren)) mode) |
|
735 |
else (parms, mode)); |
|
12263 | 736 |
|
19780 | 737 |
(* add registrations of (name, ps), recursively; adjust hyps of witnesses *) |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
738 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
739 |
fun add_regs (name, ps) (ths, ids) = |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
740 |
let |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
741 |
val {params, regs, ...} = the_locale thy name; |
19278 | 742 |
val ps' = map #1 params; |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
743 |
val ren = map #1 ps' ~~ map (fn (x, _) => (x, NONE)) ps; |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
744 |
(* dummy syntax, since required by rename *) |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
745 |
val ps'' = map (fn ((p, _), (_, T)) => (p, T)) (ps ~~ ps'); |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
746 |
val [env] = unify_parms ctxt ps [map (apsnd SOME) ps'']; |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
747 |
(* propagate parameter types, to keep them consistent *) |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
748 |
val regs' = map (fn ((name, ps), ths) => |
18137 | 749 |
((name, map (Element.rename ren) ps), ths)) regs; |
17496 | 750 |
val new_regs = gen_rems (eq_fst (op =)) (regs', ids); |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
751 |
val new_ids = map fst new_regs; |
17485 | 752 |
val new_idTs = map (apsnd (map (fn p => (p, (the o AList.lookup (op =) ps) p)))) new_ids; |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
753 |
|
19780 | 754 |
val new_ths = new_regs |> map (#2 #> map |
755 |
(Element.instT_witness thy env #> |
|
756 |
Element.rename_witness ren #> |
|
757 |
Element.satisfy_witness ths)); |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
758 |
val new_ids' = map (fn (id, ths) => |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
759 |
(id, ([], Derived ths))) (new_ids ~~ new_ths); |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
760 |
in |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19423
diff
changeset
|
761 |
fold add_regs new_idTs (ths @ flat new_ths, ids @ new_ids') |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
762 |
end; |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
763 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
764 |
(* distribute top-level axioms over assumed ids *) |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
765 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
766 |
fun axiomify all_ps ((name, parms), (_, Assumed _)) axioms = |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
767 |
let |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
768 |
val {elems, ...} = the_locale thy name; |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19423
diff
changeset
|
769 |
val ts = maps |
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19423
diff
changeset
|
770 |
(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
|
771 |
| _ => []) |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19423
diff
changeset
|
772 |
elems; |
19018
88b4979193d8
added abbreviations: activated by init, no expressions yet;
wenzelm
parents:
18964
diff
changeset
|
773 |
val (axs1, axs2) = chop (length ts) axioms; |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
774 |
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
|
775 |
| axiomify all_ps (id, (_, Derived ths)) axioms = |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
776 |
((id, (all_ps, Derived ths)), axioms); |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
777 |
|
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
778 |
(* identifiers of an expression *) |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
779 |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
780 |
fun identify top (Locale name) = |
15596 | 781 |
(* 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
|
782 |
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
|
783 |
a list of axioms relating to the identifier, axs is empty unless |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
784 |
identify at top level (top = true); |
14215
ebf291f3b449
Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents:
13629
diff
changeset
|
785 |
parms is accumulated list of parameters *) |
12289 | 786 |
let |
20035
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
787 |
val {axiom, import, params, ...} = the_locale thy name; |
19278 | 788 |
val ps = map (#1 o #1) params; |
20035
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
789 |
val (ids', parms') = identify false import; |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
790 |
(* acyclic import dependencies *) |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
791 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
792 |
val ids'' = ids' @ [((name, ps), ([], Assumed []))]; |
19278 | 793 |
val (_, ids''') = add_regs (name, map #1 params) ([], ids''); |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
794 |
val (_, ids4) = chop (length ids' + 1) ids'''; |
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
795 |
val ids5 = ids' @ ids4 @ [((name, ps), ([], Assumed []))]; |
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
796 |
val ids_ax = if top then fst (fold_map (axiomify ps) ids5 axiom) else ids5; |
20035
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
797 |
in (ids_ax, merge_lists parms' ps) end |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
798 |
| identify top (Rename (e, xs)) = |
12273 | 799 |
let |
20035
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
800 |
val (ids', parms') = identify top e; |
12839 | 801 |
val ren = renaming xs parms' |
18678 | 802 |
handle ERROR msg => err_in_locale' ctxt msg ids'; |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
803 |
|
19061
ffbbac0261c9
removed distinct, renamed gen_distinct to distinct;
wenzelm
parents:
19025
diff
changeset
|
804 |
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
|
805 |
val parms'' = distinct (op =) (maps (#2 o #1) ids''); |
20035
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
806 |
in (ids'', parms'') end |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
807 |
| identify top (Merge es) = |
20035
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
808 |
fold (fn e => fn (ids, parms) => |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
809 |
let |
20035
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
810 |
val (ids', parms') = identify top e |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
811 |
in |
20035
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
812 |
(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
|
813 |
end) |
20035
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
814 |
es ([], []); |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
815 |
|
20035
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
816 |
fun inst_wit all_params (t, th) = let |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
817 |
val {hyps, prop, ...} = Thm.rep_thm th; |
16861 | 818 |
val ps = map (apsnd SOME) (fold Term.add_frees (prop :: hyps) []); |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
819 |
val [env] = unify_parms ctxt all_params [ps]; |
18137 | 820 |
val t' = Element.instT_term env t; |
821 |
val th' = Element.instT_thm thy env th; |
|
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
822 |
in (t', th') end; |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
823 |
|
20035
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
824 |
fun eval all_params tenv syn ((name, params), (locale_params, mode)) = |
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
825 |
let |
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
826 |
val {params = ps_mx, elems = elems_stamped, ...} = the_locale thy name; |
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
827 |
val elems = map fst elems_stamped; |
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
828 |
val ps = map fst ps_mx; |
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
829 |
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
|
830 |
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
|
831 |
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
|
832 |
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
|
833 |
val [env] = unify_parms ctxt all_params [map (apfst (Element.rename ren) o apsnd SOME) ps]; |
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
834 |
val elems' = elems |
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
835 |
|> map (Element.rename_ctxt ren) |
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
836 |
|> map (Element.map_ctxt {var = I, typ = I, term = I, fact = I, attrib = I, |
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
837 |
name = NameSpace.qualified (space_implode "_" params)}) |
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
838 |
|> map (Element.instT_ctxt thy env) |
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
839 |
in (((name, map (apsnd SOME) locale_params'), mode'), elems') end; |
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
840 |
|
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
841 |
(* parameters, their types and syntax *) |
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
842 |
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
|
843 |
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
|
844 |
(* compute identifiers and syntax, merge with previous ones *) |
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
845 |
val (ids, _) = identify true expr; |
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
846 |
val idents = gen_rems (eq_fst (op =)) (ids, prev_idents); |
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
847 |
val syntax = merge_syntax ctxt ids (syn, prev_syntax); |
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
848 |
(* type-instantiate elements *) |
80c79876d2de
Internal restructuring: identify no longer computes syntax.
ballarin
parents:
20032
diff
changeset
|
849 |
val final_elemss = map (eval all_params tenv syntax) idents; |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
850 |
in ((prev_idents @ idents, syntax), final_elemss) end; |
12046 | 851 |
|
12510 | 852 |
end; |
853 |
||
12070 | 854 |
|
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
855 |
(* activate elements *) |
12273 | 856 |
|
12510 | 857 |
local |
858 |
||
18671 | 859 |
fun axioms_export axs _ hyps = |
19780 | 860 |
Element.satisfy_thm axs |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
861 |
#> Drule.implies_intr_list (Library.drop (length axs, hyps)) |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
862 |
#> Seq.single; |
12263 | 863 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
864 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
865 |
(* 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
|
866 |
|
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
867 |
fun activate_elem _ _ ((ctxt, mode), Fixes fixes) = |
18671 | 868 |
((ctxt |> ProofContext.add_fixes_i fixes |> snd, mode), []) |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
869 |
| activate_elem _ _ ((ctxt, mode), Constrains _) = |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
870 |
((ctxt, mode), []) |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
871 |
| activate_elem ax_in_ctxt _ ((ctxt, Assumed axs), Assumes asms) = |
13399
c136276dc847
support locale ``views'' (for cumulative predicates);
wenzelm
parents:
13394
diff
changeset
|
872 |
let |
18728 | 873 |
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
|
874 |
val ts = maps (map #1 o #2) asms'; |
19018
88b4979193d8
added abbreviations: activated by init, no expressions yet;
wenzelm
parents:
18964
diff
changeset
|
875 |
val (ps, qs) = chop (length ts) axs; |
17856 | 876 |
val (_, ctxt') = |
18671 | 877 |
ctxt |> fold ProofContext.fix_frees ts |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
878 |
|> 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
|
879 |
in ((ctxt', Assumed qs), []) end |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
880 |
| activate_elem _ _ ((ctxt, Derived ths), Assumes asms) = |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
881 |
((ctxt, Derived ths), []) |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
882 |
| activate_elem _ _ ((ctxt, Assumed axs), Defines defs) = |
15596 | 883 |
let |
18728 | 884 |
val defs' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) defs; |
19732 | 885 |
val asms = defs' |> map (fn ((name, atts), (t, ps)) => |
886 |
let val ((c, _), t') = LocalDefs.cert_def ctxt t |
|
887 |
in (t', ((if name = "" then Thm.def_name c else name, atts), [(t', ps)])) end); |
|
17856 | 888 |
val (_, ctxt') = |
19732 | 889 |
ctxt |> fold (ProofContext.fix_frees o #1) asms |
890 |
|> 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
|
891 |
in ((ctxt', Assumed axs), []) end |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
892 |
| activate_elem _ _ ((ctxt, Derived ths), Defines defs) = |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
893 |
((ctxt, Derived ths), []) |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
894 |
| activate_elem _ is_ext ((ctxt, mode), Notes facts) = |
15596 | 895 |
let |
18728 | 896 |
val facts' = Attrib.map_facts (Attrib.attribute_i (ProofContext.theory_of ctxt)) facts; |
17856 | 897 |
val (res, ctxt') = ctxt |> ProofContext.note_thmss_i facts'; |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
898 |
in ((ctxt', mode), if is_ext then res else []) end; |
12502 | 899 |
|
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
900 |
fun activate_elems ax_in_ctxt (((name, ps), mode), elems) ctxt = |
17033 | 901 |
let |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
902 |
val thy = ProofContext.theory_of ctxt; |
17033 | 903 |
val ((ctxt', _), res) = |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
904 |
foldl_map (activate_elem ax_in_ctxt (name = "")) ((ProofContext.qualified_names ctxt, mode), elems) |
18678 | 905 |
handle ERROR msg => err_in_locale ctxt msg [(name, map fst ps)] |
15696 | 906 |
val ctxt'' = if name = "" then ctxt' |
907 |
else let |
|
908 |
val ps' = map (fn (n, SOME T) => Free (n, T)) ps; |
|
909 |
val ctxt'' = put_local_registration (name, ps') ("", []) ctxt' |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
910 |
in case mode of |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
911 |
Assumed axs => |
19780 | 912 |
fold (add_local_witness (name, ps') o |
913 |
Element.assume_witness thy o Element.witness_prop) axs ctxt'' |
|
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
914 |
| Derived ths => fold (add_local_witness (name, ps')) ths ctxt'' |
15696 | 915 |
end |
16144 | 916 |
in (ProofContext.restore_naming ctxt ctxt'', res) end; |
13399
c136276dc847
support locale ``views'' (for cumulative predicates);
wenzelm
parents:
13394
diff
changeset
|
917 |
|
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
918 |
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
|
919 |
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
|
920 |
let |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
921 |
val elems = map (prep_facts ctxt) raw_elems; |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19423
diff
changeset
|
922 |
val (ctxt', res) = apsnd flat |
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
923 |
(activate_elems ax_in_ctxt (((name, ps), mode), elems) ctxt); |
18137 | 924 |
val elems' = elems |> map (Element.map_ctxt |
925 |
{name = I, var = I, typ = I, term = I, fact = I, attrib = Args.closure}); |
|
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
926 |
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
|
927 |
|
12546 | 928 |
in |
929 |
||
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
930 |
(* CB: activate_facts prep_facts (ctxt, elemss), |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
931 |
where elemss is a list of pairs consisting of identifiers and |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
932 |
context elements, extends ctxt by the context elements yielding |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
933 |
ctxt' and returns (ctxt', (elemss', facts)). |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
934 |
Identifiers in the argument are of the form ((name, ps), axs) and |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
935 |
assumptions use the axioms in the identifiers to set up exporters |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
936 |
in ctxt'. elemss' does not contain identifiers and is obtained |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
937 |
from elemss and the intermediate context with prep_facts. |
15703 | 938 |
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
|
939 |
the internal/external markers from elemss. *) |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
940 |
|
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
941 |
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
|
942 |
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
|
943 |
in (ctxt', (elemss, flat factss)) end; |
15703 | 944 |
|
12510 | 945 |
end; |
946 |
||
12307 | 947 |
|
15696 | 948 |
|
18137 | 949 |
(** prepare locale elements **) |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
950 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
951 |
(* expressions *) |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
952 |
|
16458 | 953 |
fun intern_expr thy (Locale xname) = Locale (intern thy xname) |
954 |
| intern_expr thy (Merge exprs) = Merge (map (intern_expr thy) exprs) |
|
955 |
| 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
|
956 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
957 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
958 |
(* propositions and bindings *) |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
959 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
960 |
(* flatten (ctxt, prep_expr) ((ids, syn), expr) |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
961 |
normalises expr (which is either a locale |
14508
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
962 |
expression or a single context element) wrt. |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
963 |
to the list ids of already accumulated identifiers. |
19783 | 964 |
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
|
965 |
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
|
966 |
context elements generated from expr. |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
967 |
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
|
968 |
is an extension of syn. |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
969 |
For details, see flatten_expr. |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
970 |
|
15596 | 971 |
Additionally, for a locale expression, the elems are grouped into a single |
972 |
Int; individual context elements are marked Ext. In this case, the |
|
973 |
identifier-like information of the element is as follows: |
|
974 |
- for Fixes: (("", ps), []) where the ps have type info NONE |
|
975 |
- for other elements: (("", []), []). |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
976 |
The implementation of activate_facts relies on identifier names being |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
977 |
empty strings for external elements. |
15596 | 978 |
*) |
14508
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
979 |
|
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
980 |
fun flatten (ctxt, _) ((ids, syn), Elem (Fixes fixes)) = let |
18137 | 981 |
val ids' = ids @ [(("", map #1 fixes), ([], Assumed []))] |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
982 |
in |
18137 | 983 |
((ids', |
984 |
merge_syntax ctxt ids' |
|
985 |
(syn, Symtab.make (map (fn fx => (#1 fx, #3 fx)) fixes)) |
|
986 |
handle Symtab.DUPS xs => err_in_locale ctxt |
|
19931
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
ballarin
parents:
19914
diff
changeset
|
987 |
("Conflicting syntax for parameters: " ^ commas_quote xs) |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
988 |
(map #1 ids')), |
18137 | 989 |
[((("", map (rpair NONE o #1) fixes), Assumed []), Ext (Fixes fixes))]) |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
990 |
end |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
991 |
| flatten _ ((ids, syn), Elem elem) = |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
992 |
((ids @ [(("", []), ([], Assumed []))], syn), [((("", []), Assumed []), Ext elem)]) |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
993 |
| flatten (ctxt, prep_expr) ((ids, syn), Expr expr) = |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
994 |
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
|
995 |
|
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
996 |
local |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
997 |
|
12839 | 998 |
local |
999 |
||
12727 | 1000 |
fun declare_int_elem (ctxt, Fixes fixes) = |
18671 | 1001 |
(ctxt |> ProofContext.add_fixes_i (map (fn (x, T, mx) => |
1002 |
(x, Option.map (Term.map_type_tfree (TypeInfer.param 0)) T, mx)) fixes) |> snd, []) |
|
12727 | 1003 |
| declare_int_elem (ctxt, _) = (ctxt, []); |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1004 |
|
18671 | 1005 |
fun declare_ext_elem prep_vars (ctxt, Fixes fixes) = |
1006 |
let val (vars, _) = prep_vars fixes ctxt |
|
1007 |
in (ctxt |> ProofContext.add_fixes_i vars |> snd, []) end |
|
1008 |
| declare_ext_elem prep_vars (ctxt, Constrains csts) = |
|
1009 |
let val (_, ctxt') = prep_vars (map (fn (x, T) => (x, SOME T, NoSyn)) csts) ctxt |
|
1010 |
in (ctxt', []) end |
|
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1011 |
| declare_ext_elem _ (ctxt, Assumes asms) = (ctxt, map #2 asms) |
19585 | 1012 |
| declare_ext_elem _ (ctxt, Defines defs) = (ctxt, map (fn (_, (t, ps)) => [(t, ps)]) defs) |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1013 |
| declare_ext_elem _ (ctxt, Notes facts) = (ctxt, []); |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1014 |
|
18671 | 1015 |
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
|
1016 |
let val (ctxt', propps) = |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1017 |
(case elems of |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1018 |
Int es => foldl_map declare_int_elem (ctxt, es) |
18671 | 1019 |
| Ext e => foldl_map (declare_ext_elem prep_vars) (ctxt, [e])) |
18678 | 1020 |
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
|
1021 |
in (ctxt', propps) end |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1022 |
| declare_elems _ (ctxt, ((_, Derived _), elems)) = (ctxt, []); |
12727 | 1023 |
|
12839 | 1024 |
in |
1025 |
||
18671 | 1026 |
fun declare_elemss prep_vars fixed_params raw_elemss ctxt = |
12727 | 1027 |
let |
14215
ebf291f3b449
Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents:
13629
diff
changeset
|
1028 |
(* 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
|
1029 |
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
|
1030 |
distinct from types of parameters in target (fixed_params). *) |
ebf291f3b449
Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents:
13629
diff
changeset
|
1031 |
val ctxt_with_fixed = |
19900
21a99d88d925
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19873
diff
changeset
|
1032 |
fold Variable.declare_term (map Free fixed_params) ctxt; |
12727 | 1033 |
val int_elemss = |
1034 |
raw_elemss |
|
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19423
diff
changeset
|
1035 |
|> 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
|
1036 |
|> unify_elemss ctxt_with_fixed fixed_params; |
12727 | 1037 |
val (_, raw_elemss') = |
1038 |
foldl_map (fn ((_, es) :: elemss, (id, Int _)) => (elemss, (id, Int es)) | x => x) |
|
1039 |
(int_elemss, raw_elemss); |
|
18671 | 1040 |
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
|
1041 |
|
12839 | 1042 |
end; |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1043 |
|
12839 | 1044 |
local |
1045 |
||
1046 |
val norm_term = Envir.beta_norm oo Term.subst_atomic; |
|
1047 |
||
16458 | 1048 |
fun abstract_thm thy eq = |
1049 |
Thm.assume (Thm.cterm_of thy eq) |> Drule.gen_all |> Drule.abs_def; |
|
12502 | 1050 |
|
18190 | 1051 |
fun bind_def ctxt (name, ps) eq (xs, env, ths) = |
12839 | 1052 |
let |
18831 | 1053 |
val ((y, T), b) = LocalDefs.abs_def eq; |
13308 | 1054 |
val b' = norm_term env b; |
16458 | 1055 |
val th = abstract_thm (ProofContext.theory_of ctxt) eq; |
13308 | 1056 |
fun err msg = err_in_locale ctxt (msg ^ ": " ^ quote y) [(name, map fst ps)]; |
12839 | 1057 |
in |
13308 | 1058 |
conditional (exists (equal y o #1) xs) (fn () => |
1059 |
err "Attempt to define previously specified variable"); |
|
1060 |
conditional (exists (fn (Free (y', _), _) => y = y' | _ => false) env) (fn () => |
|
1061 |
err "Attempt to redefine variable"); |
|
16861 | 1062 |
(Term.add_frees b' xs, (Free (y, T), b') :: env, th :: ths) |
12839 | 1063 |
end; |
12575 | 1064 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1065 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1066 |
(* CB: for finish_elems (Int and Ext), |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1067 |
extracts specification, only of assumed elements *) |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1068 |
|
18190 | 1069 |
fun eval_text _ _ _ (Fixes _) text = text |
1070 |
| eval_text _ _ _ (Constrains _) text = text |
|
1071 |
| eval_text _ (_, Assumed _) is_ext (Assumes asms) |
|
1072 |
(((exts, exts'), (ints, ints')), (xs, env, defs)) = |
|
13394 | 1073 |
let |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19423
diff
changeset
|
1074 |
val ts = maps (map #1 o #2) asms; |
13394 | 1075 |
val ts' = map (norm_term env) ts; |
1076 |
val spec' = |
|
1077 |
if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints')) |
|
1078 |
else ((exts, exts'), (ints @ ts, ints' @ ts')); |
|
16861 | 1079 |
in (spec', (fold Term.add_frees ts' xs, env, defs)) end |
18190 | 1080 |
| eval_text _ (_, Derived _) _ (Assumes _) text = text |
1081 |
| eval_text ctxt (id, Assumed _) _ (Defines defs) (spec, binds) = |
|
1082 |
(spec, fold (bind_def ctxt id o #1 o #2) defs binds) |
|
1083 |
| eval_text _ (_, Derived _) _ (Defines _) text = text |
|
1084 |
| eval_text _ _ _ (Notes _) text = text; |
|
13308 | 1085 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1086 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1087 |
(* for finish_elems (Int), |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1088 |
remove redundant elements of derived identifiers, |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1089 |
turn assumptions and definitions into facts, |
19780 | 1090 |
adjust hypotheses of facts using witnesses *) |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1091 |
|
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
1092 |
fun finish_derived _ _ (Assumed _) (Fixes fixes) = SOME (Fixes fixes) |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
1093 |
| finish_derived _ _ (Assumed _) (Constrains csts) = SOME (Constrains csts) |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
1094 |
| finish_derived _ _ (Assumed _) (Assumes asms) = SOME (Assumes asms) |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
1095 |
| finish_derived _ _ (Assumed _) (Defines defs) = SOME (Defines defs) |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
1096 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1097 |
| finish_derived _ _ (Derived _) (Fixes _) = NONE |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1098 |
| finish_derived _ _ (Derived _) (Constrains _) = NONE |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1099 |
| finish_derived sign wits (Derived _) (Assumes asms) = asms |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
1100 |
|> map (apsnd (map (fn (a, _) => ([Thm.assume (cterm_of sign a)], [])))) |
19780 | 1101 |
|> Notes |> Element.map_ctxt_values I I (Element.satisfy_thm wits) |> SOME |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1102 |
| finish_derived sign wits (Derived _) (Defines defs) = defs |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
1103 |
|> map (apsnd (fn (d, _) => [([Thm.assume (cterm_of sign d)], [])])) |
19780 | 1104 |
|> Notes |> Element.map_ctxt_values I I (Element.satisfy_thm wits) |> SOME |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1105 |
|
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
1106 |
| finish_derived _ wits _ (Notes facts) = (Notes facts) |
19780 | 1107 |
|> Element.map_ctxt_values I I (Element.satisfy_thm wits) |> SOME; |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1108 |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1109 |
(* CB: for finish_elems (Ext) *) |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1110 |
|
13308 | 1111 |
fun closeup _ false elem = elem |
1112 |
| closeup ctxt true elem = |
|
12839 | 1113 |
let |
13308 | 1114 |
fun close_frees t = |
19900
21a99d88d925
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19873
diff
changeset
|
1115 |
let val frees = rev (filter_out (Variable.is_fixed ctxt o #1) (Term.add_frees t [])) |
13308 | 1116 |
in Term.list_all_free (frees, t) end; |
1117 |
||
1118 |
fun no_binds [] = [] |
|
18678 | 1119 |
| no_binds _ = error "Illegal term bindings in locale element"; |
13308 | 1120 |
in |
1121 |
(case elem of |
|
1122 |
Assumes asms => Assumes (asms |> map (fn (a, propps) => |
|
19585 | 1123 |
(a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps))) |
13308 | 1124 |
| Defines defs => Defines (defs |> map (fn (a, (t, ps)) => |
18831 | 1125 |
(a, (close_frees (#2 (LocalDefs.cert_def ctxt t)), no_binds ps)))) |
13308 | 1126 |
| e => e) |
1127 |
end; |
|
12839 | 1128 |
|
12502 | 1129 |
|
12839 | 1130 |
fun finish_ext_elem parms _ (Fixes fixes, _) = Fixes (map (fn (x, _, mx) => |
17271 | 1131 |
(x, AList.lookup (op =) parms x, mx)) fixes) |
18899
a8e913c93578
theorem(_in_locale): Element.statement, Obtain.statement;
wenzelm
parents:
18890
diff
changeset
|
1132 |
| finish_ext_elem parms _ (Constrains _, _) = Constrains [] |
12839 | 1133 |
| finish_ext_elem _ close (Assumes asms, propp) = |
1134 |
close (Assumes (map #1 asms ~~ propp)) |
|
1135 |
| finish_ext_elem _ close (Defines defs, propp) = |
|
19585 | 1136 |
close (Defines (map #1 defs ~~ map (fn [(t, ps)] => (t, ps)) propp)) |
12839 | 1137 |
| finish_ext_elem _ _ (Notes facts, _) = Notes facts; |
1138 |
||
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1139 |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1140 |
(* CB: finish_parms introduces type info from parms to identifiers *) |
15531 | 1141 |
(* CB: only needed for types that have been NONE so far??? |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1142 |
If so, which are these??? *) |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1143 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1144 |
fun finish_parms parms (((name, ps), mode), elems) = |
19932 | 1145 |
(((name, map (fn (x, _) => (x, AList.lookup (op = : string * string -> bool) parms x)) ps), mode), elems); |
12839 | 1146 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1147 |
fun finish_elems ctxt parms _ ((text, wits), ((id, Int e), _)) = |
12839 | 1148 |
let |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1149 |
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
|
1150 |
val wits' = case mode of Assumed _ => wits | Derived ths => wits @ ths; |
18190 | 1151 |
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
|
1152 |
val es' = map_filter |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1153 |
(finish_derived (ProofContext.theory_of ctxt) wits' mode) es; |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1154 |
in ((text', wits'), (id', map Int es')) end |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1155 |
| finish_elems ctxt parms do_close ((text, wits), ((id, Ext e), [propp])) = |
13308 | 1156 |
let |
1157 |
val e' = finish_ext_elem parms (closeup ctxt do_close) (e, propp); |
|
18190 | 1158 |
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
|
1159 |
in ((text', wits), (id, [Ext e'])) end |
12839 | 1160 |
|
1161 |
in |
|
12510 | 1162 |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1163 |
(* CB: only called by prep_elemss *) |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1164 |
|
13375 | 1165 |
fun finish_elemss ctxt parms do_close = |
1166 |
foldl_map (apsnd (finish_parms parms) o finish_elems ctxt parms do_close); |
|
12839 | 1167 |
|
1168 |
end; |
|
1169 |
||
16736
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
1170 |
|
19942
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1171 |
(* Remove duplicate Defines elements: temporary workaround to fix Afp/Category. *) |
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1172 |
|
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1173 |
fun defs_ord (defs1, defs2) = |
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1174 |
list_ord (fn ((_, (d1, _)), (_, (d2, _))) => |
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1175 |
Term.fast_term_ord (d1, d2)) (defs1, defs2); |
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1176 |
structure Defstab = |
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1177 |
TableFun(type key = ((string * Attrib.src list) * (term * term list)) list |
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1178 |
val ord = defs_ord); |
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1179 |
|
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1180 |
fun rem_dup_defs es ds = |
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1181 |
fold_map (fn e as (Defines defs) => (fn ds => |
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1182 |
if Defstab.defined ds defs |
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1183 |
then (Defines [], ds) |
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1184 |
else (e, Defstab.update (defs, ()) ds)) |
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1185 |
| e => (fn ds => (e, ds))) es ds; |
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1186 |
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
|
1187 |
| rem_dup_elemss (Ext e) ds = (Ext e, ds); |
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1188 |
fun rem_dup_defines raw_elemss = |
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1189 |
fold_map (fn (id as (_, (Assumed _)), es) => (fn ds => |
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1190 |
apfst (pair id) (rem_dup_elemss es ds)) |
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1191 |
| (id as (_, (Derived _)), es) => (fn ds => |
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1192 |
((id, es), ds))) raw_elemss Defstab.empty |> #1; |
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1193 |
|
16736
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
1194 |
(* CB: type inference and consistency checks for locales. |
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
1195 |
|
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
1196 |
Works by building a context (through declare_elemss), extracting the |
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
1197 |
required information and adjusting the context elements (finish_elemss). |
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
1198 |
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
|
1199 |
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
|
1200 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1201 |
Only elements of assumed identifiers are considered. |
16736
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
1202 |
*) |
15127 | 1203 |
|
18671 | 1204 |
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
|
1205 |
let |
15127 | 1206 |
(* CB: contexts computed in the course of this function are discarded. |
1207 |
They are used for type inference and consistency checks only. *) |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1208 |
(* CB: fixed_params are the parameters (with types) of the target locale, |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1209 |
empty list if there is no target. *) |
14508
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1210 |
(* 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
|
1211 |
context elements, the latter marked as internal or external. *) |
19942
dc92e3aebc71
Improved handling of defines imported in duplicate.
ballarin
parents:
19932
diff
changeset
|
1212 |
val raw_elemss = rem_dup_defines raw_elemss; |
18671 | 1213 |
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
|
1214 |
(* 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
|
1215 |
the fixes elements in raw_elemss, |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1216 |
raw_proppss contains assumptions and definitions from the |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1217 |
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
|
1218 |
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
|
1219 |
let |
e57731ba01dd
discontinued unflat in favour of burrow and burrow_split
haftmann
parents:
18421
diff
changeset
|
1220 |
(* 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
|
1221 |
(* CB: process patterns (conclusion and external elements only) *) |
e57731ba01dd
discontinued unflat in favour of burrow and burrow_split
haftmann
parents:
18421
diff
changeset
|
1222 |
val (ctxt, all_propp) = |
19900
21a99d88d925
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19873
diff
changeset
|
1223 |
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
|
1224 |
(* 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
|
1225 |
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
|
1226 |
(* 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
|
1227 |
val all_propp' = map2 (curry (op ~~)) |
e57731ba01dd
discontinued unflat in favour of burrow and burrow_split
haftmann
parents:
18421
diff
changeset
|
1228 |
(#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
|
1229 |
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
|
1230 |
in (propp, (ctxt, concl)) end |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1231 |
|
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1232 |
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
|
1233 |
(fold_burrow o fold_burrow) prep_prop raw_proppss (raw_ctxt, raw_concl); |
12502 | 1234 |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1235 |
(* CB: obtain all parameters from identifier part of raw_elemss *) |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1236 |
val xs = map #1 (params_of' raw_elemss); |
12727 | 1237 |
val typing = unify_frozen ctxt 0 |
19900
21a99d88d925
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19873
diff
changeset
|
1238 |
(map (Variable.default_type raw_ctxt) xs) |
21a99d88d925
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19873
diff
changeset
|
1239 |
(map (Variable.default_type ctxt) xs); |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1240 |
val parms = param_types (xs ~~ typing); |
14508
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1241 |
(* CB: parms are the parameters from raw_elemss, with correct typing. *) |
12273 | 1242 |
|
14508
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1243 |
(* CB: extract information from assumes and defines elements |
16169
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16144
diff
changeset
|
1244 |
(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
|
1245 |
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
|
1246 |
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
|
1247 |
((((([], []), ([], [])), ([], [], [])), []), raw_elemss ~~ proppss); |
14508
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1248 |
(* CB: text has the following structure: |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1249 |
(((exts, exts'), (ints, ints')), (xs, env, defs)) |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
|