author | wenzelm |
Wed, 25 Jan 2006 00:21:40 +0100 | |
changeset 18782 | e4d9d718b8bb |
parent 18768 | 6e97b57cdcba |
child 18795 | 303793f49b0f |
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 |
|
14215
ebf291f3b449
Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents:
13629
diff
changeset
|
8 |
Draws some 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 |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
10 |
meta-logic. Furthermore, we provide structured import of contexts |
14215
ebf291f3b449
Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents:
13629
diff
changeset
|
11 |
(with merge and rename operations), 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. |
11896 | 23 |
*) |
24 |
||
16169
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16144
diff
changeset
|
25 |
(* TODO: |
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16144
diff
changeset
|
26 |
- beta-eta normalisation of interpretation parameters |
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16144
diff
changeset
|
27 |
- dangling type frees in locales |
16620
2a7f46324218
Proper treatment of beta-redexes in witness theorems.
ballarin
parents:
16458
diff
changeset
|
28 |
- test subsumption of interpretations when merging theories |
17138 | 29 |
- var vs. fixes in locale to be interpreted (interpretation in locale) |
30 |
(implicit locale expressions generated by multiple registrations) |
|
16169
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16144
diff
changeset
|
31 |
*) |
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16144
diff
changeset
|
32 |
|
11896 | 33 |
signature LOCALE = |
34 |
sig |
|
12273 | 35 |
datatype expr = |
36 |
Locale of string | |
|
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
37 |
Rename of expr * (string * mixfix option) option list | |
12273 | 38 |
Merge of expr list |
39 |
val empty: expr |
|
18137 | 40 |
datatype 'a element = Elem of 'a | Expr of expr |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
41 |
|
16458 | 42 |
val intern: theory -> xstring -> string |
43 |
val extern: theory -> string -> xstring |
|
18617 | 44 |
val init: string -> theory -> Proof.context |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
45 |
|
18617 | 46 |
(* Processing of locale statements *) (* FIXME export more abstract version *) |
18137 | 47 |
val read_context_statement: xstring option -> Element.context element list -> |
18617 | 48 |
(string * (string list * string list)) list list -> Proof.context -> |
49 |
string option * (cterm list * cterm list) * Proof.context * Proof.context * |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
50 |
(term * (term list * term list)) list list |
18137 | 51 |
val cert_context_statement: string option -> Element.context_i element list -> |
18617 | 52 |
(term * (term list * term list)) list list -> Proof.context -> |
53 |
string option * (cterm list * cterm list) * Proof.context * Proof.context * |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
54 |
(term * (term list * term list)) list list |
15596 | 55 |
|
56 |
(* Diagnostic functions *) |
|
12758 | 57 |
val print_locales: theory -> unit |
18137 | 58 |
val print_locale: theory -> bool -> expr -> Element.context list -> unit |
17138 | 59 |
val print_global_registrations: bool -> string -> theory -> unit |
18617 | 60 |
val print_local_registrations': bool -> string -> Proof.context -> unit |
61 |
val print_local_registrations: bool -> string -> Proof.context -> unit |
|
18137 | 62 |
|
18671 | 63 |
(* FIXME avoid duplicates *) |
18137 | 64 |
val add_locale_context: bool -> bstring -> expr -> Element.context list -> theory |
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
65 |
-> ((Element.context_i list list * Element.context_i list list) |
18617 | 66 |
* Proof.context) * theory |
18137 | 67 |
val add_locale_context_i: bool -> bstring -> expr -> Element.context_i list -> theory |
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
68 |
-> ((Element.context_i list list * Element.context_i list list) |
18617 | 69 |
* Proof.context) * theory |
18137 | 70 |
val add_locale: bool -> bstring -> expr -> Element.context list -> theory -> theory |
71 |
val add_locale_i: bool -> bstring -> expr -> Element.context_i list -> theory -> theory |
|
15596 | 72 |
|
15696 | 73 |
(* Storing results *) |
15703 | 74 |
val smart_note_thmss: string -> string option -> |
18728 | 75 |
((bstring * attribute list) * (thm list * attribute list) list) list -> |
18421
464c93701351
re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents:
18377
diff
changeset
|
76 |
theory -> (bstring * thm list) list * theory |
15703 | 77 |
val note_thmss: string -> xstring -> |
78 |
((bstring * Attrib.src list) * (thmref * Attrib.src list) list) list -> |
|
18617 | 79 |
theory -> (bstring * thm list) list * (theory * Proof.context) |
15703 | 80 |
val note_thmss_i: string -> string -> |
81 |
((bstring * Attrib.src list) * (thm list * Attrib.src list) list) list -> |
|
18617 | 82 |
theory -> (bstring * thm list) list * (theory * Proof.context) |
18137 | 83 |
|
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
84 |
val theorem: string -> Method.text option -> (thm list list -> theory -> theory) -> |
18137 | 85 |
string * Attrib.src list -> Element.context element list -> |
17355 | 86 |
((string * Attrib.src list) * (string * (string list * string list)) list) list -> |
87 |
theory -> Proof.state |
|
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
88 |
val theorem_i: string -> Method.text option -> (thm list list -> theory -> theory) -> |
18728 | 89 |
string * attribute list -> Element.context_i element list -> |
90 |
((string * attribute list) * (term * (term list * term list)) list) list -> |
|
17355 | 91 |
theory -> Proof.state |
17856 | 92 |
val theorem_in_locale: string -> Method.text option -> |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
93 |
(thm list list -> thm list list -> theory -> theory) -> |
18137 | 94 |
xstring -> string * Attrib.src list -> Element.context element list -> |
17355 | 95 |
((string * Attrib.src list) * (string * (string list * string list)) list) list -> |
96 |
theory -> Proof.state |
|
17856 | 97 |
val theorem_in_locale_i: string -> Method.text option -> |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
98 |
(thm list list -> thm list list -> theory -> theory) -> |
18137 | 99 |
string -> string * Attrib.src list -> Element.context_i element list -> |
17355 | 100 |
((string * Attrib.src list) * (term * (term list * term list)) list) list -> |
101 |
theory -> Proof.state |
|
102 |
val smart_theorem: string -> xstring option -> |
|
18137 | 103 |
string * Attrib.src list -> Element.context element list -> |
17355 | 104 |
((string * Attrib.src list) * (string * (string list * string list)) list) list -> |
105 |
theory -> Proof.state |
|
106 |
val interpretation: string * Attrib.src list -> expr -> string option list -> |
|
107 |
theory -> Proof.state |
|
108 |
val interpretation_in_locale: xstring * expr -> theory -> Proof.state |
|
109 |
val interpret: string * Attrib.src list -> expr -> string option list -> |
|
110 |
bool -> Proof.state -> Proof.state |
|
11896 | 111 |
end; |
12839 | 112 |
|
12289 | 113 |
structure Locale: LOCALE = |
11896 | 114 |
struct |
115 |
||
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
116 |
|
12273 | 117 |
(** locale elements and expressions **) |
11896 | 118 |
|
18137 | 119 |
datatype ctxt = datatype Element.ctxt; |
17355 | 120 |
|
12273 | 121 |
datatype expr = |
122 |
Locale of string | |
|
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
123 |
Rename of expr * (string * mixfix option) option list | |
12273 | 124 |
Merge of expr list; |
11896 | 125 |
|
12273 | 126 |
val empty = Merge []; |
127 |
||
18137 | 128 |
datatype 'a element = |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
129 |
Elem of 'a | Expr of expr; |
12273 | 130 |
|
18137 | 131 |
fun map_elem f (Elem e) = Elem (f e) |
132 |
| map_elem _ (Expr e) = Expr e; |
|
133 |
||
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
134 |
type witness = term * thm; (*hypothesis as fact*) |
12070 | 135 |
type locale = |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
136 |
{predicate: cterm list * witness list, |
16736
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
137 |
(* CB: For locales with "(open)" this entry is ([], []). |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
138 |
For new-style locales, which declare predicates, if the locale declares |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
139 |
no predicates, this is also ([], []). |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
140 |
If the locale declares predicates, the record field is |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
141 |
([statement], axioms), where statement is the locale predicate applied |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
142 |
to the (assumed) locale parameters. Axioms contains the projections |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
143 |
from the locale predicate to the normalised assumptions of the locale |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
144 |
(cf. [1], normalisation of locale expressions.) |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
145 |
*) |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
146 |
import: expr, (*dynamic import*) |
18137 | 147 |
elems: (Element.context_i * stamp) list, (*static content*) |
18671 | 148 |
params: ((string * typ) * mixfix) list * string list, |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
149 |
(*all/local params*) |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
150 |
regs: ((string * string list) * (term * thm) list) list} |
16736
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
151 |
(* Registrations: indentifiers and witness theorems of locales interpreted |
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
152 |
in the locale. |
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
153 |
*) |
12063 | 154 |
|
11896 | 155 |
|
15703 | 156 |
(* CB: an internal (Int) locale element was either imported or included, |
157 |
an external (Ext) element appears directly in the locale text. *) |
|
158 |
||
159 |
datatype ('a, 'b) int_ext = Int of 'a | Ext of 'b; |
|
160 |
||
161 |
||
162 |
||
16736
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
163 |
(** management of registrations in theories and proof contexts **) |
11896 | 164 |
|
15837 | 165 |
structure Registrations : |
166 |
sig |
|
167 |
type T |
|
168 |
val empty: T |
|
169 |
val join: T * T -> T |
|
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
170 |
val dest: T -> (term list * ((string * Attrib.src list) * witness list)) list |
16458 | 171 |
val lookup: theory -> T * term list -> |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
172 |
((string * Attrib.src list) * witness list) option |
16458 | 173 |
val insert: theory -> term list * (string * Attrib.src list) -> T -> |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
174 |
T * (term list * ((string * Attrib.src list) * witness list)) list |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
175 |
val add_witness: term list -> witness -> T -> T |
15837 | 176 |
end = |
177 |
struct |
|
178 |
(* a registration consists of theorems instantiating locale assumptions |
|
179 |
and prefix and attributes, indexed by parameter instantiation *) |
|
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
180 |
type T = ((string * Attrib.src list) * witness list) Termtab.table; |
15837 | 181 |
|
182 |
val empty = Termtab.empty; |
|
183 |
||
184 |
(* term list represented as single term, for simultaneous matching *) |
|
185 |
fun termify ts = |
|
18343 | 186 |
Term.list_comb (Const ("", map fastype_of ts ---> propT), ts); |
15837 | 187 |
fun untermify t = |
188 |
let fun ut (Const _) ts = ts |
|
189 |
| ut (s $ t) ts = ut s (t::ts) |
|
190 |
in ut t [] end; |
|
191 |
||
16620
2a7f46324218
Proper treatment of beta-redexes in witness theorems.
ballarin
parents:
16458
diff
changeset
|
192 |
(* joining of registrations: prefix and attributes of left theory, |
15837 | 193 |
thms are equal, no attempt to subsumption testing *) |
16458 | 194 |
fun join (r1, r2) = Termtab.join (fn _ => fn (reg, _) => SOME reg) (r1, r2); |
15837 | 195 |
|
196 |
fun dest regs = map (apfst untermify) (Termtab.dest regs); |
|
197 |
||
198 |
(* registrations that subsume t *) |
|
17203 | 199 |
fun subsumers thy t regs = |
200 |
List.filter (fn (t', _) => Pattern.matches thy (t', t)) (Termtab.dest regs); |
|
15837 | 201 |
|
202 |
(* look up registration, pick one that subsumes the query *) |
|
203 |
fun lookup sign (regs, ts) = |
|
204 |
let |
|
205 |
val t = termify ts; |
|
17203 | 206 |
val subs = subsumers sign t regs; |
15837 | 207 |
in (case subs of |
208 |
[] => NONE |
|
209 |
| ((t', (attn, thms)) :: _) => let |
|
18190 | 210 |
val (tinst, inst) = Pattern.match sign (t', t) (Vartab.empty, Vartab.empty); |
15837 | 211 |
(* thms contain Frees, not Vars *) |
212 |
val tinst' = tinst |> Vartab.dest |
|
213 |
|> map (fn ((x, 0), (_, T)) => (x, Type.unvarifyT T)) |
|
214 |
|> Symtab.make; |
|
215 |
val inst' = inst |> Vartab.dest |
|
216 |
|> map (fn ((x, 0), (_, t)) => (x, Logic.unvarify t)) |
|
217 |
|> Symtab.make; |
|
218 |
in |
|
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
219 |
SOME (attn, map (fn (t, th) => |
18137 | 220 |
(Element.inst_term (tinst', inst') t, |
221 |
Element.inst_thm sign (tinst', inst') th)) thms) |
|
15837 | 222 |
end) |
223 |
end; |
|
224 |
||
225 |
(* add registration if not subsumed by ones already present, |
|
226 |
additionally returns registrations that are strictly subsumed *) |
|
227 |
fun insert sign (ts, attn) regs = |
|
228 |
let |
|
229 |
val t = termify ts; |
|
17203 | 230 |
val subs = subsumers sign t regs ; |
15837 | 231 |
in (case subs of |
232 |
[] => let |
|
233 |
val sups = |
|
17203 | 234 |
List.filter (fn (t', _) => Pattern.matches sign (t, t')) (Termtab.dest regs); |
15837 | 235 |
val sups' = map (apfst untermify) sups |
17412 | 236 |
in (Termtab.update (t, (attn, [])) regs, sups') end |
15837 | 237 |
| _ => (regs, [])) |
238 |
end; |
|
239 |
||
240 |
(* add witness theorem to registration, |
|
16169
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16144
diff
changeset
|
241 |
only if instantiation is exact, otherwise exception Option raised *) |
15837 | 242 |
fun add_witness ts thm regs = |
243 |
let |
|
244 |
val t = termify ts; |
|
18343 | 245 |
val (x, thms) = the (Termtab.lookup regs t); |
15837 | 246 |
in |
17412 | 247 |
Termtab.update (t, (x, thm::thms)) regs |
15837 | 248 |
end; |
249 |
end; |
|
250 |
||
16736
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
251 |
|
15837 | 252 |
(** theory data **) |
15596 | 253 |
|
16458 | 254 |
structure GlobalLocalesData = TheoryDataFun |
255 |
(struct |
|
12014 | 256 |
val name = "Isar/locales"; |
15837 | 257 |
type T = NameSpace.T * locale Symtab.table * Registrations.T Symtab.table; |
15596 | 258 |
(* 1st entry: locale namespace, |
259 |
2nd entry: locales of the theory, |
|
15837 | 260 |
3rd entry: registrations, indexed by locale name *) |
11896 | 261 |
|
15596 | 262 |
val empty = (NameSpace.empty, Symtab.empty, Symtab.empty); |
12063 | 263 |
val copy = I; |
16458 | 264 |
val extend = I; |
12289 | 265 |
|
16736
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
266 |
fun join_locs _ ({predicate, import, elems, params, regs}: locale, |
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
267 |
{elems = elems', regs = regs', ...}: locale) = |
15596 | 268 |
SOME {predicate = predicate, import = import, |
17496 | 269 |
elems = gen_merge_lists (eq_snd (op =)) elems elems', |
16736
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
270 |
params = params, |
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
271 |
regs = merge_alists regs regs'}; |
16458 | 272 |
fun merge _ ((space1, locs1, regs1), (space2, locs2, regs2)) = |
15596 | 273 |
(NameSpace.merge (space1, space2), Symtab.join join_locs (locs1, locs2), |
16458 | 274 |
Symtab.join (K (SOME o Registrations.join)) (regs1, regs2)); |
12289 | 275 |
|
15596 | 276 |
fun print _ (space, locs, _) = |
16346 | 277 |
Pretty.strs ("locales:" :: map #1 (NameSpace.extern_table (space, locs))) |
12014 | 278 |
|> Pretty.writeln; |
16458 | 279 |
end); |
11896 | 280 |
|
18708 | 281 |
val _ = Context.add_setup GlobalLocalesData.init; |
15801 | 282 |
|
283 |
||
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
284 |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
285 |
(** context data **) |
11896 | 286 |
|
16458 | 287 |
structure LocalLocalesData = ProofDataFun |
288 |
(struct |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
289 |
val name = "Isar/locales"; |
15837 | 290 |
type T = Registrations.T Symtab.table; |
291 |
(* registrations, indexed by locale name *) |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
292 |
fun init _ = Symtab.empty; |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
293 |
fun print _ _ = (); |
16458 | 294 |
end); |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
295 |
|
18708 | 296 |
val _ = Context.add_setup LocalLocalesData.init; |
12289 | 297 |
|
12277 | 298 |
|
299 |
(* access locales *) |
|
300 |
||
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
301 |
val print_locales = GlobalLocalesData.print; |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
302 |
|
16458 | 303 |
val intern = NameSpace.intern o #1 o GlobalLocalesData.get; |
304 |
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
|
305 |
|
16144 | 306 |
fun declare_locale name thy = |
307 |
thy |> GlobalLocalesData.map (fn (space, locs, regs) => |
|
16458 | 308 |
(Sign.declare_name thy name space, locs, regs)); |
11896 | 309 |
|
15596 | 310 |
fun put_locale name loc = |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
311 |
GlobalLocalesData.map (fn (space, locs, regs) => |
17412 | 312 |
(space, Symtab.update (name, loc) locs, regs)); |
313 |
||
314 |
fun get_locale thy name = Symtab.lookup (#2 (GlobalLocalesData.get thy)) name; |
|
11896 | 315 |
|
12014 | 316 |
fun the_locale thy name = |
317 |
(case get_locale thy name of |
|
15531 | 318 |
SOME loc => loc |
319 |
| NONE => error ("Unknown locale " ^ quote name)); |
|
11896 | 320 |
|
12046 | 321 |
|
15596 | 322 |
(* access registrations *) |
323 |
||
15696 | 324 |
(* Ids of global registrations are varified, |
325 |
Ids of local registrations are not. |
|
326 |
Thms of registrations are never varified. *) |
|
327 |
||
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
328 |
(* retrieve registration from theory or context *) |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
329 |
|
15696 | 330 |
fun gen_get_registrations get thy_ctxt name = |
17412 | 331 |
case Symtab.lookup (get thy_ctxt) name of |
15696 | 332 |
NONE => [] |
15837 | 333 |
| SOME reg => Registrations.dest reg; |
15696 | 334 |
|
335 |
val get_global_registrations = |
|
336 |
gen_get_registrations (#3 o GlobalLocalesData.get); |
|
337 |
val get_local_registrations = |
|
338 |
gen_get_registrations LocalLocalesData.get; |
|
339 |
||
16458 | 340 |
fun gen_get_registration get thy_of thy_ctxt (name, ps) = |
17412 | 341 |
case Symtab.lookup (get thy_ctxt) name of |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
342 |
NONE => NONE |
16458 | 343 |
| 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
|
344 |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
345 |
val get_global_registration = |
16458 | 346 |
gen_get_registration (#3 o GlobalLocalesData.get) I; |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
347 |
val get_local_registration = |
16458 | 348 |
gen_get_registration LocalLocalesData.get ProofContext.theory_of; |
15596 | 349 |
|
18343 | 350 |
val test_global_registration = is_some oo get_global_registration; |
351 |
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
|
352 |
fun smart_test_registration ctxt id = |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
353 |
let |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
354 |
val thy = ProofContext.theory_of ctxt; |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
355 |
in |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
356 |
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
|
357 |
end; |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
358 |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
359 |
|
15837 | 360 |
(* 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
|
361 |
|
16458 | 362 |
fun gen_put_registration map_data thy_of (name, ps) attn thy_ctxt = |
15837 | 363 |
map_data (fn regs => |
364 |
let |
|
16458 | 365 |
val thy = thy_of thy_ctxt; |
18343 | 366 |
val reg = the_default Registrations.empty (Symtab.lookup regs name); |
16458 | 367 |
val (reg', sups) = Registrations.insert thy (ps, attn) reg; |
15837 | 368 |
val _ = if not (null sups) then warning |
369 |
("Subsumed interpretation(s) of locale " ^ |
|
16458 | 370 |
quote (extern thy name) ^ |
15837 | 371 |
"\nby interpretation(s) with the following prefix(es):\n" ^ |
372 |
commas_quote (map (fn (_, ((s, _), _)) => s) sups)) |
|
373 |
else (); |
|
17412 | 374 |
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
|
375 |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
376 |
val put_global_registration = |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
377 |
gen_put_registration (fn f => |
16458 | 378 |
GlobalLocalesData.map (fn (space, locs, regs) => (space, locs, f regs))) I; |
15837 | 379 |
val put_local_registration = |
16458 | 380 |
gen_put_registration LocalLocalesData.map ProofContext.theory_of; |
15596 | 381 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
382 |
fun put_registration_in_locale name id thy = |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
383 |
let |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
384 |
val {predicate, import, elems, params, regs} = the_locale thy name; |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
385 |
in |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
386 |
put_locale name {predicate = predicate, import = import, |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
387 |
elems = elems, params = params, regs = regs @ [(id, [])]} thy |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
388 |
end; |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
389 |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
390 |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
391 |
(* add witness theorem to registration in theory or context, |
15596 | 392 |
ignored if registration not present *) |
393 |
||
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
394 |
fun add_witness (name, ps) thm = |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
395 |
Symtab.map_entry name (Registrations.add_witness ps thm); |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
396 |
|
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
397 |
fun add_global_witness id thm = |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
398 |
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
|
399 |
|
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
400 |
val add_local_witness = LocalLocalesData.map oo add_witness; |
15596 | 401 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
402 |
fun add_witness_in_locale name id thm thy = |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
403 |
let |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
404 |
val {predicate, import, elems, params, regs} = the_locale thy name; |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
405 |
fun add (id', thms) = |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
406 |
if id = id' then (id', thm :: thms) else (id', thms); |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
407 |
in |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
408 |
put_locale name {predicate = predicate, import = import, |
18137 | 409 |
elems = elems, params = params, regs = map add regs} thy |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
410 |
end; |
15596 | 411 |
|
14215
ebf291f3b449
Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents:
13629
diff
changeset
|
412 |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
413 |
(* printing of registrations *) |
15596 | 414 |
|
17138 | 415 |
fun gen_print_registrations get_regs mk_ctxt msg show_wits loc thy_ctxt = |
15596 | 416 |
let |
15703 | 417 |
val ctxt = mk_ctxt thy_ctxt; |
418 |
val thy = ProofContext.theory_of ctxt; |
|
419 |
||
420 |
val prt_term = Pretty.quote o ProofContext.pretty_term ctxt; |
|
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
421 |
fun prt_inst ts = |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
422 |
Pretty.enclose "(" ")" (Pretty.breaks (map prt_term ts)); |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
423 |
fun prt_attn (prfx, atts) = |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
424 |
Pretty.breaks (Pretty.str prfx :: Args.pretty_attribs ctxt atts); |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
425 |
fun prt_thm (_, th) = Pretty.quote (ProofContext.pretty_thm ctxt th); |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
426 |
fun prt_thms thms = |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
427 |
Pretty.enclose "[" "]" (Pretty.breaks (map prt_thm thms)); |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
428 |
fun prt_reg (ts, (("", []), thms)) = |
17138 | 429 |
if show_wits |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
430 |
then Pretty.block [prt_inst ts, Pretty.fbrk, prt_thms thms] |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
431 |
else prt_inst ts |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
432 |
| prt_reg (ts, (attn, thms)) = |
17138 | 433 |
if show_wits |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
434 |
then Pretty.block ((prt_attn attn @ |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
435 |
[Pretty.str ":", Pretty.brk 1, prt_inst ts, Pretty.fbrk, |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
436 |
prt_thms thms])) |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
437 |
else Pretty.block ((prt_attn attn @ |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
438 |
[Pretty.str ":", Pretty.brk 1, prt_inst ts])); |
15703 | 439 |
|
16458 | 440 |
val loc_int = intern thy loc; |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
441 |
val regs = get_regs thy_ctxt; |
17412 | 442 |
val loc_regs = Symtab.lookup regs loc_int; |
15596 | 443 |
in |
444 |
(case loc_regs of |
|
17355 | 445 |
NONE => Pretty.str ("no interpretations in " ^ msg) |
15763
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
ballarin
parents:
15749
diff
changeset
|
446 |
| SOME r => let |
15837 | 447 |
val r' = Registrations.dest r; |
15763
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
ballarin
parents:
15749
diff
changeset
|
448 |
val r'' = Library.sort_wrt (fn (_, ((prfx, _), _)) => prfx) r'; |
17355 | 449 |
in Pretty.big_list ("interpretations in " ^ msg ^ ":") |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
450 |
(map prt_reg r'') |
15763
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
ballarin
parents:
15749
diff
changeset
|
451 |
end) |
15596 | 452 |
|> Pretty.writeln |
453 |
end; |
|
454 |
||
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
455 |
val print_global_registrations = |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
456 |
gen_print_registrations (#3 o GlobalLocalesData.get) |
15703 | 457 |
ProofContext.init "theory"; |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
458 |
val print_local_registrations' = |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
459 |
gen_print_registrations LocalLocalesData.get |
15703 | 460 |
I "context"; |
17138 | 461 |
fun print_local_registrations show_wits loc ctxt = |
462 |
(print_global_registrations show_wits loc (ProofContext.theory_of ctxt); |
|
463 |
print_local_registrations' show_wits loc ctxt); |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
464 |
|
15596 | 465 |
|
12277 | 466 |
(* diagnostics *) |
12273 | 467 |
|
12277 | 468 |
fun err_in_locale ctxt msg ids = |
469 |
let |
|
16458 | 470 |
val thy = ProofContext.theory_of ctxt; |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
471 |
fun prt_id (name, parms) = |
16458 | 472 |
[Pretty.block (Pretty.breaks (map Pretty.str (extern thy name :: parms)))]; |
15570 | 473 |
val prt_ids = List.concat (separate [Pretty.str " +", Pretty.brk 1] (map prt_id ids)); |
12502 | 474 |
val err_msg = |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
475 |
if forall (equal "" o #1) ids then msg |
12502 | 476 |
else msg ^ "\n" ^ Pretty.string_of (Pretty.block |
477 |
(Pretty.str "The error(s) above occurred in locale:" :: Pretty.brk 1 :: prt_ids)); |
|
18678 | 478 |
in error err_msg end; |
12063 | 479 |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
480 |
fun err_in_locale' ctxt msg ids' = err_in_locale ctxt msg (map fst ids'); |
12277 | 481 |
|
482 |
||
12046 | 483 |
|
18137 | 484 |
(** witnesses -- protected facts **) |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
485 |
|
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
486 |
fun assume_protected thy t = |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
487 |
(t, Goal.protect (Thm.assume (Thm.cterm_of thy t))); |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
488 |
|
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
489 |
fun prove_protected thy t tac = |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
490 |
(t, Goal.prove thy [] [] (Logic.protect t) (fn _ => |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
491 |
Tactic.rtac Drule.protectI 1 THEN tac)); |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
492 |
|
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
493 |
val conclude_protected = Goal.conclude #> Goal.norm_hhf; |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
494 |
|
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
495 |
fun satisfy_protected pths thm = |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
496 |
let |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
497 |
fun satisfy chyp = |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
498 |
(case find_first (fn (t, _) => Thm.term_of chyp aconv t) pths of |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
499 |
NONE => I |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
500 |
| SOME (_, th) => Drule.implies_intr_protected [chyp] #> Goal.comp_hhf th); |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
501 |
in fold satisfy (#hyps (Thm.crep_thm thm)) thm end; |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
502 |
|
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
503 |
|
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
504 |
|
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
505 |
(** structured contexts: rename + merge + implicit type instantiation **) |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
506 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
507 |
(* parameter types *) |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
508 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
509 |
fun frozen_tvars ctxt Ts = |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
510 |
let |
16861 | 511 |
val tvars = rev (fold Term.add_tvarsT Ts []); |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
512 |
val tfrees = map TFree |
14695 | 513 |
(Term.invent_names (ProofContext.used_types ctxt) "'a" (length tvars) ~~ map #2 tvars); |
18343 | 514 |
in map (fn ((xi, S), T) => (xi, (S, T))) (tvars ~~ tfrees) end; |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
515 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
516 |
fun unify_frozen ctxt maxidx Ts Us = |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
517 |
let |
18343 | 518 |
fun paramify NONE i = (NONE, i) |
519 |
| 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
|
520 |
|
18343 | 521 |
val (Ts', maxidx') = fold_map paramify Ts maxidx; |
522 |
val (Us', maxidx'') = fold_map paramify Us maxidx'; |
|
16947 | 523 |
val thy = ProofContext.theory_of ctxt; |
14215
ebf291f3b449
Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents:
13629
diff
changeset
|
524 |
|
18190 | 525 |
fun unify (SOME T, SOME U) env = (Sign.typ_unify thy (U, T) env |
526 |
handle Type.TUNIFY => raise TYPE ("unify_frozen: failed to unify types", [U, T], [])) |
|
527 |
| unify _ env = env; |
|
528 |
val (unifier, _) = fold unify (Ts' ~~ Us') (Vartab.empty, maxidx''); |
|
15570 | 529 |
val Vs = map (Option.map (Envir.norm_type unifier)) Us'; |
530 |
val unifier' = Vartab.extend (unifier, frozen_tvars ctxt (List.mapPartial I Vs)); |
|
531 |
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
|
532 |
|
17496 | 533 |
fun params_of elemss = gen_distinct (eq_fst (op =)) (List.concat (map (snd o fst) elemss)); |
534 |
fun params_of' elemss = gen_distinct (eq_fst (op =)) (List.concat (map (snd o fst o fst) elemss)); |
|
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
535 |
fun params_syn_of syn elemss = |
17496 | 536 |
gen_distinct (eq_fst (op =)) (List.concat (map (snd o fst) elemss)) |> |
18343 | 537 |
map (apfst (fn x => (x, the (Symtab.lookup syn x)))); |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
538 |
|
14508
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
539 |
|
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
540 |
(* CB: param_types has the following type: |
15531 | 541 |
('a * 'b option) list -> ('a * 'b) list *) |
15570 | 542 |
fun param_types ps = List.mapPartial (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
|
543 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
544 |
|
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
545 |
fun merge_syntax ctxt ids ss = Symtab.merge (op =) ss |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
546 |
handle Symtab.DUPS xs => err_in_locale ctxt |
16105 | 547 |
("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
|
548 |
|
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
549 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
550 |
(* Distinction of assumed vs. derived identifiers. |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
551 |
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
|
552 |
assumptions of the specification fragment (for locales with |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
553 |
predicates). The latter have witness theorems relating assumptions of the |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
554 |
specification fragment to assumptions of other (assumed) specification |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
555 |
fragments. *) |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
556 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
557 |
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
|
558 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
559 |
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
|
560 |
| map_mode f (Derived x) = Derived (f x); |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
561 |
|
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
562 |
|
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
563 |
(* flatten expressions *) |
11896 | 564 |
|
12510 | 565 |
local |
12502 | 566 |
|
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
567 |
fun unique_parms ctxt elemss = |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
568 |
let |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
569 |
val param_decls = |
15570 | 570 |
List.concat (map (fn (((name, (ps, qs)), _), _) => map (rpair (name, ps)) qs) elemss) |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
571 |
|> Symtab.make_multi |> Symtab.dest; |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
572 |
in |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
573 |
(case find_first (fn (_, ids) => length ids > 1) param_decls of |
15531 | 574 |
SOME (q, ids) => err_in_locale ctxt ("Multiple declaration of parameter " ^ quote q) |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
575 |
(map (apsnd (map fst)) ids) |
15531 | 576 |
| NONE => map (apfst (apfst (apsnd #1))) elemss) |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
577 |
end; |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
578 |
|
18137 | 579 |
fun unify_parms ctxt fixed_parms raw_parmss = |
12502 | 580 |
let |
16458 | 581 |
val thy = ProofContext.theory_of ctxt; |
12502 | 582 |
val maxidx = length raw_parmss; |
583 |
val idx_parmss = (0 upto maxidx - 1) ~~ raw_parmss; |
|
584 |
||
585 |
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
|
586 |
fun varify_parms (i, ps) = map (apsnd (varify i)) (param_types ps); |
15570 | 587 |
val parms = fixed_parms @ List.concat (map varify_parms idx_parmss); |
12502 | 588 |
|
18137 | 589 |
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
|
590 |
handle Type.TUNIFY => |
18137 | 591 |
let val prt = Sign.string_of_typ thy in |
592 |
raise TYPE ("unify_parms: failed to unify types " ^ |
|
593 |
prt U ^ " and " ^ prt T, [U, T], []) |
|
594 |
end; |
|
595 |
fun unify_list (T :: Us) = fold (unify T) Us |
|
596 |
| unify_list [] = I; |
|
597 |
val (unifier, _) = fold unify_list (map #2 (Symtab.dest (Symtab.make_multi parms))) |
|
598 |
(Vartab.empty, maxidx); |
|
12502 | 599 |
|
17496 | 600 |
val parms' = map (apsnd (Envir.norm_type unifier)) (gen_distinct (eq_fst (op =)) parms); |
12502 | 601 |
val unifier' = Vartab.extend (unifier, frozen_tvars ctxt (map #2 parms')); |
602 |
||
603 |
fun inst_parms (i, ps) = |
|
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
604 |
foldr Term.add_typ_tfrees [] (List.mapPartial snd ps) |
15570 | 605 |
|> List.mapPartial (fn (a, S) => |
12502 | 606 |
let val T = Envir.norm_type unifier' (TVar ((a, i), S)) |
18137 | 607 |
in if T = TFree (a, S) then NONE else SOME (a, T) end) |
608 |
|> Symtab.make; |
|
609 |
in map inst_parms idx_parmss end; |
|
12502 | 610 |
|
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
611 |
in |
12502 | 612 |
|
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
613 |
fun unify_elemss _ _ [] = [] |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
614 |
| unify_elemss _ [] [elems] = [elems] |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
615 |
| unify_elemss ctxt fixed_parms elemss = |
12502 | 616 |
let |
18137 | 617 |
val thy = ProofContext.theory_of ctxt; |
17756 | 618 |
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
|
619 |
fun inst ((((name, ps), mode), elems), env) = |
18137 | 620 |
(((name, map (apsnd (Option.map (Element.instT_type env))) ps), |
621 |
map_mode (map (fn (t, th) => |
|
622 |
(Element.instT_term env t, Element.instT_thm thy env th))) mode), |
|
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 |
(* like unify_elemss, but does not touch mode, additional |
16736
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
627 |
parameter c_parms for enforcing further constraints (eg. syntax) *) |
18343 | 628 |
(* FIXME avoid code duplication *) |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
629 |
|
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
630 |
fun unify_elemss' _ _ [] [] = [] |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
631 |
| unify_elemss' _ [] [elems] [] = [elems] |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
632 |
| unify_elemss' ctxt fixed_parms elemss c_parms = |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
633 |
let |
18137 | 634 |
val thy = ProofContext.theory_of ctxt; |
635 |
val envs = |
|
636 |
unify_parms ctxt fixed_parms (map (snd o fst o fst) elemss @ map single c_parms); |
|
17033 | 637 |
fun inst ((((name, ps), (ps', mode)), elems), env) = |
18137 | 638 |
(((name, map (apsnd (Option.map (Element.instT_type env))) ps), (ps', mode)), |
639 |
map (Element.instT_ctxt thy env) elems); |
|
18343 | 640 |
in map inst (elemss ~~ Library.take (length elemss, envs)) end; |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
641 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
642 |
|
15596 | 643 |
(* flatten_expr: |
644 |
Extend list of identifiers by those new in locale expression expr. |
|
645 |
Compute corresponding list of lists of locale elements (one entry per |
|
646 |
identifier). |
|
647 |
||
648 |
Identifiers represent locale fragments and are in an extended form: |
|
649 |
((name, ps), (ax_ps, axs)) |
|
650 |
(name, ps) is the locale name with all its parameters. |
|
651 |
(ax_ps, axs) is the locale axioms with its parameters; |
|
652 |
axs are always taken from the top level of the locale hierarchy, |
|
653 |
hence axioms may contain additional parameters from later fragments: |
|
654 |
ps subset of ax_ps. axs is either singleton or empty. |
|
655 |
||
656 |
Elements are enriched by identifier-like information: |
|
657 |
(((name, ax_ps), axs), elems) |
|
658 |
The parameters in ax_ps are the axiom parameters, but enriched by type |
|
659 |
info: now each entry is a pair of string and typ option. Axioms are |
|
660 |
type-instantiated. |
|
661 |
||
662 |
*) |
|
663 |
||
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
664 |
fun flatten_expr ctxt ((prev_idents, prev_syntax), expr) = |
12014 | 665 |
let |
666 |
val thy = ProofContext.theory_of ctxt; |
|
12263 | 667 |
|
15531 | 668 |
fun renaming (SOME x :: xs) (y :: ys) = (y, x) :: renaming xs ys |
669 |
| renaming (NONE :: xs) (y :: ys) = renaming xs ys |
|
12273 | 670 |
| renaming [] _ = [] |
18678 | 671 |
| renaming xs [] = error ("Too many arguments in renaming: " ^ |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
672 |
commas (map (fn NONE => "_" | SOME x => quote (fst x)) xs)); |
12289 | 673 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
674 |
fun rename_parms top ren ((name, ps), (parms, mode)) = |
18137 | 675 |
let val ps' = map (Element.rename ren) ps in |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
676 |
(case duplicates ps' of |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
677 |
[] => ((name, ps'), |
18137 | 678 |
if top then (map (Element.rename ren) parms, |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
679 |
map_mode (map (fn (t, th) => |
18137 | 680 |
(Element.rename_term ren t, Element.rename_thm ren th))) mode) |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
681 |
else (parms, mode)) |
12289 | 682 |
| dups => err_in_locale ctxt ("Duplicate parameters: " ^ commas_quote dups) [(name, ps')]) |
683 |
end; |
|
12263 | 684 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
685 |
(* add registrations of (name, ps), recursively; |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
686 |
adjust hyps of witness theorems *) |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
687 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
688 |
fun add_regs (name, ps) (ths, ids) = |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
689 |
let |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
690 |
val {params, regs, ...} = the_locale thy name; |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
691 |
val ps' = map #1 (#1 params); |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
692 |
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
|
693 |
(* dummy syntax, since required by rename *) |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
694 |
val ps'' = map (fn ((p, _), (_, T)) => (p, T)) (ps ~~ ps'); |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
695 |
val [env] = unify_parms ctxt ps [map (apsnd SOME) ps'']; |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
696 |
(* propagate parameter types, to keep them consistent *) |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
697 |
val regs' = map (fn ((name, ps), ths) => |
18137 | 698 |
((name, map (Element.rename ren) ps), ths)) regs; |
17496 | 699 |
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
|
700 |
val new_ids = map fst new_regs; |
17485 | 701 |
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
|
702 |
|
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
703 |
val new_ths = new_regs |> map (fn (_, ths') => ths' |> map (fn (t, th) => |
18137 | 704 |
(t |> Element.instT_term env |> Element.rename_term ren, |
705 |
th |> Element.instT_thm thy env |> Element.rename_thm ren |> satisfy_protected ths))); |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
706 |
val new_ids' = map (fn (id, ths) => |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
707 |
(id, ([], Derived ths))) (new_ids ~~ new_ths); |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
708 |
in |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
709 |
fold add_regs new_idTs (ths @ List.concat new_ths, ids @ new_ids') |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
710 |
end; |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
711 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
712 |
(* distribute top-level axioms over assumed ids *) |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
713 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
714 |
fun axiomify all_ps ((name, parms), (_, Assumed _)) axioms = |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
715 |
let |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
716 |
val {elems, ...} = the_locale thy name; |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
717 |
val ts = List.concat (map |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
718 |
(fn (Assumes asms, _) => List.concat (map (map #1 o #2) asms) |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
719 |
| _ => []) |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
720 |
elems); |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
721 |
val (axs1, axs2) = splitAt (length ts, axioms); |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
722 |
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
|
723 |
| axiomify all_ps (id, (_, Derived ths)) axioms = |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
724 |
((id, (all_ps, Derived ths)), axioms); |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
725 |
|
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
726 |
(* identifiers of an expression *) |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
727 |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
728 |
fun identify top (Locale name) = |
15596 | 729 |
(* 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
|
730 |
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
|
731 |
a list of axioms relating to the identifier, axs is empty unless |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
732 |
identify at top level (top = true); |
14215
ebf291f3b449
Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents:
13629
diff
changeset
|
733 |
parms is accumulated list of parameters *) |
12289 | 734 |
let |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
735 |
val {predicate = (_, axioms), import, params, ...} = |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
736 |
the_locale thy name; |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
737 |
val ps = map (#1 o #1) (#1 params); |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
738 |
val (ids', parms', _) = identify false import; |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
739 |
(* acyclic import dependencies *) |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
740 |
val ids'' = ids' @ [((name, ps), ([], Assumed []))]; |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
741 |
val (_, ids''') = add_regs (name, map #1 (#1 params)) ([], ids''); |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
742 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
743 |
val ids_ax = if top then fst |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
744 |
(fold_map (axiomify ps) ids''' axioms) |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
745 |
else ids'''; |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
746 |
val syn = Symtab.make (map (apfst fst) (#1 params)); |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
747 |
in (ids_ax, merge_lists parms' ps, syn) end |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
748 |
| identify top (Rename (e, xs)) = |
12273 | 749 |
let |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
750 |
val (ids', parms', syn') = identify top e; |
12839 | 751 |
val ren = renaming xs parms' |
18678 | 752 |
handle ERROR msg => err_in_locale' ctxt msg ids'; |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
753 |
|
17496 | 754 |
val ids'' = gen_distinct (eq_fst (op =)) (map (rename_parms top ren) ids'); |
15570 | 755 |
val parms'' = distinct (List.concat (map (#2 o #1) ids'')); |
18137 | 756 |
val syn'' = syn' |> Symtab.dest |> map (Element.rename_var ren) |> Symtab.make; |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
757 |
(* check for conflicting syntax? *) |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
758 |
in (ids'', parms'', syn'') end |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
759 |
| identify top (Merge es) = |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
760 |
fold (fn e => fn (ids, parms, syn) => |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
761 |
let |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
762 |
val (ids', parms', syn') = identify top e |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
763 |
in |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
764 |
(merge_alists ids ids', |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
765 |
merge_lists parms parms', |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
766 |
merge_syntax ctxt ids' (syn, syn')) |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
767 |
end) |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
768 |
es ([], [], Symtab.empty); |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
769 |
|
12014 | 770 |
|
18137 | 771 |
(* CB: enrich identifiers by parameter types and |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
772 |
the corresponding elements (with renamed parameters), |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
773 |
also takes care of parameter syntax *) |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
774 |
|
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
775 |
fun eval syn ((name, xs), axs) = |
12273 | 776 |
let |
13308 | 777 |
val {params = (ps, qs), elems, ...} = the_locale thy name; |
16620
2a7f46324218
Proper treatment of beta-redexes in witness theorems.
ballarin
parents:
16458
diff
changeset
|
778 |
val ps' = map (apsnd SOME o #1) ps; |
18671 | 779 |
fun lookup_syn x = (case Symtab.lookup syn x of SOME Structure => NONE | opt => opt); |
780 |
val ren = map #1 ps' ~~ map (fn x => (x, lookup_syn x)) xs; |
|
13308 | 781 |
val (params', elems') = |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
782 |
if null ren then ((ps', qs), map #1 elems) |
18137 | 783 |
else ((map (apfst (Element.rename ren)) ps', map (Element.rename ren) qs), |
784 |
map (Element.rename_ctxt ren o #1) elems); |
|
785 |
val elems'' = elems' |> map (Element.map_ctxt |
|
786 |
{var = I, typ = I, term = I, fact = I, attrib = I, |
|
787 |
name = NameSpace.qualified (space_implode "_" xs)}); |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
788 |
in (((name, params'), axs), elems'') end; |
12307 | 789 |
|
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
790 |
(* type constraint for renamed parameter with syntax *) |
18343 | 791 |
fun mixfix_type mx = |
18671 | 792 |
SOME (Type.freeze_type (#1 (TypeInfer.paramify_dummies (TypeInfer.mixfixT mx) 0))); |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
793 |
|
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
794 |
(* compute identifiers and syntax, merge with previous ones *) |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
795 |
val (ids, _, syn) = identify true expr; |
17496 | 796 |
val idents = gen_rems (eq_fst (op =)) (ids, prev_idents); |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
797 |
val syntax = merge_syntax ctxt ids (syn, prev_syntax); |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
798 |
(* add types to params, check for unique params and unify them *) |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
799 |
val raw_elemss = unique_parms ctxt (map (eval syntax) idents); |
18671 | 800 |
val elemss = unify_elemss' ctxt [] raw_elemss (map (apsnd mixfix_type) (Symtab.dest syntax)); |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
801 |
(* replace params in ids by params from axioms, |
17033 | 802 |
adjust types in mode *) |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
803 |
val all_params' = params_of' elemss; |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
804 |
val all_params = param_types all_params'; |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
805 |
val elemss' = map (fn (((name, _), (ps, mode)), elems) => |
17485 | 806 |
(((name, map (fn p => (p, AList.lookup (op =) all_params p)) ps), mode), elems)) |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
807 |
elemss; |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
808 |
fun inst_th (t, th) = let |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
809 |
val {hyps, prop, ...} = Thm.rep_thm th; |
16861 | 810 |
val ps = map (apsnd SOME) (fold Term.add_frees (prop :: hyps) []); |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
811 |
val [env] = unify_parms ctxt all_params [ps]; |
18137 | 812 |
val t' = Element.instT_term env t; |
813 |
val th' = Element.instT_thm thy env th; |
|
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
814 |
in (t', th') end; |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
815 |
val final_elemss = map (fn ((id, mode), elems) => |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
816 |
((id, map_mode (map inst_th) mode), elems)) elemss'; |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
817 |
|
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
818 |
in ((prev_idents @ idents, syntax), final_elemss) end; |
12046 | 819 |
|
12510 | 820 |
end; |
821 |
||
12070 | 822 |
|
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
823 |
(* activate elements *) |
12273 | 824 |
|
12510 | 825 |
local |
826 |
||
18671 | 827 |
fun axioms_export axs _ hyps = |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
828 |
satisfy_protected axs |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
829 |
#> Drule.implies_intr_list (Library.drop (length axs, hyps)) |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
830 |
#> Seq.single; |
12263 | 831 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
832 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
833 |
(* 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
|
834 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
835 |
fun activate_elem _ ((ctxt, mode), Fixes fixes) = |
18671 | 836 |
((ctxt |> ProofContext.add_fixes_i fixes |> snd, mode), []) |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
837 |
| activate_elem _ ((ctxt, mode), Constrains _) = |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
838 |
((ctxt, mode), []) |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
839 |
| activate_elem _ ((ctxt, Assumed axs), Assumes asms) = |
13399
c136276dc847
support locale ``views'' (for cumulative predicates);
wenzelm
parents:
13394
diff
changeset
|
840 |
let |
18728 | 841 |
val asms' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) asms; |
15703 | 842 |
val ts = List.concat (map (map #1 o #2) asms'); |
843 |
val (ps, qs) = splitAt (length ts, axs); |
|
17856 | 844 |
val (_, ctxt') = |
18671 | 845 |
ctxt |> fold ProofContext.fix_frees ts |
846 |
|> ProofContext.add_assms_i (axioms_export ps) asms'; |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
847 |
in ((ctxt', Assumed qs), []) end |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
848 |
| activate_elem _ ((ctxt, Derived ths), Assumes asms) = |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
849 |
((ctxt, Derived ths), []) |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
850 |
| activate_elem _ ((ctxt, Assumed axs), Defines defs) = |
15596 | 851 |
let |
18728 | 852 |
val defs' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) defs; |
17856 | 853 |
val (_, ctxt') = |
18671 | 854 |
ctxt |> ProofContext.add_assms_i ProofContext.def_export |
15703 | 855 |
(defs' |> map (fn ((name, atts), (t, ps)) => |
13399
c136276dc847
support locale ``views'' (for cumulative predicates);
wenzelm
parents:
13394
diff
changeset
|
856 |
let val (c, t') = ProofContext.cert_def ctxt t |
c136276dc847
support locale ``views'' (for cumulative predicates);
wenzelm
parents:
13394
diff
changeset
|
857 |
in ((if name = "" then Thm.def_name c else name, atts), [(t', (ps, []))]) end)) |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
858 |
in ((ctxt', Assumed axs), []) end |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
859 |
| activate_elem _ ((ctxt, Derived ths), Defines defs) = |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
860 |
((ctxt, Derived ths), []) |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
861 |
| activate_elem is_ext ((ctxt, mode), Notes facts) = |
15596 | 862 |
let |
18728 | 863 |
val facts' = Attrib.map_facts (Attrib.attribute_i (ProofContext.theory_of ctxt)) facts; |
17856 | 864 |
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
|
865 |
in ((ctxt', mode), if is_ext then res else []) end; |
12502 | 866 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
867 |
fun activate_elems (((name, ps), mode), elems) ctxt = |
17033 | 868 |
let |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
869 |
val thy = ProofContext.theory_of ctxt; |
17033 | 870 |
val ((ctxt', _), res) = |
871 |
foldl_map (activate_elem (name = "")) ((ProofContext.qualified_names ctxt, mode), elems) |
|
18678 | 872 |
handle ERROR msg => err_in_locale ctxt msg [(name, map fst ps)] |
15696 | 873 |
val ctxt'' = if name = "" then ctxt' |
874 |
else let |
|
875 |
val ps' = map (fn (n, SOME T) => Free (n, T)) ps; |
|
876 |
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
|
877 |
in case mode of |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
878 |
Assumed axs => |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
879 |
fold (add_local_witness (name, ps') o assume_protected thy o #1) axs ctxt'' |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
880 |
| Derived ths => fold (add_local_witness (name, ps')) ths ctxt'' |
15696 | 881 |
end |
16144 | 882 |
in (ProofContext.restore_naming ctxt ctxt'', res) end; |
13399
c136276dc847
support locale ``views'' (for cumulative predicates);
wenzelm
parents:
13394
diff
changeset
|
883 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
884 |
fun activate_elemss prep_facts = |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
885 |
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
|
886 |
let |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
887 |
val elems = map (prep_facts ctxt) raw_elems; |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
888 |
val (ctxt', res) = apsnd List.concat |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
889 |
(activate_elems (((name, ps), mode), elems) ctxt); |
18137 | 890 |
val elems' = elems |> map (Element.map_ctxt |
891 |
{name = I, var = I, typ = I, term = I, fact = I, attrib = Args.closure}); |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
892 |
in ((((name, ps), elems'), res), ctxt') end); |
12834
e5bec3268932
added locale_facts(_i) interface (useful for simple ML proof tools);
wenzelm
parents:
12806
diff
changeset
|
893 |
|
12546 | 894 |
in |
895 |
||
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
896 |
(* CB: activate_facts prep_facts (ctxt, elemss), |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
897 |
where elemss is a list of pairs consisting of identifiers and |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
898 |
context elements, extends ctxt by the context elements yielding |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
899 |
ctxt' and returns (ctxt', (elemss', facts)). |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
900 |
Identifiers in the argument are of the form ((name, ps), axs) and |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
901 |
assumptions use the axioms in the identifiers to set up exporters |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
902 |
in ctxt'. elemss' does not contain identifiers and is obtained |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
903 |
from elemss and the intermediate context with prep_facts. |
15703 | 904 |
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
|
905 |
the internal/external markers from elemss. *) |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
906 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
907 |
fun activate_facts prep_facts (ctxt, args) = |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
908 |
let |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
909 |
val (res, ctxt') = activate_elemss prep_facts args ctxt; |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
910 |
in |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
911 |
(ctxt', apsnd List.concat (split_list res)) |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
912 |
end; |
12546 | 913 |
|
15703 | 914 |
fun activate_note prep_facts (ctxt, args) = |
915 |
let |
|
916 |
val (ctxt', ([(_, [Notes args'])], facts)) = |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
917 |
activate_facts prep_facts (ctxt, [((("", []), Assumed []), [Ext (Notes args)])]); |
15703 | 918 |
in (ctxt', (args', facts)) end; |
919 |
||
12510 | 920 |
end; |
921 |
||
12307 | 922 |
|
15696 | 923 |
|
18137 | 924 |
(** prepare locale elements **) |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
925 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
926 |
(* expressions *) |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
927 |
|
16458 | 928 |
fun intern_expr thy (Locale xname) = Locale (intern thy xname) |
929 |
| intern_expr thy (Merge exprs) = Merge (map (intern_expr thy) exprs) |
|
930 |
| 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
|
931 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
932 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
933 |
(* propositions and bindings *) |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
934 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
935 |
(* flatten (ctxt, prep_expr) ((ids, syn), expr) |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
936 |
normalises expr (which is either a locale |
14508
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
937 |
expression or a single context element) wrt. |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
938 |
to the list ids of already accumulated identifiers. |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
939 |
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
|
940 |
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
|
941 |
context elements generated from expr. |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
942 |
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
|
943 |
is an extension of syn. |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
944 |
For details, see flatten_expr. |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
945 |
|
15596 | 946 |
Additionally, for a locale expression, the elems are grouped into a single |
947 |
Int; individual context elements are marked Ext. In this case, the |
|
948 |
identifier-like information of the element is as follows: |
|
949 |
- for Fixes: (("", ps), []) where the ps have type info NONE |
|
950 |
- for other elements: (("", []), []). |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
951 |
The implementation of activate_facts relies on identifier names being |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
952 |
empty strings for external elements. |
15596 | 953 |
*) |
14508
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
954 |
|
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
955 |
fun flatten (ctxt, _) ((ids, syn), Elem (Fixes fixes)) = let |
18137 | 956 |
val ids' = ids @ [(("", map #1 fixes), ([], Assumed []))] |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
957 |
in |
18137 | 958 |
((ids', |
959 |
merge_syntax ctxt ids' |
|
960 |
(syn, Symtab.make (map (fn fx => (#1 fx, #3 fx)) fixes)) |
|
961 |
handle Symtab.DUPS xs => err_in_locale ctxt |
|
962 |
("Conflicting syntax for parameters: " ^ commas_quote xs) |
|
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
963 |
(map #1 ids')), |
18137 | 964 |
[((("", map (rpair NONE o #1) fixes), Assumed []), Ext (Fixes fixes))]) |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
965 |
end |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
966 |
| flatten _ ((ids, syn), Elem elem) = |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
967 |
((ids @ [(("", []), ([], Assumed []))], syn), [((("", []), Assumed []), Ext elem)]) |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
968 |
| flatten (ctxt, prep_expr) ((ids, syn), Expr expr) = |
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
969 |
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
|
970 |
|
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
971 |
local |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
972 |
|
12839 | 973 |
local |
974 |
||
12727 | 975 |
fun declare_int_elem (ctxt, Fixes fixes) = |
18671 | 976 |
(ctxt |> ProofContext.add_fixes_i (map (fn (x, T, mx) => |
977 |
(x, Option.map (Term.map_type_tfree (TypeInfer.param 0)) T, mx)) fixes) |> snd, []) |
|
12727 | 978 |
| declare_int_elem (ctxt, _) = (ctxt, []); |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
979 |
|
18671 | 980 |
fun declare_ext_elem prep_vars (ctxt, Fixes fixes) = |
981 |
let val (vars, _) = prep_vars fixes ctxt |
|
982 |
in (ctxt |> ProofContext.add_fixes_i vars |> snd, []) end |
|
983 |
| declare_ext_elem prep_vars (ctxt, Constrains csts) = |
|
984 |
let val (_, ctxt') = prep_vars (map (fn (x, T) => (x, SOME T, NoSyn)) csts) ctxt |
|
985 |
in (ctxt', []) end |
|
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
986 |
| declare_ext_elem _ (ctxt, Assumes asms) = (ctxt, map #2 asms) |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
987 |
| declare_ext_elem _ (ctxt, Defines defs) = (ctxt, map (fn (_, (t, ps)) => [(t, (ps, []))]) defs) |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
988 |
| declare_ext_elem _ (ctxt, Notes facts) = (ctxt, []); |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
989 |
|
18671 | 990 |
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
|
991 |
let val (ctxt', propps) = |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
992 |
(case elems of |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
993 |
Int es => foldl_map declare_int_elem (ctxt, es) |
18671 | 994 |
| Ext e => foldl_map (declare_ext_elem prep_vars) (ctxt, [e])) |
18678 | 995 |
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
|
996 |
in (ctxt', propps) end |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
997 |
| declare_elems _ (ctxt, ((_, Derived _), elems)) = (ctxt, []); |
12727 | 998 |
|
12839 | 999 |
in |
1000 |
||
18671 | 1001 |
fun declare_elemss prep_vars fixed_params raw_elemss ctxt = |
12727 | 1002 |
let |
14215
ebf291f3b449
Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents:
13629
diff
changeset
|
1003 |
(* 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
|
1004 |
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
|
1005 |
distinct from types of parameters in target (fixed_params). *) |
ebf291f3b449
Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents:
13629
diff
changeset
|
1006 |
val ctxt_with_fixed = |
16028 | 1007 |
fold ProofContext.declare_term (map Free fixed_params) ctxt; |
12727 | 1008 |
val int_elemss = |
1009 |
raw_elemss |
|
15570 | 1010 |
|> List.mapPartial (fn (id, Int es) => SOME (id, es) | _ => NONE) |
14215
ebf291f3b449
Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents:
13629
diff
changeset
|
1011 |
|> unify_elemss ctxt_with_fixed fixed_params; |
12727 | 1012 |
val (_, raw_elemss') = |
1013 |
foldl_map (fn ((_, es) :: elemss, (id, Int _)) => (elemss, (id, Int es)) | x => x) |
|
1014 |
(int_elemss, raw_elemss); |
|
18671 | 1015 |
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
|
1016 |
|
12839 | 1017 |
end; |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1018 |
|
12839 | 1019 |
local |
1020 |
||
1021 |
val norm_term = Envir.beta_norm oo Term.subst_atomic; |
|
1022 |
||
16458 | 1023 |
fun abstract_thm thy eq = |
1024 |
Thm.assume (Thm.cterm_of thy eq) |> Drule.gen_all |> Drule.abs_def; |
|
12502 | 1025 |
|
18190 | 1026 |
fun bind_def ctxt (name, ps) eq (xs, env, ths) = |
12839 | 1027 |
let |
18782 | 1028 |
val ((y, T), b) = ProofContext.abs_def eq; |
13308 | 1029 |
val b' = norm_term env b; |
16458 | 1030 |
val th = abstract_thm (ProofContext.theory_of ctxt) eq; |
13308 | 1031 |
fun err msg = err_in_locale ctxt (msg ^ ": " ^ quote y) [(name, map fst ps)]; |
12839 | 1032 |
in |
13308 | 1033 |
conditional (exists (equal y o #1) xs) (fn () => |
1034 |
err "Attempt to define previously specified variable"); |
|
1035 |
conditional (exists (fn (Free (y', _), _) => y = y' | _ => false) env) (fn () => |
|
1036 |
err "Attempt to redefine variable"); |
|
16861 | 1037 |
(Term.add_frees b' xs, (Free (y, T), b') :: env, th :: ths) |
12839 | 1038 |
end; |
12575 | 1039 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1040 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1041 |
(* CB: for finish_elems (Int and Ext), |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1042 |
extracts specification, only of assumed elements *) |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1043 |
|
18190 | 1044 |
fun eval_text _ _ _ (Fixes _) text = text |
1045 |
| eval_text _ _ _ (Constrains _) text = text |
|
1046 |
| eval_text _ (_, Assumed _) is_ext (Assumes asms) |
|
1047 |
(((exts, exts'), (ints, ints')), (xs, env, defs)) = |
|
13394 | 1048 |
let |
15570 | 1049 |
val ts = List.concat (map (map #1 o #2) asms); |
13394 | 1050 |
val ts' = map (norm_term env) ts; |
1051 |
val spec' = |
|
1052 |
if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints')) |
|
1053 |
else ((exts, exts'), (ints @ ts, ints' @ ts')); |
|
16861 | 1054 |
in (spec', (fold Term.add_frees ts' xs, env, defs)) end |
18190 | 1055 |
| eval_text _ (_, Derived _) _ (Assumes _) text = text |
1056 |
| eval_text ctxt (id, Assumed _) _ (Defines defs) (spec, binds) = |
|
1057 |
(spec, fold (bind_def ctxt id o #1 o #2) defs binds) |
|
1058 |
| eval_text _ (_, Derived _) _ (Defines _) text = text |
|
1059 |
| eval_text _ _ _ (Notes _) text = text; |
|
13308 | 1060 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1061 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1062 |
(* for finish_elems (Int), |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1063 |
remove redundant elements of derived identifiers, |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1064 |
turn assumptions and definitions into facts, |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1065 |
adjust hypotheses of facts using witness theorems *) |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1066 |
|
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
1067 |
fun finish_derived _ _ (Assumed _) (Fixes fixes) = SOME (Fixes fixes) |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
1068 |
| finish_derived _ _ (Assumed _) (Constrains csts) = SOME (Constrains csts) |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
1069 |
| finish_derived _ _ (Assumed _) (Assumes asms) = SOME (Assumes asms) |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
1070 |
| finish_derived _ _ (Assumed _) (Defines defs) = SOME (Defines defs) |
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
1071 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1072 |
| finish_derived _ _ (Derived _) (Fixes _) = NONE |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1073 |
| finish_derived _ _ (Derived _) (Constrains _) = NONE |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1074 |
| finish_derived sign wits (Derived _) (Assumes asms) = asms |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
1075 |
|> map (apsnd (map (fn (a, _) => ([Thm.assume (cterm_of sign a)], [])))) |
18137 | 1076 |
|> Notes |> Element.map_ctxt_values I I (satisfy_protected wits) |> SOME |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1077 |
| finish_derived sign wits (Derived _) (Defines defs) = defs |
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
1078 |
|> map (apsnd (fn (d, _) => [([Thm.assume (cterm_of sign d)], [])])) |
18137 | 1079 |
|> Notes |> Element.map_ctxt_values I I (satisfy_protected wits) |> SOME |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1080 |
|
17096
8327b71282ce
Improved generation of witnesses in interpretation.
ballarin
parents:
17033
diff
changeset
|
1081 |
| finish_derived _ wits _ (Notes facts) = (Notes facts) |
18137 | 1082 |
|> Element.map_ctxt_values I I (satisfy_protected wits) |> SOME; |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1083 |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1084 |
(* CB: for finish_elems (Ext) *) |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1085 |
|
13308 | 1086 |
fun closeup _ false elem = elem |
1087 |
| closeup ctxt true elem = |
|
12839 | 1088 |
let |
13308 | 1089 |
fun close_frees t = |
1090 |
let val frees = rev (filter_out (ProofContext.is_fixed ctxt o #1) |
|
16861 | 1091 |
(Term.add_frees t [])) |
13308 | 1092 |
in Term.list_all_free (frees, t) end; |
1093 |
||
1094 |
fun no_binds [] = [] |
|
18678 | 1095 |
| no_binds _ = error "Illegal term bindings in locale element"; |
13308 | 1096 |
in |
1097 |
(case elem of |
|
1098 |
Assumes asms => Assumes (asms |> map (fn (a, propps) => |
|
1099 |
(a, map (fn (t, (ps, qs)) => (close_frees t, (no_binds ps, no_binds qs))) propps))) |
|
1100 |
| Defines defs => Defines (defs |> map (fn (a, (t, ps)) => |
|
1101 |
(a, (close_frees (#2 (ProofContext.cert_def ctxt t)), no_binds ps)))) |
|
1102 |
| e => e) |
|
1103 |
end; |
|
12839 | 1104 |
|
12502 | 1105 |
|
12839 | 1106 |
fun finish_ext_elem parms _ (Fixes fixes, _) = Fixes (map (fn (x, _, mx) => |
17271 | 1107 |
(x, AList.lookup (op =) parms x, mx)) fixes) |
16169
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16144
diff
changeset
|
1108 |
| finish_ext_elem parms _ (Constrains csts, _) = |
18671 | 1109 |
(* FIXME fails if x is not a parameter *) |
17271 | 1110 |
Constrains (map (fn (x, _) => (x, (the o AList.lookup (op =) parms) x)) csts) |
12839 | 1111 |
| finish_ext_elem _ close (Assumes asms, propp) = |
1112 |
close (Assumes (map #1 asms ~~ propp)) |
|
1113 |
| finish_ext_elem _ close (Defines defs, propp) = |
|
12727 | 1114 |
close (Defines (map #1 defs ~~ map (fn [(t, (ps, []))] => (t, ps)) propp)) |
12839 | 1115 |
| finish_ext_elem _ _ (Notes facts, _) = Notes facts; |
1116 |
||
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1117 |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1118 |
(* CB: finish_parms introduces type info from parms to identifiers *) |
15531 | 1119 |
(* CB: only needed for types that have been NONE so far??? |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1120 |
If so, which are these??? *) |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1121 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1122 |
fun finish_parms parms (((name, ps), mode), elems) = |
17485 | 1123 |
(((name, map (fn (x, _) => (x, AList.lookup (op =) parms x)) ps), mode), elems); |
12839 | 1124 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1125 |
fun finish_elems ctxt parms _ ((text, wits), ((id, Int e), _)) = |
12839 | 1126 |
let |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1127 |
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
|
1128 |
val wits' = case mode of Assumed _ => wits | Derived ths => wits @ ths; |
18190 | 1129 |
val text' = fold (eval_text ctxt id' false) es text; |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1130 |
val es' = List.mapPartial |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1131 |
(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
|
1132 |
in ((text', wits'), (id', map Int es')) end |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1133 |
| finish_elems ctxt parms do_close ((text, wits), ((id, Ext e), [propp])) = |
13308 | 1134 |
let |
1135 |
val e' = finish_ext_elem parms (closeup ctxt do_close) (e, propp); |
|
18190 | 1136 |
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
|
1137 |
in ((text', wits), (id, [Ext e'])) end |
12839 | 1138 |
|
1139 |
in |
|
12510 | 1140 |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1141 |
(* CB: only called by prep_elemss *) |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1142 |
|
13375 | 1143 |
fun finish_elemss ctxt parms do_close = |
1144 |
foldl_map (apsnd (finish_parms parms) o finish_elems ctxt parms do_close); |
|
12839 | 1145 |
|
1146 |
end; |
|
1147 |
||
16736
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
1148 |
|
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
1149 |
(* CB: type inference and consistency checks for locales. |
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
1150 |
|
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
1151 |
Works by building a context (through declare_elemss), extracting the |
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
1152 |
required information and adjusting the context elements (finish_elemss). |
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
1153 |
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
|
1154 |
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
|
1155 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1156 |
Only elements of assumed identifiers are considered. |
16736
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
1157 |
*) |
15127 | 1158 |
|
18671 | 1159 |
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
|
1160 |
let |
15127 | 1161 |
(* CB: contexts computed in the course of this function are discarded. |
1162 |
They are used for type inference and consistency checks only. *) |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1163 |
(* CB: fixed_params are the parameters (with types) of the target locale, |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1164 |
empty list if there is no target. *) |
14508
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1165 |
(* 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
|
1166 |
context elements, the latter marked as internal or external. *) |
18671 | 1167 |
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
|
1168 |
(* 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
|
1169 |
the fixes elements in raw_elemss, |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1170 |
raw_proppss contains assumptions and definitions from the |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1171 |
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
|
1172 |
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
|
1173 |
let |
e57731ba01dd
discontinued unflat in favour of burrow and burrow_split
haftmann
parents:
18421
diff
changeset
|
1174 |
(* 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
|
1175 |
(* CB: process patterns (conclusion and external elements only) *) |
e57731ba01dd
discontinued unflat in favour of burrow and burrow_split
haftmann
parents:
18421
diff
changeset
|
1176 |
val (ctxt, all_propp) = |
e57731ba01dd
discontinued unflat in favour of burrow and burrow_split
haftmann
parents:
18421
diff
changeset
|
1177 |
prepp (fold ProofContext.declare_term (map Free fixed_params) raw_ctxt, raw_concl @ raw_propp); |
e57731ba01dd
discontinued unflat in favour of burrow and burrow_split
haftmann
parents:
18421
diff
changeset
|
1178 |
(* CB: add type information from conclusion and external elements to context *) |
e57731ba01dd
discontinued unflat in favour of burrow and burrow_split
haftmann
parents:
18421
diff
changeset
|
1179 |
val ctxt = fold ProofContext.declare_term (List.concat (map (map fst) all_propp)) ctxt; |
e57731ba01dd
discontinued unflat in favour of burrow and burrow_split
haftmann
parents:
18421
diff
changeset
|
1180 |
(* 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
|
1181 |
val all_propp' = map2 (curry (op ~~)) |
e57731ba01dd
discontinued unflat in favour of burrow and burrow_split
haftmann
parents:
18421
diff
changeset
|
1182 |
(#1 (#2 (ProofContext.bind_propp_schematic_i (ctxt, all_propp)))) (map (map snd) all_propp); |
e57731ba01dd
discontinued unflat in favour of burrow and burrow_split
haftmann
parents:
18421
diff
changeset
|
1183 |
val (concl, propp) = splitAt (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
|
1184 |
in (propp, (ctxt, concl)) end |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1185 |
|
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1186 |
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
|
1187 |
(fold_burrow o fold_burrow) prep_prop raw_proppss (raw_ctxt, raw_concl); |
12502 | 1188 |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1189 |
(* CB: obtain all parameters from identifier part of raw_elemss *) |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1190 |
val xs = map #1 (params_of' raw_elemss); |
12727 | 1191 |
val typing = unify_frozen ctxt 0 |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1192 |
(map (ProofContext.default_type raw_ctxt) xs) |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1193 |
(map (ProofContext.default_type ctxt) xs); |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1194 |
val parms = param_types (xs ~~ typing); |
14508
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1195 |
(* CB: parms are the parameters from raw_elemss, with correct typing. *) |
12273 | 1196 |
|
14508
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1197 |
(* CB: extract information from assumes and defines elements |
16169
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16144
diff
changeset
|
1198 |
(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
|
1199 |
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
|
1200 |
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
|
1201 |
((((([], []), ([], [])), ([], [], [])), []), raw_elemss ~~ proppss); |
14508
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1202 |
(* CB: text has the following structure: |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1203 |
(((exts, exts'), (ints, ints')), (xs, env, defs)) |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1204 |
where |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1205 |
exts: external assumptions (terms in external assumes elements) |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1206 |
exts': dito, normalised wrt. env |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1207 |
ints: internal assumptions (terms in internal assumes elements) |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1208 |
ints': dito, normalised wrt. env |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1209 |
xs: the free variables in exts' and ints' and rhss of definitions, |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1210 |
this includes parameters except defined parameters |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1211 |
env: list of term pairs encoding substitutions, where the first term |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1212 |
is a free variable; substitutions represent defines elements and |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1213 |
the rhs is normalised wrt. the previous env |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1214 |
defs: theorems representing the substitutions from defines elements |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1215 |
(thms are normalised wrt. env). |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1216 |
elemss is an updated version of raw_elemss: |
16169
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16144
diff
changeset
|
1217 |
- type info added to Fixes and modified in Constrains |
14508
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1218 |
- axiom and definition statement replaced by corresponding one |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1219 |
from proppss in Assumes and Defines |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1220 |
- Facts unchanged |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1221 |
*) |
13308 | 1222 |
in ((parms, elemss, concl), text) end; |
12502 | 1223 |
|
1224 |
in |
|
1225 |
||
18671 | 1226 |
fun read_elemss x = prep_elemss ProofContext.read_vars ProofContext.read_propp_schematic x; |
1227 |
fun cert_elemss x = prep_elemss ProofContext.cert_vars ProofContext.cert_propp_schematic x; |
|
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1228 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1229 |
end; |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1230 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1231 |
|
15703 | 1232 |
(* facts and attributes *) |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1233 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1234 |
local |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1235 |
|
18678 | 1236 |
fun prep_name name = |
1237 |
if NameSpace.is_qualified name then error ("Illegal qualified name: " ^ quote name) |
|
15703 | 1238 |
else name; |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1239 |
|
15703 | 1240 |
fun prep_facts _ _ ctxt (Int elem) = |
18137 | 1241 |
Element.map_ctxt_values I I (Thm.transfer (ProofContext.theory_of ctxt)) elem |
1242 |
| prep_facts get intern ctxt (Ext elem) = elem |> Element.map_ctxt |
|
15703 | 1243 |
{var = I, typ = I, term = I, |
18678 | 1244 |
name = prep_name, |
15703 | 1245 |
fact = get ctxt, |
16458 | 1246 |
attrib = Args.assignable o intern (ProofContext.theory_of ctxt)}; |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1247 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1248 |
in |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1249 |
|
15703 | 1250 |
fun read_facts x = prep_facts ProofContext.get_thms Attrib.intern_src x; |
1251 |
fun cert_facts x = prep_facts (K I) (K I) x; |
|
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1252 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1253 |
end; |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1254 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1255 |
|
12546 | 1256 |
(* full context statements: import + elements + conclusion *) |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1257 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1258 |
local |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1259 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1260 |
fun prep_context_statement prep_expr prep_elemss prep_facts |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1261 |
do_close fixed_params import elements raw_concl context = |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1262 |
let |
16458 | 1263 |
val thy = ProofContext.theory_of context; |
13375 | 1264 |
|
16458 | 1265 |
val ((import_ids, import_syn), raw_import_elemss) = |
1266 |
flatten (context, prep_expr thy) (([], Symtab.empty), Expr import); |
|
14215
ebf291f3b449
Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents:
13629
diff
changeset
|
1267 |
(* CB: normalise "includes" among elements *) |
16458 | 1268 |
val ((ids, syn), raw_elemsss) = foldl_map (flatten (context, prep_expr thy)) |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
1269 |
((import_ids, import_syn), elements); |
15696 | 1270 |
|
1271 |
val raw_elemss = List.concat raw_elemsss; |
|
14508
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1272 |
(* CB: raw_import_elemss @ raw_elemss is the normalised list of |
859b11514537
Experimental command for instantiation of locales in proof contexts:
ballarin
parents:
14446
diff
changeset
|
1273 |
context elements obtained from import and elements. *) |
13375 | 1274 |
val ((parms, all_elemss, concl), (spec, (_, _, defs))) = prep_elemss do_close |
13336 | 1275 |
context fixed_params (raw_import_elemss @ raw_elemss) raw_concl; |
15696 | 1276 |
(* replace extended ids (for axioms) by ids *) |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1277 |
val all_elemss' = map (fn (((_, ps), _), (((n, ps'), mode), elems)) => |
17485 | 1278 |
(((n, map (fn p => (p, (the o AList.lookup (op =) ps') p)) ps), mode), elems)) |
15696 | 1279 |
(ids ~~ all_elemss); |
1280 |
||
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1281 |
(* CB: all_elemss and parms contain the correct parameter types *) |
15696 | 1282 |
val (ps,qs) = splitAt(length raw_import_elemss, all_elemss') |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1283 |
val (import_ctxt, (import_elemss, _)) = |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1284 |
activate_facts prep_facts (context, ps); |
14215
ebf291f3b449
Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents:
13629
diff
changeset
|
1285 |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1286 |
val (ctxt, (elemss, _)) = |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1287 |
activate_facts prep_facts (import_ctxt, qs); |
15212 | 1288 |
val stmt = gen_distinct Term.aconv |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1289 |
(List.concat (map (fn ((_, Assumed axs), _) => |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1290 |
List.concat (map (#hyps o Thm.rep_thm o #2) axs) |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1291 |
| ((_, Derived _), _) => []) qs)); |
16458 | 1292 |
val cstmt = map (cterm_of thy) stmt; |
12834
e5bec3268932
added locale_facts(_i) interface (useful for simple ML proof tools);
wenzelm
parents:
12806
diff
changeset
|
1293 |
in |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
1294 |
((((import_ctxt, import_elemss), (ctxt, elemss, syn)), (parms, spec, defs)), (cstmt, concl)) |
12834
e5bec3268932
added locale_facts(_i) interface (useful for simple ML proof tools);
wenzelm
parents:
12806
diff
changeset
|
1295 |
end; |
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1296 |
|
15703 | 1297 |
val gen_context = prep_context_statement intern_expr read_elemss read_facts; |
1298 |
val gen_context_i = prep_context_statement (K I) cert_elemss cert_facts; |
|
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1299 |
|
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1300 |
fun gen_statement prep_locale prep_ctxt raw_locale elems concl ctxt = |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1301 |
let |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1302 |
val thy = ProofContext.theory_of ctxt; |
16458 | 1303 |
val locale = Option.map (prep_locale thy) raw_locale; |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1304 |
val (target_stmt, fixed_params, import) = |
15531 | 1305 |
(case locale of NONE => ([], [], empty) |
1306 |
| SOME name => |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1307 |
let val {predicate = (stmt, _), params = (ps, _), ...} = |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1308 |
the_locale thy name |
16620
2a7f46324218
Proper treatment of beta-redexes in witness theorems.
ballarin
parents:
16458
diff
changeset
|
1309 |
in (stmt, map fst ps, Locale name) end); |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
1310 |
val ((((locale_ctxt, locale_elemss), (elems_ctxt, _, _)), _), (elems_stmt, concl')) = |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1311 |
prep_ctxt false fixed_params import elems concl ctxt; |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1312 |
in (locale, (target_stmt, elems_stmt), locale_ctxt, elems_ctxt, concl') end; |
13399
c136276dc847
support locale ``views'' (for cumulative predicates);
wenzelm
parents:
13394
diff
changeset
|
1313 |
|
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1314 |
in |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1315 |
|
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1316 |
fun read_context import body ctxt = #1 (gen_context true [] import (map Elem body) [] ctxt); |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1317 |
fun cert_context import body ctxt = #1 (gen_context_i true [] import (map Elem body) [] ctxt); |
14215
ebf291f3b449
Improvements to Isar/Locales: premises generated by "includes" elements
ballarin
parents:
13629
diff
changeset
|
1318 |
|
12529
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1319 |
val read_context_statement = gen_statement intern gen_context; |
d99716a53f59
simultaneous type-inference of complete context/statement specifications;
wenzelm
parents:
12514
diff
changeset
|
1320 |
val cert_context_statement = gen_statement (K I) gen_context_i; |
12502 | 1321 |
|
18617 | 1322 |
fun init loc thy = #3 (cert_context_statement (SOME loc) [] [] (ProofContext.init thy)); |
1323 |
||
12502 | 1324 |
end; |
11896 | 1325 |
|
1326 |
||
13336 | 1327 |
(* print locale *) |
12070 | 1328 |
|
17228 | 1329 |
fun print_locale thy show_facts import body = |
12070 | 1330 |
let |
18137 | 1331 |
val (((_, import_elemss), (ctxt, elemss, _)), _) = |
1332 |
read_context import body (ProofContext.init thy); |
|
1333 |
val prt_elem = Element.pretty_ctxt ctxt; |
|
15570 | 1334 |
val all_elems = List.concat (map #2 (import_elemss @ elemss)); |
12277 | 1335 |
in |
18137 | 1336 |
Pretty.big_list "locale elements:" (all_elems |
17316 | 1337 |
|> (if show_facts then I else filter (fn Notes _ => false | _ => true)) |
1338 |
|> map (Pretty.chunks o prt_elem)) |
|
13336 | 1339 |
|> Pretty.writeln |
12277 | 1340 |
end; |
12070 | 1341 |
|
1342 |
||
12706 | 1343 |
|
16144 | 1344 |
(** store results **) |
12702 | 1345 |
|
16144 | 1346 |
(* note_thmss_qualified *) |
15696 | 1347 |
|
14564 | 1348 |
fun note_thmss_qualified kind name args thy = |
12706 | 1349 |
thy |
17449
429ca1e21289
interpretation: use goal commands without target -- no storing of results;
wenzelm
parents:
17438
diff
changeset
|
1350 |
|> Theory.add_path (Sign.base_name name) |
429ca1e21289
interpretation: use goal commands without target -- no storing of results;
wenzelm
parents:
17438
diff
changeset
|
1351 |
|> Theory.no_base_names |
14564 | 1352 |
|> PureThy.note_thmss_i (Drule.kind kind) args |
18377 | 1353 |
||> Theory.restore_naming thy; |
16144 | 1354 |
|
12706 | 1355 |
|
15696 | 1356 |
(* accesses of interpreted theorems *) |
1357 |
||
16144 | 1358 |
local |
1359 |
||
1360 |
(*fully qualified name in theory is T.p.r.n where |
|
1361 |
T: theory name, p: interpretation prefix, r: renaming prefix, n: name*) |
|
1362 |
fun global_accesses _ [] = [] |
|
1363 |
| global_accesses "" [T, n] = [[T, n], [n]] |
|
1364 |
| global_accesses "" [T, r, n] = [[T, r, n], [T, n], [r, n], [n]] |
|
1365 |
| global_accesses _ [T, p, n] = [[T, p, n], [p, n]] |
|
1366 |
| global_accesses _ [T, p, r, n] = [[T, p, r, n], [T, p, n], [p, r, n], [p, n]] |
|
1367 |
| global_accesses _ names = error ("Bad name declaration " ^ quote (NameSpace.pack names)); |
|
15696 | 1368 |
|
16144 | 1369 |
(*fully qualified name in context is p.r.n where |
1370 |
p: interpretation prefix, r: renaming prefix, n: name*) |
|
1371 |
fun local_accesses _ [] = [] |
|
1372 |
| local_accesses "" [n] = [[n]] |
|
1373 |
| local_accesses "" [r, n] = [[r, n], [n]] |
|
1374 |
| local_accesses _ [p, n] = [[p, n]] |
|
1375 |
| local_accesses _ [p, r, n] = [[p, r, n], [p, n]] |
|
1376 |
| local_accesses _ names = error ("Bad name declaration " ^ quote (NameSpace.pack names)); |
|
1377 |
||
1378 |
in |
|
15696 | 1379 |
|
16144 | 1380 |
fun global_note_accesses_i kind prfx args thy = |
1381 |
thy |
|
1382 |
|> Theory.qualified_names |
|
1383 |
|> Theory.custom_accesses (global_accesses prfx) |
|
1384 |
|> PureThy.note_thmss_i kind args |
|
18377 | 1385 |
||> Theory.restore_naming thy; |
15696 | 1386 |
|
16144 | 1387 |
fun local_note_accesses_i prfx args ctxt = |
1388 |
ctxt |
|
1389 |
|> ProofContext.qualified_names |
|
1390 |
|> ProofContext.custom_accesses (local_accesses prfx) |
|
17856 | 1391 |
|> ProofContext.note_thmss_i args |> swap |
16144 | 1392 |
|>> ProofContext.restore_naming ctxt; |
1393 |
||
1394 |
end; |
|
15696 | 1395 |
|
1396 |
||
17138 | 1397 |
(* collect witness theorems for global registration; |
1398 |
requires parameters and flattened list of (assumed!) identifiers |
|
1399 |
instead of recomputing it from the target *) |
|
1400 |
||
1401 |
fun collect_global_witnesses thy parms ids vts = let |
|
1402 |
val ts = map Logic.unvarify vts; |
|
1403 |
val (parms, parmTs) = split_list parms; |
|
1404 |
val parmvTs = map Type.varifyT parmTs; |
|
1405 |
val vtinst = fold (Sign.typ_match thy) (parmvTs ~~ map Term.fastype_of ts) Vartab.empty; |
|
1406 |
val tinst = Vartab.dest vtinst |> map (fn ((x, 0), (_, T)) => (x, T)) |
|
18137 | 1407 |
|> Symtab.make; |
17138 | 1408 |
(* replace parameter names in ids by instantiations *) |
1409 |
val vinst = Symtab.make (parms ~~ vts); |
|
17412 | 1410 |
fun vinst_names ps = map (the o Symtab.lookup vinst) ps; |
17138 | 1411 |
val inst = Symtab.make (parms ~~ ts); |
1412 |
val ids' = map (apsnd vinst_names) ids; |
|
18343 | 1413 |
val wits = List.concat (map (snd o the o get_global_registration thy) ids'); |
18137 | 1414 |
in ((tinst, inst), wits) end; |
17138 | 1415 |
|
1416 |
||
15696 | 1417 |
(* store instantiations of args for all registered interpretations |
1418 |
of the theory *) |
|
1419 |
||
1420 |
fun note_thmss_registrations kind target args thy = |
|
15596 | 1421 |
let |
17138 | 1422 |
val parms = the_locale thy target |> #params |> fst |> map fst; |
16458 | 1423 |
val ids = flatten (ProofContext.init thy, intern_expr thy) |
17033 | 1424 |
(([], Symtab.empty), Expr (Locale target)) |> fst |> fst |
1425 |
|> List.mapPartial (fn (id, (_, Assumed _)) => SOME id | _ => NONE) |
|
15696 | 1426 |
|
1427 |
val regs = get_global_registrations thy target; |
|
1428 |
||
1429 |
(* add args to thy for all registrations *) |
|
15596 | 1430 |
|
18190 | 1431 |
fun activate (vts, ((prfx, atts2), _)) thy = |
15696 | 1432 |
let |
18137 | 1433 |
val (insts, prems) = collect_global_witnesses thy parms ids vts; |
1434 |
val inst_atts = |
|
1435 |
Args.map_values I (Element.instT_type (#1 insts)) |
|
1436 |
(Element.inst_term insts) (Element.inst_thm thy insts); |
|
1437 |
val args' = args |> map (fn ((n, atts), [(ths, [])]) => |
|
17033 | 1438 |
((NameSpace.qualified prfx n, |
18728 | 1439 |
map (Attrib.attribute_i thy) (map inst_atts atts @ atts2)), |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1440 |
[(map (Drule.standard o satisfy_protected prems o |
18137 | 1441 |
Element.inst_thm thy insts) ths, [])])); |
18377 | 1442 |
in global_note_accesses_i (Drule.kind kind) prfx args' thy |> snd end; |
18190 | 1443 |
in fold activate regs thy end; |
15596 | 1444 |
|
1445 |
||
18421
464c93701351
re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents:
18377
diff
changeset
|
1446 |
fun smart_note_thmss kind NONE = PureThy.note_thmss_i (Drule.kind kind) |
464c93701351
re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents:
18377
diff
changeset
|
1447 |
| smart_note_thmss kind (SOME loc) = note_thmss_qualified kind loc; |
12958 | 1448 |
|
1449 |
||
1450 |
local |
|
1451 |
||
15696 | 1452 |
(* add facts to locale in theory *) |
1453 |
||
12958 | 1454 |
fun put_facts loc args thy = |
18698 | 1455 |
let val {predicate, import, elems, params, regs} = the_locale thy loc |
15703 | 1456 |
in |
1457 |
thy |> put_locale loc {predicate = predicate, import = import, |
|
18698 | 1458 |
elems = elems @ [(Notes args, stamp ())], params = params, regs = regs} |
15703 | 1459 |
end; |
12958 | 1460 |
|
15696 | 1461 |
(* add theorem to locale and theory, |
1462 |
base for theorems (in loc) and declare (in loc) *) |
|
1463 |
||
15703 | 1464 |
fun gen_note_thmss prep_locale prep_facts kind raw_loc args thy = |
12706 | 1465 |
let |
1466 |
val thy_ctxt = ProofContext.init thy; |
|
16458 | 1467 |
val loc = prep_locale thy raw_loc; |
15703 | 1468 |
val (_, (stmt, _), loc_ctxt, _, _) = cert_context_statement (SOME loc) [] [] thy_ctxt; |
18038 | 1469 |
val export = ProofContext.export_view stmt loc_ctxt thy_ctxt; |
15703 | 1470 |
|
17109
606c269d1e26
added add_locale_context(_i), which returns the body context for presentation;
wenzelm
parents:
17096
diff
changeset
|
1471 |
val (ctxt', (args', facts)) = activate_note prep_facts (loc_ctxt, args); |
15703 | 1472 |
val facts' = |
1473 |
map (rpair [] o #1 o #1) args' ~~ |
|
1474 |
map (single o Thm.no_attributes o map export o #2) facts; |
|
17109
606c269d1e26
added add_locale_context(_i), which returns the body context for presentation;
wenzelm
parents:
17096
diff
changeset
|
1475 |
|
18377 | 1476 |
val (result, thy') = |
17109
606c269d1e26
added add_locale_context(_i), which returns the body context for presentation;
wenzelm
parents:
17096
diff
changeset
|
1477 |
thy |
606c269d1e26
added add_locale_context(_i), which returns the body context for presentation;
wenzelm
parents:
17096
diff
changeset
|
1478 |
|> put_facts loc args' |
606c269d1e26
added add_locale_context(_i), which returns the body context for presentation;
wenzelm
parents:
17096
diff
changeset
|
1479 |
|> note_thmss_registrations kind loc args' |
606c269d1e26
added add_locale_context(_i), which returns the body context for presentation;
wenzelm
parents:
17096
diff
changeset
|
1480 |
|> note_thmss_qualified kind loc facts'; |
18421
464c93701351
re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents:
18377
diff
changeset
|
1481 |
in (result, (thy', ctxt')) end; |
12706 | 1482 |
|
1483 |
in |
|
1484 |
||
15703 | 1485 |
val note_thmss = gen_note_thmss intern read_facts; |
1486 |
val note_thmss_i = gen_note_thmss (K I) cert_facts; |
|
12711 | 1487 |
|
17355 | 1488 |
fun add_thmss kind loc args (ctxt, thy) = |
12958 | 1489 |
let |
15703 | 1490 |
val (ctxt', (args', facts)) = activate_note cert_facts |
17355 | 1491 |
(ctxt, map (apsnd Thm.simple_fact) args); |
15703 | 1492 |
val thy' = |
1493 |
thy |
|
1494 |
|> put_facts loc args' |
|
1495 |
|> note_thmss_registrations kind loc args'; |
|
17355 | 1496 |
in (facts, (ctxt', thy')) end; |
12702 | 1497 |
|
12706 | 1498 |
end; |
12063 | 1499 |
|
11896 | 1500 |
|
18137 | 1501 |
|
1502 |
(** define locales **) |
|
1503 |
||
13336 | 1504 |
(* predicate text *) |
15596 | 1505 |
(* CB: generate locale predicates and delta predicates *) |
13336 | 1506 |
|
13375 | 1507 |
local |
1508 |
||
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1509 |
(* introN: name of theorems for introduction rules of locale and |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1510 |
delta predicates; |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1511 |
axiomsN: name of theorem set with destruct rules for locale predicates, |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1512 |
also name suffix of delta predicates. *) |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1513 |
|
13375 | 1514 |
val introN = "intro"; |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1515 |
val axiomsN = "axioms"; |
13375 | 1516 |
|
16458 | 1517 |
fun atomize_spec thy ts = |
13375 | 1518 |
let |
18502 | 1519 |
val t = Logic.mk_conjunction_list ts; |
16458 | 1520 |
val body = ObjectLogic.atomize_term thy t; |
13375 | 1521 |
val bodyT = Term.fastype_of body; |
1522 |
in |
|
16458 | 1523 |
if bodyT = propT then (t, propT, Thm.reflexive (Thm.cterm_of thy t)) |
18782 | 1524 |
else (body, bodyT, ObjectLogic.atomize_cterm (Thm.cterm_of thy t)) |
13375 | 1525 |
end; |
1526 |
||
13394 | 1527 |
fun aprop_tr' n c = (c, fn args => |
1528 |
if length args = n then Syntax.const "_aprop" $ Term.list_comb (Syntax.free c, args) |
|
1529 |
else raise Match); |
|
13336 | 1530 |
|
15104 | 1531 |
(* CB: define one predicate including its intro rule and axioms |
1532 |
- bname: predicate name |
|
1533 |
- parms: locale parameters |
|
1534 |
- defs: thms representing substitutions from defines elements |
|
1535 |
- ts: terms representing locale assumptions (not normalised wrt. defs) |
|
1536 |
- norm_ts: terms representing locale assumptions (normalised wrt. defs) |
|
1537 |
- thy: the theory |
|
1538 |
*) |
|
1539 |
||
13420 | 1540 |
fun def_pred bname parms defs ts norm_ts thy = |
13375 | 1541 |
let |
16458 | 1542 |
val name = Sign.full_name thy bname; |
13375 | 1543 |
|
16458 | 1544 |
val (body, bodyT, body_eq) = atomize_spec thy norm_ts; |
13394 | 1545 |
val env = Term.add_term_free_names (body, []); |
15570 | 1546 |
val xs = List.filter (fn (x, _) => x mem_string env) parms; |
13394 | 1547 |
val Ts = map #2 xs; |
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
1548 |
val extraTs = (Term.term_tfrees body \\ foldr Term.add_typ_tfrees [] Ts) |
13394 | 1549 |
|> Library.sort_wrt #1 |> map TFree; |
13399
c136276dc847
support locale ``views'' (for cumulative predicates);
wenzelm
parents:
13394
diff
changeset
|
1550 |
val predT = map Term.itselfT extraTs ---> Ts ---> bodyT; |
13336 | 1551 |
|
13394 | 1552 |
val args = map Logic.mk_type extraTs @ map Free xs; |
1553 |
val head = Term.list_comb (Const (name, predT), args); |
|
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1554 |
val statement = ObjectLogic.ensure_propT thy head; |
13375 | 1555 |
|
18358 | 1556 |
val ([pred_def], defs_thy) = |
13375 | 1557 |
thy |
13394 | 1558 |
|> (if bodyT <> propT then I else |
1559 |
Theory.add_trfuns ([], [], map (aprop_tr' (length args)) (NameSpace.accesses' name), [])) |
|
18671 | 1560 |
|> Theory.add_consts_i [(bname, predT, NoSyn)] |
13375 | 1561 |
|> PureThy.add_defs_i false [((Thm.def_name bname, Logic.mk_equals (head, body)), [])]; |
13394 | 1562 |
|
16458 | 1563 |
val cert = Thm.cterm_of defs_thy; |
13375 | 1564 |
|
17973 | 1565 |
val intro = Drule.standard (Goal.prove defs_thy [] norm_ts statement (fn _ => |
13375 | 1566 |
Tactic.rewrite_goals_tac [pred_def] THEN |
1567 |
Tactic.compose_tac (false, body_eq RS Drule.equal_elim_rule1, 1) 1 THEN |
|
16947 | 1568 |
Tactic.compose_tac (false, Drule.conj_intr_list (map (Thm.assume o cert) norm_ts), 0) 1)); |
13375 | 1569 |
|
1570 |
val conjuncts = |
|
17257
0ab67cb765da
introduced binding priority 1 for linear combinators etc.
haftmann
parents:
17228
diff
changeset
|
1571 |
(Drule.equal_elim_rule1 OF [Thm.symmetric body_eq, |
0ab67cb765da
introduced binding priority 1 for linear combinators etc.
haftmann
parents:
17228
diff
changeset
|
1572 |
Tactic.rewrite_rule [pred_def] (Thm.assume (cert statement))]) |
18502 | 1573 |
|> Drule.conj_elim_precise [length ts] |> hd; |
17257
0ab67cb765da
introduced binding priority 1 for linear combinators etc.
haftmann
parents:
17228
diff
changeset
|
1574 |
val axioms = ts ~~ conjuncts |> map (fn (t, ax) => |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1575 |
prove_protected defs_thy t |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1576 |
(Tactic.rewrite_goals_tac defs THEN |
13375 | 1577 |
Tactic.compose_tac (false, ax, 0) 1)); |
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1578 |
in ((statement, intro, axioms), defs_thy) end; |
13375 | 1579 |
|
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1580 |
fun assumes_to_notes (Assumes asms) axms = |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1581 |
axms |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1582 |
|> fold_map (fn (a, spec) => fn axs => |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1583 |
let val (ps, qs) = splitAt (length spec, axs) |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1584 |
in ((a, [(ps, [])]), qs) end |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1585 |
) asms |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1586 |
|-> (fn asms' => pair (Notes asms')) |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1587 |
| assumes_to_notes e axms = (e, axms); |
13394 | 1588 |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1589 |
(* CB: changes only "new" elems, these have identifier ("", _). *) |
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1590 |
|
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1591 |
fun change_elemss axioms (import_elemss, body_elemss) = |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1592 |
let |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1593 |
fun change (id as ("", _), es)= |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1594 |
fold_map assumes_to_notes |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1595 |
(map (Element.map_ctxt_values I I (satisfy_protected axioms)) es) |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1596 |
#-> (fn es' => pair (id, es')) |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1597 |
| change e = pair e; |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1598 |
in |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1599 |
axioms |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1600 |
|> map (conclude_protected o #2) |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1601 |
|> fold_map change import_elemss |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1602 |
||>> fold_map change body_elemss |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1603 |
|> fst |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1604 |
end; |
13394 | 1605 |
|
1606 |
in |
|
13375 | 1607 |
|
15104 | 1608 |
(* CB: main predicate definition function *) |
1609 |
||
13394 | 1610 |
fun define_preds bname (parms, ((exts, exts'), (ints, ints')), defs) elemss thy = |
1611 |
let |
|
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1612 |
val ((elemss', more_ts), thy') = |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1613 |
if null exts then ((elemss, []), thy) |
13394 | 1614 |
else |
1615 |
let |
|
15531 | 1616 |
val aname = if null ints then bname else bname ^ "_" ^ axiomsN; |
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1617 |
val ((statement, intro, axioms), def_thy) = |
13394 | 1618 |
thy |> def_pred aname parms defs exts exts'; |
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1619 |
val elemss' = |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1620 |
elemss |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1621 |
|> change_elemss axioms |
18671 | 1622 |
|> apsnd (fn elems => elems @ |
1623 |
[(("", []), [Assumes [((bname ^ "_" ^ axiomsN, []), [(statement, ([], []))])]])]); |
|
13394 | 1624 |
in |
18377 | 1625 |
def_thy |
1626 |
|> note_thmss_qualified "" aname [((introN, []), [([intro], [])])] |
|
1627 |
|> snd |
|
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1628 |
|> pair (elemss', [statement]) |
13394 | 1629 |
end; |
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1630 |
val (predicate, thy'') = |
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1631 |
if null ints then (([], []), thy') |
13394 | 1632 |
else |
1633 |
let |
|
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1634 |
val ((statement, intro, axioms), def_thy) = |
13394 | 1635 |
thy' |> def_pred bname parms defs (ints @ more_ts) (ints' @ more_ts); |
16458 | 1636 |
val cstatement = Thm.cterm_of def_thy statement; |
13394 | 1637 |
in |
18377 | 1638 |
def_thy |
1639 |
|> note_thmss_qualified "" bname |
|
1640 |
[((introN, []), [([intro], [])]), |
|
1641 |
((axiomsN, []), [(map (Drule.standard o conclude_protected o #2) axioms, [])])] |
|
1642 |
|> snd |
|
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1643 |
|> pair ([cstatement], axioms) |
13394 | 1644 |
end; |
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1645 |
in ((elemss', predicate), thy'') end; |
13375 | 1646 |
|
1647 |
end; |
|
13336 | 1648 |
|
1649 |
||
13297 | 1650 |
(* add_locale(_i) *) |
1651 |
||
1652 |
local |
|
1653 |
||
18343 | 1654 |
fun gen_add_locale prep_ctxt prep_expr |
1655 |
do_predicate bname raw_import raw_body thy = |
|
13297 | 1656 |
let |
16458 | 1657 |
val name = Sign.full_name thy bname; |
18343 | 1658 |
val _ = conditional (is_some (get_locale thy name)) (fn () => |
13297 | 1659 |
error ("Duplicate definition of locale " ^ quote name)); |
1660 |
||
1661 |
val thy_ctxt = ProofContext.init thy; |
|
17228 | 1662 |
val (((import_ctxt, import_elemss), (body_ctxt, body_elemss, syn)), |
1663 |
text as (parms, ((_, exts'), _), _)) = |
|
13375 | 1664 |
prep_ctxt raw_import raw_body thy_ctxt; |
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1665 |
val elemss = (import_elemss, body_elemss); |
16458 | 1666 |
val import = prep_expr thy raw_import; |
13297 | 1667 |
|
17228 | 1668 |
val extraTs = foldr Term.add_term_tfrees [] exts' \\ |
17756 | 1669 |
foldr Term.add_typ_tfrees [] (map snd parms); |
17228 | 1670 |
val _ = if null extraTs then () |
17437 | 1671 |
else warning ("Additional type variable(s) in locale specification " ^ quote bname); |
17228 | 1672 |
|
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1673 |
val ((elemss', predicate as (predicate_statement, predicate_axioms)), pred_thy) = |
18343 | 1674 |
if do_predicate then thy |> define_preds bname text elemss |
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1675 |
else ((elemss, ([], [])), thy); |
13420 | 1676 |
|
18137 | 1677 |
fun axiomify axioms elemss = |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1678 |
(axioms, elemss) |> foldl_map (fn (axs, (id, elems)) => let |
15570 | 1679 |
val ts = List.concat (List.mapPartial (fn (Assumes asms) => |
1680 |
SOME (List.concat (map (map #1 o #2) asms)) | _ => NONE) elems); |
|
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1681 |
val (axs1, axs2) = splitAt (length ts, axs); |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1682 |
in (axs2, ((id, Assumed axs1), elems)) end) |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1683 |
|> snd; |
18569 | 1684 |
val pred_ctxt = ProofContext.init pred_thy; |
15206
09d78ec709c7
Modified locales: improved implementation of "includes".
ballarin
parents:
15127
diff
changeset
|
1685 |
val (ctxt, (_, facts)) = activate_facts (K I) |
18569 | 1686 |
(pred_ctxt, axiomify predicate_axioms ((op @) elemss')); |
18546 | 1687 |
val export = ProofContext.export_view predicate_statement ctxt thy_ctxt; |
13420 | 1688 |
val facts' = facts |> map (fn (a, ths) => ((a, []), [(map export ths, [])])); |
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1689 |
val elems' = List.concat (map #2 (List.filter (equal "" o #1 o #1) ((op @) elemss'))) |
13297 | 1690 |
in |
13375 | 1691 |
pred_thy |
18377 | 1692 |
|> note_thmss_qualified "" name facts' |> snd |
13297 | 1693 |
|> declare_locale name |
16458 | 1694 |
|> put_locale name {predicate = predicate, import = import, |
17142
76a5a2cc3171
add_locale_context(_i) now exporting elements (still some refinements to be done)
haftmann
parents:
17138
diff
changeset
|
1695 |
elems = map (fn e => (e, stamp ())) elems', |
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1696 |
params = (params_of ((op @) elemss') |> |
17412 | 1697 |
map (fn (x, SOME T) => ((x, T), the (Symtab.lookup syn x))), map #1 (params_of body_elemss)), |
16736
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
1698 |
regs = []} |
18550
59b89f625d68
add_local_context now yields imported and body elements seperatly; additional slight clenup in code
haftmann
parents:
18546
diff
changeset
|
1699 |
|> pair ((fn (e1, e2) => (map snd e1, map snd e2)) elemss', body_ctxt) |
13297 | 1700 |
end; |
1701 |
||
1702 |
in |
|
1703 |
||
17109
606c269d1e26
added add_locale_context(_i), which returns the body context for presentation;
wenzelm
parents:
17096
diff
changeset
|
1704 |
val add_locale_context = gen_add_locale read_context intern_expr; |
606c269d1e26
added add_locale_context(_i), which returns the body context for presentation;
wenzelm
parents:
17096
diff
changeset
|
1705 |
val add_locale_context_i = gen_add_locale cert_context (K I); |
17142
76a5a2cc3171
add_locale_context(_i) now exporting elements (still some refinements to be done)
haftmann
parents:
17138
diff
changeset
|
1706 |
fun add_locale b = #2 oooo add_locale_context b; |
76a5a2cc3171
add_locale_context(_i) now exporting elements (still some refinements to be done)
haftmann
parents:
17138
diff
changeset
|
1707 |
fun add_locale_i b = #2 oooo add_locale_context_i b; |
13297 | 1708 |
|
1709 |
end; |
|
1710 |
||
15801 | 1711 |
val _ = Context.add_setup |
18708 | 1712 |
(add_locale_i true "var" empty [Fixes [(Syntax.internal "x", NONE, NoSyn)]] #> |
1713 |
add_locale_i true "struct" empty [Fixes [(Syntax.internal "S", NONE, Structure)]]); |
|
15801 | 1714 |
|
13297 | 1715 |
|
12730 | 1716 |
|
17355 | 1717 |
(** locale goals **) |
1718 |
||
1719 |
local |
|
1720 |
||
18137 | 1721 |
fun intern_attrib thy = map_elem (Element.map_ctxt |
1722 |
{name = I, var = I, typ = I, term = I, fact = I, attrib = Attrib.intern_src thy}); |
|
1723 |
||
17355 | 1724 |
fun global_goal prep_att = |
1725 |
Proof.global_goal ProofDisplay.present_results prep_att ProofContext.bind_propp_schematic_i; |
|
1726 |
||
17856 | 1727 |
fun gen_theorem prep_att prep_elem prep_stmt kind before_qed after_qed a raw_elems concl thy = |
17355 | 1728 |
let |
1729 |
val thy_ctxt = ProofContext.init thy; |
|
1730 |
val elems = map (prep_elem thy) raw_elems; |
|
1731 |
val (_, (_, view), _, ctxt, propp) = prep_stmt NONE elems (map snd concl) thy_ctxt; |
|
1732 |
val ctxt' = ctxt |> ProofContext.add_view thy_ctxt view; |
|
1733 |
val stmt = map fst concl ~~ propp; |
|
17856 | 1734 |
in global_goal prep_att kind before_qed after_qed (SOME "") a stmt ctxt' end; |
17449
429ca1e21289
interpretation: use goal commands without target -- no storing of results;
wenzelm
parents:
17438
diff
changeset
|
1735 |
|
429ca1e21289
interpretation: use goal commands without target -- no storing of results;
wenzelm
parents:
17438
diff
changeset
|
1736 |
fun gen_theorem_in_locale prep_locale prep_src prep_elem prep_stmt no_target |
17856 | 1737 |
kind before_qed after_qed raw_locale (name, atts) raw_elems concl thy = |
17355 | 1738 |
let |
1739 |
val locale = prep_locale thy raw_locale; |
|
1740 |
val locale_atts = map (prep_src thy) atts; |
|
1741 |
val locale_attss = map (map (prep_src thy) o snd o fst) concl; |
|
17449
429ca1e21289
interpretation: use goal commands without target -- no storing of results;
wenzelm
parents:
17438
diff
changeset
|
1742 |
val target = if no_target then NONE else SOME (extern thy locale); |
17355 | 1743 |
val elems = map (prep_elem thy) raw_elems; |
1744 |
val names = map (fst o fst) concl; |
|
1745 |
||
17449
429ca1e21289
interpretation: use goal commands without target -- no storing of results;
wenzelm
parents:
17438
diff
changeset
|
1746 |
val thy_ctxt = ProofContext.init thy; |
17355 | 1747 |
val (_, (locale_view, elems_view), locale_ctxt, elems_ctxt, propp) = |
1748 |
prep_stmt (SOME raw_locale) elems (map snd concl) thy_ctxt; |
|
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1749 |
val elems_ctxt' = elems_ctxt |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1750 |
|> ProofContext.add_view locale_ctxt elems_view |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1751 |
|> ProofContext.add_view thy_ctxt locale_view; |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1752 |
val locale_ctxt' = locale_ctxt |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1753 |
|> ProofContext.add_view thy_ctxt locale_view; |
18137 | 1754 |
|
17355 | 1755 |
val stmt = map (apsnd (K []) o fst) concl ~~ propp; |
1756 |
||
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1757 |
fun after_qed' results = |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1758 |
let val locale_results = results |> (map o map) |
18782 | 1759 |
(ProofContext.export_standard elems_ctxt' locale_ctxt') in |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1760 |
curry (add_thmss kind locale ((names ~~ locale_attss) ~~ locale_results)) locale_ctxt |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1761 |
#-> (fn res => |
17355 | 1762 |
if name = "" andalso null locale_atts then I |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1763 |
else #2 o add_thmss kind locale [((name, locale_atts), List.concat (map #2 res))]) |
17355 | 1764 |
#> #2 |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1765 |
#> after_qed locale_results results |
17355 | 1766 |
end; |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1767 |
in global_goal (K I) kind before_qed after_qed' target (name, []) stmt elems_ctxt' end; |
17355 | 1768 |
|
1769 |
in |
|
1770 |
||
18728 | 1771 |
val theorem = gen_theorem Attrib.attribute intern_attrib read_context_statement; |
17355 | 1772 |
val theorem_i = gen_theorem (K I) (K I) cert_context_statement; |
18137 | 1773 |
|
1774 |
val theorem_in_locale = gen_theorem_in_locale intern Attrib.intern_src intern_attrib |
|
1775 |
read_context_statement false; |
|
1776 |
||
1777 |
val theorem_in_locale_i = gen_theorem_in_locale (K I) (K I) (K I) |
|
1778 |
cert_context_statement false; |
|
1779 |
||
1780 |
val theorem_in_locale_no_target = gen_theorem_in_locale (K I) (K I) (K I) |
|
1781 |
cert_context_statement true; |
|
17355 | 1782 |
|
17384 | 1783 |
fun smart_theorem kind NONE a [] concl = |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1784 |
Proof.theorem kind NONE (K I) (SOME "") a concl o ProofContext.init |
17355 | 1785 |
| smart_theorem kind NONE a elems concl = |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1786 |
theorem kind NONE (K I) a elems concl |
17355 | 1787 |
| smart_theorem kind (SOME loc) a elems concl = |
17856 | 1788 |
theorem_in_locale kind NONE (K (K I)) loc a elems concl; |
17355 | 1789 |
|
1790 |
end; |
|
1791 |
||
1792 |
||
15598
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
ballarin
parents:
15596
diff
changeset
|
1793 |
(** Interpretation commands **) |
15596 | 1794 |
|
1795 |
local |
|
1796 |
||
1797 |
(* extract proof obligations (assms and defs) from elements *) |
|
1798 |
||
17138 | 1799 |
fun extract_asms_elem (Fixes _) ts = ts |
1800 |
| extract_asms_elem (Constrains _) ts = ts |
|
1801 |
| extract_asms_elem (Assumes asms) ts = |
|
15596 | 1802 |
ts @ List.concat (map (fn (_, ams) => map (fn (t, _) => t) ams) asms) |
17138 | 1803 |
| extract_asms_elem (Defines defs) ts = |
16169
b59202511b8a
Locales: new element constrains, parameter renaming with syntax,
ballarin
parents:
16144
diff
changeset
|
1804 |
ts @ map (fn (_, (def, _)) => def) defs |
17138 | 1805 |
| extract_asms_elem (Notes _) ts = ts; |
15596 | 1806 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1807 |
fun extract_asms_elems ((id, Assumed _), elems) = |
17138 | 1808 |
(id, fold extract_asms_elem elems []) |
1809 |
| extract_asms_elems ((id, Derived _), _) = (id, []); |
|
15596 | 1810 |
|
17138 | 1811 |
fun extract_asms_elemss elemss = map extract_asms_elems elemss; |
15596 | 1812 |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
1813 |
(* activate instantiated facts in theory or context *) |
15596 | 1814 |
|
17138 | 1815 |
fun gen_activate_facts_elemss get_reg note attrib std put_reg add_wit |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1816 |
attn all_elemss new_elemss propss thmss thy_ctxt = |
17033 | 1817 |
let |
1818 |
fun activate_elem disch (prfx, atts) (Notes facts) thy_ctxt = |
|
1819 |
let |
|
1820 |
val facts' = facts |
|
1821 |
(* discharge hyps in attributes *) |
|
17109
606c269d1e26
added add_locale_context(_i), which returns the body context for presentation;
wenzelm
parents:
17096
diff
changeset
|
1822 |
|> Attrib.map_facts (attrib thy_ctxt o Args.map_values I I I disch) |
17033 | 1823 |
(* insert interpretation attributes *) |
1824 |
|> map (apfst (apsnd (fn a => a @ map (attrib thy_ctxt) atts))) |
|
1825 |
(* discharge hyps *) |
|
1826 |
|> map (apsnd (map (apfst (map disch)))) |
|
1827 |
(* prefix names *) |
|
1828 |
|> map (apfst (apfst (NameSpace.qualified prfx))) |
|
17138 | 1829 |
in fst (note prfx facts' thy_ctxt) end |
17033 | 1830 |
| activate_elem _ _ _ thy_ctxt = thy_ctxt; |
1831 |
||
17138 | 1832 |
fun activate_elems disch ((id, _), elems) thy_ctxt = |
17033 | 1833 |
let |
18343 | 1834 |
val ((prfx, atts2), _) = the (get_reg thy_ctxt id) |
17033 | 1835 |
handle Option => sys_error ("Unknown registration of " ^ |
1836 |
quote (fst id) ^ " while activating facts."); |
|
1837 |
in |
|
1838 |
fold (activate_elem disch (prfx, atts2)) elems thy_ctxt |
|
1839 |
end; |
|
15596 | 1840 |
|
17033 | 1841 |
val thy_ctxt' = thy_ctxt |
1842 |
(* add registrations *) |
|
1843 |
|> fold (fn ((id, _), _) => put_reg id attn) new_elemss |
|
1844 |
(* add witnesses of Assumed elements *) |
|
1845 |
|> fold (fn (id, thms) => fold (add_wit id) thms) |
|
1846 |
(map fst propss ~~ thmss); |
|
15596 | 1847 |
|
17033 | 1848 |
val prems = List.concat (List.mapPartial |
1849 |
(fn ((id, Assumed _), _) => Option.map snd (get_reg thy_ctxt' id) |
|
1850 |
| ((_, Derived _), _) => NONE) all_elemss); |
|
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1851 |
val disch = satisfy_protected prems; |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1852 |
val disch' = std o disch; (* FIXME *) |
17033 | 1853 |
|
1854 |
val thy_ctxt'' = thy_ctxt' |
|
1855 |
(* add witnesses of Derived elements *) |
|
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1856 |
|> fold (fn (id, thms) => fold (add_wit id o apsnd disch) thms) |
17033 | 1857 |
(List.mapPartial (fn ((_, Assumed _), _) => NONE |
1858 |
| ((id, Derived thms), _) => SOME (id, thms)) all_elemss) |
|
1859 |
in |
|
1860 |
thy_ctxt'' |
|
1861 |
(* add facts to theory or context *) |
|
1862 |
|> fold (activate_elems disch') new_elemss |
|
1863 |
end; |
|
15596 | 1864 |
|
17355 | 1865 |
fun global_activate_facts_elemss x = gen_activate_facts_elemss |
15696 | 1866 |
(fn thy => fn (name, ps) => |
1867 |
get_global_registration thy (name, map Logic.varify ps)) |
|
18377 | 1868 |
(swap ooo global_note_accesses_i (Drule.kind "")) |
18728 | 1869 |
Attrib.attribute_i Drule.standard |
17033 | 1870 |
(fn (name, ps) => put_global_registration (name, map Logic.varify ps)) |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1871 |
(fn (n, ps) => fn (t, th) => |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1872 |
add_global_witness (n, map Logic.varify ps) |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
1873 |
(Logic.unvarify t, Drule.freeze_all th)) x; |
17355 | 1874 |
|
1875 |
fun local_activate_facts_elemss x = gen_activate_facts_elemss |
|
15696 | 1876 |
get_local_registration |
16144 | 1877 |
local_note_accesses_i |
18728 | 1878 |
(Attrib.attribute_i o ProofContext.theory_of) I |
17033 | 1879 |
put_local_registration |
17355 | 1880 |
add_local_witness x; |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
1881 |
|
17033 | 1882 |
fun gen_prep_registration mk_ctxt is_local read_terms test_reg activate |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
1883 |
attn expr insts thy_ctxt = |
15596 | 1884 |
let |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
1885 |
val ctxt = mk_ctxt thy_ctxt; |
16458 | 1886 |
val thy = ProofContext.theory_of ctxt; |
15596 | 1887 |
|
15696 | 1888 |
val ctxt' = ctxt |> ProofContext.theory_of |> ProofContext.init; |
16458 | 1889 |
val ((ids, _), raw_elemss) = flatten (ctxt', intern_expr thy) |
16102
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
ballarin
parents:
16028
diff
changeset
|
1890 |
(([], Symtab.empty), Expr expr); |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1891 |
val ((parms, all_elemss, _), (_, (_, defs, _))) = |
16736
1e792b32abef
Preparations for interpretation of locales in locales.
ballarin
parents:
16620
diff
changeset
|
1892 |
read_elemss false ctxt' [] raw_elemss []; |
15596 | 1893 |
|
1894 |
(** compute instantiation **) |
|
1895 |
||
15696 | 1896 |
(* user input *) |
15596 | 1897 |
val insts = if length parms < length insts |
1898 |
then error "More arguments than parameters in instantiation." |
|
1899 |
else insts @ replicate (length parms - length insts) NONE; |
|
1900 |
val (ps, pTs) = split_list parms; |
|
1901 |
val pvTs = map Type.varifyT pTs; |
|
15598
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
ballarin
parents:
15596
diff
changeset
|
1902 |
|
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
ballarin
parents:
15596
diff
changeset
|
1903 |
(* instantiations given by user *) |
15596 | 1904 |
val given = List.mapPartial (fn (_, (NONE, _)) => NONE |
1905 |
| (x, (SOME inst, T)) => SOME (x, (inst, T))) (ps ~~ (insts ~~ pvTs)); |
|
1906 |
val (given_ps, given_insts) = split_list given; |
|
1907 |
val tvars = foldr Term.add_typ_tvars [] pvTs; |
|
1908 |
val used = foldr Term.add_typ_varnames [] pvTs; |
|
17485 | 1909 |
fun sorts (a, i) = AList.lookup (op =) tvars (a, i); |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
1910 |
val (vs, vinst) = read_terms thy_ctxt sorts used given_insts; |
15696 | 1911 |
val vars = foldl Term.add_term_tvar_ixns [] vs \\ map fst tvars; |
16861 | 1912 |
val vars' = fold Term.add_term_varnames vs vars; |
15696 | 1913 |
val _ = if null vars' then () |
1914 |
else error ("Illegal schematic variable(s) in instantiation: " ^ |
|
1915 |
commas_quote (map Syntax.string_of_vname vars')); |
|
15598
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
ballarin
parents:
15596
diff
changeset
|
1916 |
(* replace new types (which are TFrees) by ones with new names *) |
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
ballarin
parents:
15596
diff
changeset
|
1917 |
val new_Tnames = foldr Term.add_term_tfree_names [] vs; |
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
ballarin
parents:
15596
diff
changeset
|
1918 |
val new_Tnames' = Term.invent_names used "'a" (length new_Tnames); |
15696 | 1919 |
val renameT = |
1920 |
if is_local then I |
|
1921 |
else Type.unvarifyT o Term.map_type_tfree (fn (a, s) => |
|
17485 | 1922 |
TFree ((the o AList.lookup (op =) (new_Tnames ~~ new_Tnames')) a, s)); |
15696 | 1923 |
val rename = |
1924 |
if is_local then I |
|
1925 |
else Term.map_term_types renameT; |
|
15598
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
ballarin
parents:
15596
diff
changeset
|
1926 |
|
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
ballarin
parents:
15596
diff
changeset
|
1927 |
val tinst = Symtab.make (map |
15696 | 1928 |
(fn ((x, 0), T) => (x, T |> renameT) |
16850 | 1929 |
| ((_, n), _) => sys_error "Internal error var in prep_registration") vinst); |
15696 | 1930 |
val inst = Symtab.make (given_ps ~~ map rename vs); |
15596 | 1931 |
|
1932 |
(* defined params without user input *) |
|
1933 |
val not_given = List.mapPartial (fn (x, (NONE, T)) => SOME (x, T) |
|
1934 |
| (_, (SOME _, _)) => NONE) (ps ~~ (insts ~~ pTs)); |
|
18137 | 1935 |
fun add_def (p, pT) inst = |
15596 | 1936 |
let |
1937 |
val (t, T) = case find_first (fn (Free (a, _), _) => a = p) defs of |
|
1938 |
NONE => error ("Instance missing for parameter " ^ quote p) |
|
1939 |
| SOME (Free (_, T), t) => (t, T); |
|
18137 | 1940 |
val d = Element.inst_term (tinst, inst) t; |
1941 |
in Symtab.update_new (p, d) inst end; |
|
1942 |
val inst = fold add_def not_given inst; |
|
1943 |
val insts = (tinst, inst); |
|
1944 |
(* Note: insts contain no vars. *) |
|
15596 | 1945 |
|
1946 |
(** compute proof obligations **) |
|
1947 |
||
15598
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
ballarin
parents:
15596
diff
changeset
|
1948 |
(* restore "small" ids *) |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1949 |
val ids' = map (fn ((n, ps), (_, mode)) => |
17485 | 1950 |
((n, map (fn p => Free (p, (the o AList.lookup (op =) parms) p)) ps), mode)) ids; |
15596 | 1951 |
(* instantiate ids and elements *) |
18137 | 1952 |
val inst_elemss = (ids' ~~ all_elemss) |> map (fn (((n, ps), _), ((_, mode), elems)) => |
1953 |
((n, map (Element.inst_term insts) ps), |
|
1954 |
map (fn Int e => Element.inst_ctxt thy insts e) elems) |
|
1955 |
|> apfst (fn id => (id, map_mode (map (fn (t, th) => |
|
1956 |
(Element.inst_term insts t, Element.inst_thm thy insts th))) mode))); |
|
15596 | 1957 |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
1958 |
(* remove fragments already registered with theory or context *) |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1959 |
val new_inst_elemss = List.filter (fn ((id, _), _) => |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
1960 |
not (test_reg thy_ctxt id)) inst_elemss; |
15703 | 1961 |
val new_ids = map #1 new_inst_elemss; |
15596 | 1962 |
|
1963 |
val propss = extract_asms_elemss new_inst_elemss; |
|
1964 |
||
15703 | 1965 |
val bind_attrib = Attrib.crude_closure ctxt o Args.assignable; |
16458 | 1966 |
val attn' = apsnd (map (bind_attrib o Attrib.intern_src thy)) attn; |
15596 | 1967 |
|
17033 | 1968 |
in (propss, activate attn' inst_elemss new_inst_elemss propss) end; |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
1969 |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
1970 |
val prep_global_registration = gen_prep_registration |
15696 | 1971 |
ProofContext.init false |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
1972 |
(fn thy => fn sorts => fn used => |
16458 | 1973 |
Sign.read_def_terms (thy, K NONE, sorts) used true) |
15696 | 1974 |
(fn thy => fn (name, ps) => |
1975 |
test_global_registration thy (name, map Logic.varify ps)) |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
1976 |
global_activate_facts_elemss; |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
1977 |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
1978 |
val prep_local_registration = gen_prep_registration |
15696 | 1979 |
I true |
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
1980 |
(fn ctxt => ProofContext.read_termTs ctxt (K false) (K NONE)) |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
1981 |
smart_test_registration |
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15623
diff
changeset
|
1982 |
local_activate_facts_elemss; |
15596 | 1983 |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1984 |
fun prep_registration_in_locale target expr thy = |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1985 |
(* target already in internal form *) |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1986 |
let |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1987 |
val ctxt = ProofContext.init thy; |
17138 | 1988 |
val ((raw_target_ids, target_syn), _) = flatten (ctxt, I) |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1989 |
(([], Symtab.empty), Expr (Locale target)); |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1990 |
val fixed = the_locale thy target |> #params |> #1 |> map #1; |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1991 |
val ((all_ids, syn), raw_elemss) = flatten (ctxt, intern_expr thy) |
17138 | 1992 |
((raw_target_ids, target_syn), Expr expr); |
1993 |
val (target_ids, ids) = splitAt (length raw_target_ids, all_ids); |
|
1994 |
val ((parms, elemss, _), _) = read_elemss false ctxt fixed raw_elemss []; |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1995 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1996 |
(** compute proof obligations **) |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1997 |
|
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1998 |
(* restore "small" ids, with mode *) |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
1999 |
val ids' = map (apsnd snd) ids; |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
2000 |
(* remove Int markers *) |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
2001 |
val elemss' = map (fn (_, es) => |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
2002 |
map (fn Int e => e) es) elemss |
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
2003 |
(* extract assumptions and defs *) |
17138 | 2004 |
val ids_elemss = ids' ~~ elemss'; |
2005 |
val propss = extract_asms_elemss ids_elemss; |
|
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
2006 |
|
17138 | 2007 |
(** activation function: |
2008 |
- add registrations to the target locale |
|
2009 |
- add induced registrations for all global registrations of |
|
2010 |
the target, unless already present |
|
2011 |
- add facts of induced registrations to theory **) |
|
2012 |
||
2013 |
val t_ids = List.mapPartial |
|
2014 |
(fn (id, (_, Assumed _)) => SOME id | _ => NONE) target_ids; |
|
2015 |
||
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2016 |
fun activate thmss thy = let |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2017 |
val satisfy = satisfy_protected (List.concat thmss); |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2018 |
val ids_elemss_thmss = ids_elemss ~~ thmss; |
17138 | 2019 |
val regs = get_global_registrations thy target; |
2020 |
||
2021 |
fun activate_id (((id, Assumed _), _), thms) thy = |
|
17033 | 2022 |
thy |> put_registration_in_locale target id |
17138 | 2023 |
|> fold (add_witness_in_locale target id) thms |
2024 |
| activate_id _ thy = thy; |
|
2025 |
||
2026 |
fun activate_reg (vts, ((prfx, atts2), _)) thy = let |
|
18137 | 2027 |
val (insts, wits) = collect_global_witnesses thy fixed t_ids vts; |
2028 |
fun inst_parms ps = map |
|
17485 | 2029 |
(the o AList.lookup (op =) (map #1 fixed ~~ vts)) ps; |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2030 |
val disch = satisfy_protected wits; |
17138 | 2031 |
val new_elemss = List.filter (fn (((name, ps), _), _) => |
2032 |
not (test_global_registration thy (name, inst_parms ps))) (ids_elemss); |
|
2033 |
fun activate_assumed_id (((_, Derived _), _), _) thy = thy |
|
2034 |
| activate_assumed_id ((((name, ps), Assumed _), _), thms) thy = let |
|
2035 |
val ps' = inst_parms ps; |
|
2036 |
in |
|
2037 |
if test_global_registration thy (name, ps') |
|
2038 |
then thy |
|
2039 |
else thy |
|
2040 |
|> put_global_registration (name, ps') (prfx, atts2) |
|
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2041 |
|> fold (fn (t, th) => fn thy => add_global_witness (name, ps') |
18137 | 2042 |
(Element.inst_term insts t, disch (Element.inst_thm thy insts th)) thy) thms |
17138 | 2043 |
end; |
2044 |
||
2045 |
fun activate_derived_id ((_, Assumed _), _) thy = thy |
|
2046 |
| activate_derived_id (((name, ps), Derived ths), _) thy = let |
|
2047 |
val ps' = inst_parms ps; |
|
2048 |
in |
|
2049 |
if test_global_registration thy (name, ps') |
|
2050 |
then thy |
|
2051 |
else thy |
|
2052 |
|> put_global_registration (name, ps') (prfx, atts2) |
|
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2053 |
|> fold (fn (t, th) => fn thy => add_global_witness (name, ps') |
18137 | 2054 |
(Element.inst_term insts t, |
2055 |
disch (Element.inst_thm thy insts (satisfy th))) thy) ths |
|
17138 | 2056 |
end; |
2057 |
||
2058 |
fun activate_elem (Notes facts) thy = |
|
2059 |
let |
|
2060 |
val facts' = facts |
|
18728 | 2061 |
|> Attrib.map_facts (Attrib.attribute_i thy o |
18137 | 2062 |
Args.map_values I (Element.instT_type (#1 insts)) |
2063 |
(Element.inst_term insts) |
|
2064 |
(disch o Element.inst_thm thy insts o satisfy)) |
|
18728 | 2065 |
|> map (apfst (apsnd (fn a => a @ map (Attrib.attribute thy) atts2))) |
18137 | 2066 |
|> map (apsnd (map (apfst (map (disch o Element.inst_thm thy insts o satisfy))))) |
17138 | 2067 |
|> map (apfst (apfst (NameSpace.qualified prfx))) |
2068 |
in |
|
18377 | 2069 |
thy |
2070 |
|> global_note_accesses_i (Drule.kind "") prfx facts' |
|
2071 |
|> snd |
|
17138 | 2072 |
end |
2073 |
| activate_elem _ thy = thy; |
|
2074 |
||
2075 |
fun activate_elems (_, elems) thy = fold activate_elem elems thy; |
|
2076 |
||
2077 |
in thy |> fold activate_assumed_id ids_elemss_thmss |
|
2078 |
|> fold activate_derived_id ids_elemss |
|
2079 |
|> fold activate_elems new_elemss end; |
|
17033 | 2080 |
in |
17138 | 2081 |
thy |> fold activate_id ids_elemss_thmss |
2082 |
|> fold activate_reg regs |
|
17033 | 2083 |
end; |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
2084 |
|
17033 | 2085 |
in (propss, activate) end; |
17000
552df70f52c2
First version of interpretation in locales. Not yet fully functional.
ballarin
parents:
16947
diff
changeset
|
2086 |
|
17355 | 2087 |
fun prep_propp propss = propss |> map (fn ((name, _), props) => |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2088 |
(("", []), map (rpair ([], []) o Logic.protect) props)); |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2089 |
|
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2090 |
fun prep_result propps thmss = |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2091 |
ListPair.map (fn ((_, props), thms) => (props ~~ thms)) (propps, thmss); |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2092 |
|
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2093 |
val refine_protected = |
18213 | 2094 |
Proof.refine (Method.Basic (K (Method.RAW_METHOD |
18502 | 2095 |
(K (ALLGOALS |
2096 |
(PRECISE_CONJUNCTS ~1 (ALLGOALS |
|
2097 |
(PRECISE_CONJUNCTS ~1 (ALLGOALS (Tactic.rtac Drule.protectI)))))))))) |
|
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2098 |
#> Seq.hd; |
17437 | 2099 |
|
17438 | 2100 |
fun goal_name thy kind target propss = |
2101 |
kind ^ Proof.goal_names (Option.map (extern thy) target) "" |
|
2102 |
(propss |> map (fn ((name, _), _) => extern thy name)); |
|
17355 | 2103 |
|
2104 |
in |
|
2105 |
||
2106 |
fun interpretation (prfx, atts) expr insts thy = |
|
2107 |
let |
|
2108 |
val (propss, activate) = prep_global_registration (prfx, atts) expr insts thy; |
|
17449
429ca1e21289
interpretation: use goal commands without target -- no storing of results;
wenzelm
parents:
17438
diff
changeset
|
2109 |
val kind = goal_name thy "interpretation" NONE propss; |
18502 | 2110 |
val after_qed = activate o prep_result propss; |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2111 |
in |
18137 | 2112 |
ProofContext.init thy |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2113 |
|> Proof.theorem_i kind NONE after_qed NONE ("", []) (prep_propp propss) |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2114 |
|> refine_protected |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2115 |
end; |
17355 | 2116 |
|
2117 |
fun interpretation_in_locale (raw_target, expr) thy = |
|
2118 |
let |
|
2119 |
val target = intern thy raw_target; |
|
2120 |
val (propss, activate) = prep_registration_in_locale target expr thy; |
|
17449
429ca1e21289
interpretation: use goal commands without target -- no storing of results;
wenzelm
parents:
17438
diff
changeset
|
2121 |
val kind = goal_name thy "interpretation" (SOME target) propss; |
18137 | 2122 |
val after_qed = K (activate o prep_result propss); |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2123 |
in |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2124 |
thy |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2125 |
|> theorem_in_locale_no_target kind NONE after_qed target ("", []) [] (prep_propp propss) |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2126 |
|> refine_protected |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2127 |
end; |
17449
429ca1e21289
interpretation: use goal commands without target -- no storing of results;
wenzelm
parents:
17438
diff
changeset
|
2128 |
|
429ca1e21289
interpretation: use goal commands without target -- no storing of results;
wenzelm
parents:
17438
diff
changeset
|
2129 |
fun interpret (prfx, atts) expr insts int state = |
429ca1e21289
interpretation: use goal commands without target -- no storing of results;
wenzelm
parents:
17438
diff
changeset
|
2130 |
let |
429ca1e21289
interpretation: use goal commands without target -- no storing of results;
wenzelm
parents:
17438
diff
changeset
|
2131 |
val ctxt = Proof.context_of state; |
429ca1e21289
interpretation: use goal commands without target -- no storing of results;
wenzelm
parents:
17438
diff
changeset
|
2132 |
val (propss, activate) = prep_local_registration (prfx, atts) expr insts ctxt; |
429ca1e21289
interpretation: use goal commands without target -- no storing of results;
wenzelm
parents:
17438
diff
changeset
|
2133 |
val kind = goal_name (Proof.theory_of state) "interpret" NONE propss; |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2134 |
fun after_qed results = |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2135 |
Proof.map_context (K (ctxt |> activate (prep_result propss results))) |
17449
429ca1e21289
interpretation: use goal commands without target -- no storing of results;
wenzelm
parents:
17438
diff
changeset
|
2136 |
#> Proof.put_facts NONE |
429ca1e21289
interpretation: use goal commands without target -- no storing of results;
wenzelm
parents:
17438
diff
changeset
|
2137 |
#> Seq.single; |
17355 | 2138 |
in |
18123
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2139 |
state |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2140 |
|> Proof.local_goal (ProofDisplay.print_results int) (K I) ProofContext.bind_propp_i |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2141 |
kind NONE after_qed (prep_propp propss) |
1bb572e8cee9
avoid prove_plain, export_plain, simplified after_qed;
wenzelm
parents:
18038
diff
changeset
|
2142 |
|> refine_protected |
17437 | 2143 |
end; |
15596 | 2144 |
|
11896 | 2145 |
end; |
17355 | 2146 |
|
2147 |
end; |