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