author | paulson |
Tue, 12 Sep 2000 10:27:16 +0200 | |
changeset 9930 | c02d48a47ed1 |
parent 8908 | 25f2bdc02123 |
child 10349 | 78434c9a54fd |
permissions | -rw-r--r-- |
5244 | 1 |
(* Title: Pure/locale.ML |
2 |
ID: $Id$ |
|
3 |
Author: Florian Kammueller, University of Cambridge |
|
4 |
||
5782
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
5 |
Locales. The theory section 'locale' declarings constants, assumptions and |
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
6 |
definitions that have local scope. Typical form is |
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
7 |
|
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
8 |
locale Locale_name = |
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
9 |
fixes (*variables that are fixed in the locale's scope*) |
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
10 |
v :: T |
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
11 |
assumes (*meta-hypothesis that hold in the locale*) |
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
12 |
Asm_name "meta-formula" |
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
13 |
defines (*local definitions of fixed variables in terms of others*) |
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
14 |
v_def "v x == ...x..." |
5244 | 15 |
|
16 |
TODO: |
|
6022 | 17 |
- operations on locales: renaming. |
5244 | 18 |
*) |
19 |
||
20 |
signature BASIC_LOCALE = |
|
21 |
sig |
|
22 |
val print_locales: theory -> unit |
|
8459 | 23 |
val pretty_goals_marker: string -> int -> thm -> Pretty.T list |
24 |
val pretty_goals: int -> thm -> Pretty.T list |
|
25 |
val print_goals_marker: string -> int -> thm -> unit |
|
5244 | 26 |
val print_goals: int -> thm -> unit |
6190
f0c14e527d68
added thm, thms, Open_locale, Close_locale, Print_scope;
wenzelm
parents:
6022
diff
changeset
|
27 |
val thm: xstring -> thm |
f0c14e527d68
added thm, thms, Open_locale, Close_locale, Print_scope;
wenzelm
parents:
6022
diff
changeset
|
28 |
val thms: xstring -> thm list |
f0c14e527d68
added thm, thms, Open_locale, Close_locale, Print_scope;
wenzelm
parents:
6022
diff
changeset
|
29 |
val Open_locale: xstring -> unit |
f0c14e527d68
added thm, thms, Open_locale, Close_locale, Print_scope;
wenzelm
parents:
6022
diff
changeset
|
30 |
val Close_locale: xstring -> unit |
f0c14e527d68
added thm, thms, Open_locale, Close_locale, Print_scope;
wenzelm
parents:
6022
diff
changeset
|
31 |
val Print_scope: unit -> unit |
5244 | 32 |
end; |
33 |
||
34 |
signature LOCALE = |
|
35 |
sig |
|
36 |
include BASIC_LOCALE |
|
37 |
val get_thm: theory -> xstring -> thm |
|
38 |
val get_thms: theory -> xstring -> thm list |
|
39 |
type locale |
|
6022 | 40 |
val add_locale: bstring -> (bstring option) -> (string * string * mixfix) list -> |
5244 | 41 |
(string * string) list -> (string * string) list -> theory -> theory |
6022 | 42 |
val add_locale_i: bstring -> (bstring option) -> (string * typ * mixfix) list -> |
5244 | 43 |
(string * term) list -> (string * term) list -> theory -> theory |
44 |
val open_locale: xstring -> theory -> theory |
|
6022 | 45 |
val close_locale: xstring -> theory -> theory |
46 |
val print_scope: theory -> unit |
|
5244 | 47 |
val in_locale: term list -> Sign.sg -> bool |
48 |
val is_open_loc_sg: Sign.sg -> bool |
|
49 |
val is_open_loc: theory -> bool |
|
50 |
val read_cterm: Sign.sg -> string * typ -> cterm |
|
51 |
val get_scope: theory -> (string * locale) list |
|
52 |
val get_scope_sg: Sign.sg -> (string * locale) list |
|
53 |
val collect_consts: Sign.sg -> string list |
|
54 |
val setup: (theory -> theory) list |
|
55 |
end; |
|
56 |
||
57 |
structure Locale: LOCALE = |
|
58 |
struct |
|
59 |
||
60 |
||
61 |
(** type locale **) |
|
62 |
||
63 |
type locale = |
|
6022 | 64 |
{ancestor: string option, |
65 |
consts: (string * typ) list, |
|
5244 | 66 |
nosyn: string list, |
67 |
rules: (string * term) list, |
|
68 |
defs: (string * term) list, |
|
69 |
thms: (string * thm) list, |
|
70 |
defaults: (string * sort) list * (string * typ) list * string list}; |
|
71 |
||
6022 | 72 |
fun make_locale ancestor consts nosyn rules defs thms defaults = |
73 |
{ancestor = ancestor, consts = consts, nosyn = nosyn, rules = rules, |
|
74 |
defs = defs, thms = thms, defaults = defaults}: locale; |
|
5244 | 75 |
|
6022 | 76 |
fun pretty_locale sg (name, {ancestor, consts, rules, defs, nosyn = _, thms = _, defaults = _}) = |
5244 | 77 |
let |
78 |
val prt_typ = Pretty.quote o Sign.pretty_typ sg; |
|
79 |
val prt_term = Pretty.quote o Sign.pretty_term sg; |
|
80 |
||
81 |
fun pretty_const (c, T) = Pretty.block |
|
82 |
[Pretty.str (c ^ " ::"), Pretty.brk 1, prt_typ T]; |
|
83 |
||
84 |
fun pretty_axiom (a, t) = Pretty.block |
|
85 |
[Pretty.str (a ^ ":"), Pretty.brk 1, prt_term t]; |
|
6022 | 86 |
|
87 |
val anc = case ancestor of |
|
88 |
None => "" |
|
89 |
| Some(loc) => ((Sign.base_name loc) ^ " +") |
|
5244 | 90 |
in |
6022 | 91 |
Pretty.big_list (name ^ " = " ^ anc) |
5244 | 92 |
[Pretty.big_list "consts:" (map pretty_const consts), |
93 |
Pretty.big_list "rules:" (map pretty_axiom rules), |
|
94 |
Pretty.big_list "defs:" (map pretty_axiom defs)] |
|
95 |
end; |
|
96 |
||
97 |
||
98 |
||
99 |
(** theory data **) |
|
100 |
||
101 |
(* data kind 'Pure/locales' *) |
|
102 |
||
103 |
type locale_data = |
|
104 |
{space: NameSpace.T, |
|
105 |
locales: locale Symtab.table, |
|
106 |
scope: (string * locale) list ref}; |
|
107 |
||
108 |
fun make_locale_data space locales scope = |
|
109 |
{space = space, locales = locales, scope = scope}: locale_data; |
|
110 |
||
111 |
structure LocalesArgs = |
|
112 |
struct |
|
113 |
val name = "Pure/locales"; |
|
114 |
type T = locale_data; |
|
115 |
||
116 |
val empty = make_locale_data NameSpace.empty Symtab.empty (ref []); |
|
6546 | 117 |
fun copy {space, locales, scope = ref locs} = make_locale_data space locales (ref locs); |
5244 | 118 |
fun prep_ext {space, locales, scope = _} = make_locale_data space locales (ref []); |
119 |
fun merge ({space = space1, locales = locales1, scope = _}, |
|
120 |
{space = space2, locales = locales2, scope = _}) = |
|
121 |
make_locale_data (NameSpace.merge (space1, space2)) |
|
122 |
(Symtab.merge (K true) (locales1, locales2)) |
|
123 |
(ref []); |
|
124 |
||
125 |
fun print sg {space, locales, scope} = |
|
126 |
let |
|
127 |
fun extrn name = |
|
128 |
if ! long_names then name else NameSpace.extern space name; |
|
129 |
val locs = map (apfst extrn) (Symtab.dest locales); |
|
130 |
val scope_names = rev (map (extrn o fst) (! scope)); |
|
131 |
in |
|
8720 | 132 |
[Display.pretty_name_space ("locale name space", space), |
133 |
Pretty.big_list "locales:" (map (pretty_locale sg) locs), |
|
134 |
Pretty.strs ("current scope:" :: scope_names)] |
|
135 |
|> Pretty.chunks |> Pretty.writeln |
|
5244 | 136 |
end; |
137 |
end; |
|
138 |
||
139 |
||
140 |
structure LocalesData = TheoryDataFun(LocalesArgs); |
|
141 |
val print_locales = LocalesData.print; |
|
142 |
||
143 |
||
144 |
(* access locales *) |
|
145 |
||
146 |
fun get_locale_sg sg name = Symtab.lookup (#locales (LocalesData.get_sg sg), name); |
|
147 |
||
148 |
val get_locale = get_locale_sg o Theory.sign_of; |
|
149 |
||
150 |
fun put_locale (name, locale) thy = |
|
151 |
let |
|
152 |
val {space, locales, scope} = LocalesData.get thy; |
|
153 |
val space' = NameSpace.extend (space, [name]); |
|
154 |
val locales' = Symtab.update ((name, locale), locales); |
|
155 |
in thy |> LocalesData.put (make_locale_data space' locales' scope) end; |
|
156 |
||
157 |
fun lookup_locale thy xname = |
|
158 |
let |
|
159 |
val {space, locales, ...} = LocalesData.get thy; |
|
160 |
val name = NameSpace.intern space xname; |
|
161 |
in apsome (pair name) (get_locale thy name) end; |
|
162 |
||
163 |
||
164 |
(* access scope *) |
|
165 |
||
7469 | 166 |
fun get_scope_sg sg = |
167 |
if Sign.eq_sg (sg, Theory.sign_of ProtoPure.thy) then [] |
|
168 |
else ! (#scope (LocalesData.get_sg sg)); |
|
5244 | 169 |
|
170 |
val get_scope = get_scope_sg o Theory.sign_of; |
|
171 |
||
172 |
fun change_scope f thy = |
|
173 |
let val {scope, ...} = LocalesData.get thy |
|
174 |
in scope := f (! scope) end; |
|
175 |
||
176 |
||
177 |
||
178 |
(** scope operations **) |
|
179 |
||
180 |
(* change scope *) |
|
181 |
||
182 |
fun the_locale thy xname = |
|
183 |
(case lookup_locale thy xname of |
|
184 |
Some loc => loc |
|
185 |
| None => error ("Unknown locale " ^ quote xname)); |
|
186 |
||
187 |
fun open_locale xname thy = |
|
6022 | 188 |
let val loc = the_locale thy xname; |
189 |
val anc = #ancestor(#2(loc)); |
|
190 |
val cur_sc = get_scope thy; |
|
191 |
fun opn lc th = (change_scope (cons lc) th; th) |
|
192 |
in case anc of |
|
193 |
None => opn loc thy |
|
194 |
| Some(loc') => |
|
195 |
if loc' mem (map fst cur_sc) |
|
196 |
then opn loc thy |
|
7628
8e177a4c86a5
tidied, using "warning" function and fixing the Close_locale bug
paulson
parents:
7469
diff
changeset
|
197 |
else (warning ("Opening locale " ^ quote loc' ^ ", required by " ^ |
8e177a4c86a5
tidied, using "warning" function and fixing the Close_locale bug
paulson
parents:
7469
diff
changeset
|
198 |
quote xname); |
6022 | 199 |
opn loc (open_locale (Sign.base_name loc') thy)) |
200 |
end; |
|
5244 | 201 |
|
202 |
fun pop_locale [] = error "Currently no open locales" |
|
203 |
| pop_locale (_ :: locs) = locs; |
|
204 |
||
6022 | 205 |
fun close_locale name thy = |
7628
8e177a4c86a5
tidied, using "warning" function and fixing the Close_locale bug
paulson
parents:
7469
diff
changeset
|
206 |
let val lname = (case get_scope thy of (ln,_)::_ => ln |
8e177a4c86a5
tidied, using "warning" function and fixing the Close_locale bug
paulson
parents:
7469
diff
changeset
|
207 |
| _ => error "No locales are open!") |
8e177a4c86a5
tidied, using "warning" function and fixing the Close_locale bug
paulson
parents:
7469
diff
changeset
|
208 |
val ok = (name = Sign.base_name lname) handle _ => false |
8e177a4c86a5
tidied, using "warning" function and fixing the Close_locale bug
paulson
parents:
7469
diff
changeset
|
209 |
in if ok then (change_scope pop_locale thy; thy) |
8e177a4c86a5
tidied, using "warning" function and fixing the Close_locale bug
paulson
parents:
7469
diff
changeset
|
210 |
else error ("locale " ^ name ^ " is not top of scope; top is " ^ lname) |
6022 | 211 |
end; |
5244 | 212 |
|
6022 | 213 |
fun print_scope thy = |
214 |
Pretty.writeln (Pretty.strs ("current scope:" :: rev(map (Sign.base_name o fst) (get_scope thy)))); |
|
5244 | 215 |
|
6190
f0c14e527d68
added thm, thms, Open_locale, Close_locale, Print_scope;
wenzelm
parents:
6022
diff
changeset
|
216 |
(*implicit context versions*) |
f0c14e527d68
added thm, thms, Open_locale, Close_locale, Print_scope;
wenzelm
parents:
6022
diff
changeset
|
217 |
fun Open_locale xname = (open_locale xname (Context.the_context ()); ()); |
f0c14e527d68
added thm, thms, Open_locale, Close_locale, Print_scope;
wenzelm
parents:
6022
diff
changeset
|
218 |
fun Close_locale xname = (close_locale xname (Context.the_context ()); ()); |
f0c14e527d68
added thm, thms, Open_locale, Close_locale, Print_scope;
wenzelm
parents:
6022
diff
changeset
|
219 |
fun Print_scope () = (print_scope (Context.the_context ()); ()); |
f0c14e527d68
added thm, thms, Open_locale, Close_locale, Print_scope;
wenzelm
parents:
6022
diff
changeset
|
220 |
|
f0c14e527d68
added thm, thms, Open_locale, Close_locale, Print_scope;
wenzelm
parents:
6022
diff
changeset
|
221 |
|
5244 | 222 |
(** functions for goals.ML **) |
223 |
||
224 |
(* in_locale: check if hyps (: term list) of a proof are contained in the |
|
225 |
(current) scope. This function is needed in prepare_proof. It needs to |
|
226 |
refer to the signature, because theory is not available in prepare_proof. *) |
|
227 |
||
228 |
fun in_locale hyps sg = |
|
229 |
let val cur_sc = get_scope_sg sg; |
|
230 |
val rule_lists = map (#rules o snd) cur_sc; |
|
231 |
val def_lists = map (#defs o snd) cur_sc; |
|
232 |
val rules = map snd (foldr (op union) (rule_lists, [])); |
|
233 |
val defs = map snd (foldr (op union) (def_lists, [])); |
|
234 |
val defnrules = rules @ defs; |
|
235 |
in |
|
236 |
hyps subset defnrules |
|
237 |
end; |
|
238 |
||
239 |
||
240 |
(* is_open_loc: check if any locale is open, i.e. in the scope of the current thy *) |
|
241 |
fun is_open_loc_sg sign = |
|
242 |
let val cur_sc = get_scope_sg sign |
|
243 |
in not(null(cur_sc)) end; |
|
244 |
||
245 |
val is_open_loc = is_open_loc_sg o Theory.sign_of; |
|
246 |
||
247 |
||
248 |
(* get theorems *) |
|
249 |
||
250 |
fun get_thm_locale name ((_, {thms, ...}: locale)) = assoc (thms, name); |
|
251 |
||
252 |
fun get_thmx f get thy name = |
|
253 |
(case get_first (get_thm_locale name) (get_scope thy) of |
|
254 |
Some thm => f thm |
|
255 |
| None => get thy name); |
|
256 |
||
257 |
val get_thm = get_thmx I PureThy.get_thm; |
|
6190
f0c14e527d68
added thm, thms, Open_locale, Close_locale, Print_scope;
wenzelm
parents:
6022
diff
changeset
|
258 |
val get_thms = get_thmx (fn x => [x]) PureThy.get_thms; |
5244 | 259 |
|
6190
f0c14e527d68
added thm, thms, Open_locale, Close_locale, Print_scope;
wenzelm
parents:
6022
diff
changeset
|
260 |
fun thm name = get_thm (Context.the_context ()) name; |
f0c14e527d68
added thm, thms, Open_locale, Close_locale, Print_scope;
wenzelm
parents:
6022
diff
changeset
|
261 |
fun thms name = get_thms (Context.the_context ()) name; |
f0c14e527d68
added thm, thms, Open_locale, Close_locale, Print_scope;
wenzelm
parents:
6022
diff
changeset
|
262 |
|
5244 | 263 |
|
6022 | 264 |
(* get the defaults of a locale, for extension *) |
265 |
||
266 |
fun get_defaults thy name = |
|
267 |
let val (lname, loc) = the_locale thy name; |
|
268 |
in #defaults(loc) |
|
269 |
end; |
|
5244 | 270 |
|
271 |
||
272 |
(** define locales **) |
|
273 |
||
274 |
(* prepare types *) |
|
275 |
||
276 |
fun read_typ sg (envT, s) = |
|
277 |
let |
|
278 |
fun def_sort (x, ~1) = assoc (envT, x) |
|
279 |
| def_sort _ = None; |
|
280 |
val T = Type.no_tvars (Sign.read_typ (sg, def_sort) s) handle TYPE (msg, _, _) => error msg; |
|
281 |
in (Term.add_typ_tfrees (T, envT), T) end; |
|
282 |
||
283 |
fun cert_typ sg (envT, raw_T) = |
|
284 |
let val T = Type.no_tvars (Sign.certify_typ sg raw_T) handle TYPE (msg, _, _) => error msg |
|
285 |
in (Term.add_typ_tfrees (T, envT), T) end; |
|
286 |
||
287 |
||
288 |
(* prepare props *) |
|
289 |
||
290 |
val add_frees = foldl_aterms (fn (vs, Free v) => v ins vs | (vs, _) => vs); |
|
291 |
||
292 |
fun enter_term t (envS, envT, used) = |
|
293 |
(Term.add_term_tfrees (t, envS), add_frees (envT, t), Term.add_term_tfree_names (t, used)); |
|
294 |
||
295 |
fun read_axm sg ((envS, envT, used), (name, s)) = |
|
296 |
let |
|
297 |
fun def_sort (x, ~1) = assoc (envS, x) |
|
298 |
| def_sort _ = None; |
|
299 |
fun def_type (x, ~1) = assoc (envT, x) |
|
300 |
| def_type _ = None; |
|
301 |
val (_, t) = Theory.read_def_axm (sg, def_type, def_sort) used (name, s); |
|
302 |
in |
|
303 |
(enter_term t (envS, envT, used), t) |
|
304 |
end; |
|
305 |
||
306 |
||
307 |
fun cert_axm sg ((envS, envT, used), (name, raw_t)) = |
|
308 |
let val (_, t) = Theory.cert_axm sg (name, raw_t) |
|
309 |
in (enter_term t (envS, envT, used), t) end; |
|
310 |
||
311 |
||
312 |
(* Locale.read_cterm: read in a string as a certified term, and respect the bindings |
|
313 |
that already exist for subterms. If no locale is open, this function is equal to |
|
314 |
Thm.read_cterm *) |
|
315 |
||
316 |
fun read_cterm sign = |
|
317 |
let val cur_sc = get_scope_sg sign; |
|
318 |
val defaults = map (#defaults o snd) cur_sc; |
|
319 |
val envS = flat (map #1 defaults); |
|
320 |
val envT = flat (map #2 defaults); |
|
321 |
val used = flat (map #3 defaults); |
|
322 |
fun def_sort (x, ~1) = assoc (envS, x) |
|
323 |
| def_sort _ = None; |
|
324 |
fun def_type (x, ~1) = assoc (envT, x) |
|
325 |
| def_type _ = None; |
|
326 |
in (if (is_open_loc_sg sign) |
|
327 |
then (#1 o read_def_cterm (sign, def_type, def_sort) used true) |
|
328 |
else Thm.read_cterm sign) |
|
329 |
end; |
|
330 |
||
331 |
(* basic functions needed for definitions and display *) |
|
332 |
(* collect all locale constants of a scope, i.e. a list of locales *) |
|
333 |
fun collect_consts sg = |
|
334 |
let val cur_sc = get_scope_sg sg; |
|
335 |
val locale_list = map snd cur_sc; |
|
336 |
val const_list = flat (map #consts locale_list) |
|
337 |
in map fst const_list end; |
|
338 |
||
339 |
(* filter out the Free's in a term *) |
|
340 |
fun list_frees t = |
|
341 |
case t of Const(c,T) => [] |
|
342 |
| Var(v,T) => [] |
|
343 |
| Free(v,T)=> [Free(v,T)] |
|
344 |
| Bound x => [] |
|
345 |
| Abs(a,T,u) => list_frees u |
|
346 |
| t1 $ t2 => (list_frees t1) @ (list_frees t2); |
|
347 |
||
348 |
(* filter out all Free's in a term that are not contained |
|
349 |
in a list of strings. Used to prepare definitions. The list of strings |
|
350 |
will be the consts of the scope. We filter out the "free" Free's to be |
|
351 |
able to bind them *) |
|
352 |
fun difflist term clist = |
|
353 |
let val flist = list_frees term; |
|
354 |
fun builddiff [] sl = [] |
|
355 |
| builddiff (t :: tl) sl = |
|
356 |
let val Free(v,T) = t |
|
357 |
in |
|
358 |
if (v mem sl) |
|
359 |
then builddiff tl sl |
|
360 |
else t :: (builddiff tl sl) |
|
361 |
end; |
|
362 |
in distinct(builddiff flist clist) end; |
|
363 |
||
5782
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
364 |
(* Bind a term with !! over a list of "free" Free's. |
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
365 |
To enable definitions like x + y == .... (without quantifier). |
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
366 |
Complications, because x and y have to be removed from defaults *) |
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
367 |
fun abs_over_free clist ((defaults: (string * sort) list * (string * typ) list * string list), (s, term)) = |
5244 | 368 |
let val diffl = rev(difflist term clist); |
369 |
fun abs_o (t, (x as Free(v,T))) = all(T) $ Abs(v, T, abstract_over (x,t)) |
|
370 |
| abs_o (_ , _) = error ("Can't be: abs_over_free"); |
|
5782
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
371 |
val diffl' = map (fn (Free (s, T)) => s) diffl; |
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
372 |
val defaults' = (#1 defaults, filter (fn x => not((fst x) mem diffl')) (#2 defaults), #3 defaults) |
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
373 |
in (defaults', (s, foldl abs_o (term, diffl))) end; |
5244 | 374 |
|
375 |
(* assume a definition, i.e assume the cterm of a definiton term and then eliminate |
|
376 |
the binding !!, so that the def can be applied as rewrite. The meta hyp will still contain !! *) |
|
5782
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
377 |
fun prep_hyps clist sg = forall_elim_vars(0) o Thm.assume o (Thm.cterm_of sg); |
5244 | 378 |
|
379 |
||
380 |
(* concrete syntax *) |
|
381 |
||
382 |
fun mark_syn c = "\\<^locale>" ^ c; |
|
383 |
||
384 |
fun mk_loc_tr c ts = list_comb (Free (c, dummyT), ts); |
|
385 |
||
386 |
||
387 |
(* add_locale *) |
|
388 |
||
6022 | 389 |
fun gen_add_locale prep_typ prep_term bname bancestor raw_consts raw_rules raw_defs thy = |
5244 | 390 |
let val sign = Theory.sign_of thy; |
391 |
||
392 |
val name = Sign.full_name sign bname; |
|
393 |
||
6022 | 394 |
val (envSb, old_loc_consts, _) = |
395 |
case bancestor of |
|
396 |
Some(loc) => (get_defaults thy loc) |
|
397 |
| None => ([],[],[]); |
|
5244 | 398 |
|
6022 | 399 |
val old_nosyn = case bancestor of |
400 |
Some(loc) => #nosyn(#2(the_locale thy loc)) |
|
401 |
| None => []; |
|
402 |
||
403 |
(* Get the full name of the ancestor *) |
|
404 |
val ancestor = case bancestor of |
|
405 |
Some(loc) => Some(#1(the_locale thy loc)) |
|
406 |
| None => None; |
|
407 |
||
408 |
(* prepare locale consts *) |
|
5244 | 409 |
|
410 |
fun prep_const (envS, (raw_c, raw_T, raw_mx)) = |
|
411 |
let |
|
412 |
val c = Syntax.const_name raw_c raw_mx; |
|
413 |
val c_syn = mark_syn c; |
|
414 |
val mx = Syntax.fix_mixfix raw_c raw_mx; |
|
415 |
val (envS', T) = prep_typ sign (envS, raw_T) handle ERROR => |
|
416 |
error ("The error(s) above occured in locale constant " ^ quote c); |
|
417 |
val trfun = if mx = Syntax.NoSyn then None else Some (c_syn, mk_loc_tr c); |
|
418 |
in (envS', ((c, T), (c_syn, T, mx), trfun)) end; |
|
419 |
||
6022 | 420 |
val (envS0, loc_consts_syn) = foldl_map prep_const (envSb, raw_consts); |
5244 | 421 |
val loc_consts = map #1 loc_consts_syn; |
6022 | 422 |
val loc_consts = old_loc_consts @ loc_consts; |
5244 | 423 |
val loc_syn = map #2 loc_consts_syn; |
6022 | 424 |
val nosyn = old_nosyn @ (map (#1 o #1) (filter (fn x => (#3(#2 x)) = NoSyn) loc_consts_syn)); |
5244 | 425 |
val loc_trfuns = mapfilter #3 loc_consts_syn; |
426 |
||
427 |
||
428 |
(* 1st stage: syntax_thy *) |
|
429 |
||
430 |
val syntax_thy = |
|
431 |
thy |
|
432 |
|> Theory.add_modesyntax_i ("", true) loc_syn |
|
433 |
|> Theory.add_trfuns ([], loc_trfuns, [], []); |
|
434 |
||
435 |
val syntax_sign = Theory.sign_of syntax_thy; |
|
436 |
||
437 |
||
438 |
(* prepare rules and defs *) |
|
439 |
||
440 |
fun prep_axiom (env, (a, raw_t)) = |
|
441 |
let |
|
442 |
val (env', t) = prep_term syntax_sign (env, (a, raw_t)) handle ERROR => |
|
443 |
error ("The error(s) above occured in locale rule / definition " ^ quote a); |
|
444 |
in (env', (a, t)) end; |
|
445 |
||
446 |
val ((envS1, envT1, used1), loc_rules) = |
|
447 |
foldl_map prep_axiom ((envS0, loc_consts, map fst envS0), raw_rules); |
|
6268 | 448 |
val (defaults, loc_defs) = |
449 |
foldl_map prep_axiom ((envS1, envT1, used1), raw_defs); |
|
5244 | 450 |
|
451 |
val old_loc_consts = collect_consts syntax_sign; |
|
452 |
val new_loc_consts = (map #1 loc_consts); |
|
453 |
val all_loc_consts = old_loc_consts @ new_loc_consts; |
|
454 |
||
6268 | 455 |
val (defaults, loc_defs_terms) = |
456 |
foldl_map (abs_over_free all_loc_consts) (defaults, loc_defs); |
|
457 |
val loc_defs_thms = |
|
458 |
map (apsnd (prep_hyps (map #1 loc_consts) syntax_sign)) loc_defs_terms; |
|
459 |
val (defaults, loc_thms_terms) = |
|
460 |
foldl_map (abs_over_free all_loc_consts) (defaults, loc_rules); |
|
461 |
val loc_thms = map (apsnd (prep_hyps (map #1 loc_consts) syntax_sign)) |
|
462 |
(loc_thms_terms) |
|
5782
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
463 |
@ loc_defs_thms; |
5244 | 464 |
|
465 |
||
6268 | 466 |
(* error messages *) |
467 |
||
468 |
fun locale_error msg = error (msg ^ "\nFor locale " ^ quote name); |
|
5244 | 469 |
|
470 |
val err_dup_locale = |
|
471 |
if is_none (get_locale thy name) then [] |
|
472 |
else ["Duplicate definition of locale " ^ quote name]; |
|
473 |
||
6268 | 474 |
(* check if definientes are locale constants |
475 |
(in the same locale, so no redefining!) *) |
|
5749 | 476 |
val err_def_head = |
477 |
let fun peal_appl t = |
|
478 |
case t of |
|
479 |
t1 $ t2 => peal_appl t1 |
|
480 |
| Free(t) => t |
|
6268 | 481 |
| _ => locale_error ("Bad form of LHS in locale definition"); |
5749 | 482 |
fun lhs (_, Const ("==" , _) $ d1 $ d2) = peal_appl d1 |
6268 | 483 |
| lhs _ = locale_error ("Definitions must use the == relation"); |
5749 | 484 |
val defs = map lhs loc_defs; |
485 |
val check = defs subset loc_consts |
|
486 |
in if check then [] |
|
487 |
else ["defined item not declared fixed in locale " ^ quote name] |
|
488 |
end; |
|
489 |
||
5782
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
490 |
(* check that variables on rhs of definitions are either fixed or on lhs *) |
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
491 |
val err_var_rhs = |
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
492 |
let fun compare_var_sides (t, (_, Const ("==", _) $ d1 $ d2)) = |
6268 | 493 |
let val varl1 = difflist d1 all_loc_consts; |
494 |
val varl2 = difflist d2 all_loc_consts |
|
495 |
in t andalso (varl2 subset varl1) |
|
496 |
end |
|
497 |
| compare_var_sides (_,_) = |
|
498 |
locale_error ("Definitions must use the == relation") |
|
5782
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
499 |
val check = foldl compare_var_sides (true, loc_defs) |
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
500 |
in if check then [] |
6268 | 501 |
else ["nonfixed variable on right hand side of a locale definition in locale " ^ quote name] |
5782
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
502 |
end; |
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
503 |
|
7559f116cb10
locales now implicitly quantify over free variables
paulson
parents:
5749
diff
changeset
|
504 |
val errs = err_dup_locale @ err_def_head @ err_var_rhs; |
5244 | 505 |
in |
506 |
if null errs then () |
|
507 |
else error (cat_lines errs); |
|
508 |
||
509 |
syntax_thy |
|
5749 | 510 |
|> put_locale (name, |
6022 | 511 |
make_locale ancestor loc_consts nosyn loc_thms_terms |
512 |
loc_defs_terms loc_thms defaults) |
|
5244 | 513 |
end; |
514 |
||
515 |
||
516 |
val add_locale = gen_add_locale read_typ read_axm; |
|
517 |
val add_locale_i = gen_add_locale cert_typ cert_axm; |
|
518 |
||
519 |
(** print functions **) |
|
520 |
(* idea: substitute all locale contants (Free's) that are syntactical by their |
|
521 |
"real" constant representation (i.e. \\<^locale>constname). |
|
522 |
- function const_ssubst does this substitution |
|
523 |
- function Locale.pretty_term: |
|
524 |
if locale is open then do this substitution & then call Sign.pretty_term |
|
525 |
else call Sign.pretty_term |
|
526 |
*) |
|
527 |
(* substitutes all Free variables s in t by Const's s *) |
|
528 |
fun const_ssubst t s = |
|
529 |
case t of |
|
530 |
Free(v,T) => if v = s then Const("\\<^locale>" ^ s,T) else Free(v,T) |
|
531 |
| Const(c,T) => Const(c,T) |
|
532 |
| Var(v,T) => Var(v,T) |
|
533 |
| Bound x => Bound x |
|
534 |
| Abs(a,T,u) => Abs(a,T, const_ssubst u s) |
|
535 |
| t1 $ t2 => const_ssubst t1 s $ const_ssubst t2 s; |
|
536 |
||
537 |
(* FIXME: improve: can be expressed with foldl *) |
|
538 |
fun const_ssubst_list [] t = t |
|
539 |
| const_ssubst_list (s :: l) t = const_ssubst_list l (const_ssubst t s); |
|
540 |
||
541 |
(* Locale.pretty_term *) |
|
542 |
fun pretty_term sign = |
|
543 |
if (is_open_loc_sg sign) then |
|
544 |
let val locale_list = map snd(get_scope_sg sign); |
|
545 |
val nosyn = flat (map #nosyn locale_list); |
|
546 |
val str_list = (collect_consts sign) \\ nosyn |
|
547 |
in Sign.pretty_term sign o (const_ssubst_list str_list) |
|
548 |
end |
|
549 |
else Sign.pretty_term sign; |
|
550 |
||
551 |
||
552 |
||
553 |
(** print_goals **) |
|
554 |
||
555 |
(*print thm A1,...,An/B in "goal style" -- premises as numbered subgoals*) |
|
556 |
||
557 |
local |
|
558 |
||
559 |
(* utils *) |
|
560 |
||
561 |
fun ins_entry (x, y) [] = [(x, [y])] |
|
562 |
| ins_entry (x, y) ((pair as (x', ys')) :: pairs) = |
|
563 |
if x = x' then (x', y ins ys') :: pairs |
|
564 |
else pair :: ins_entry (x, y) pairs; |
|
565 |
||
566 |
fun add_consts (Const (c, T), env) = ins_entry (T, (c, T)) env |
|
567 |
| add_consts (t $ u, env) = add_consts (u, add_consts (t, env)) |
|
568 |
| add_consts (Abs (_, _, t), env) = add_consts (t, env) |
|
569 |
| add_consts (_, env) = env; |
|
570 |
||
571 |
fun add_vars (Free (x, T), env) = ins_entry (T, (x, ~1)) env |
|
572 |
| add_vars (Var (xi, T), env) = ins_entry (T, xi) env |
|
573 |
| add_vars (Abs (_, _, t), env) = add_vars (t, env) |
|
574 |
| add_vars (t $ u, env) = add_vars (u, add_vars (t, env)) |
|
575 |
| add_vars (_, env) = env; |
|
576 |
||
577 |
fun add_varsT (Type (_, Ts), env) = foldr add_varsT (Ts, env) |
|
578 |
| add_varsT (TFree (x, S), env) = ins_entry (S, (x, ~1)) env |
|
579 |
| add_varsT (TVar (xi, S), env) = ins_entry (S, xi) env; |
|
580 |
||
581 |
fun sort_idxs vs = map (apsnd (sort (prod_ord string_ord int_ord))) vs; |
|
582 |
fun sort_cnsts cs = map (apsnd (sort_wrt fst)) cs; |
|
583 |
||
584 |
||
585 |
(* prepare atoms *) |
|
586 |
||
587 |
fun consts_of t = sort_cnsts (add_consts (t, [])); |
|
588 |
fun vars_of t = sort_idxs (add_vars (t, [])); |
|
589 |
fun varsT_of t = rev (sort_idxs (it_term_types add_varsT (t, []))); |
|
590 |
||
591 |
in |
|
592 |
||
8459 | 593 |
fun pretty_goals_marker begin_goal maxgoals state = |
5244 | 594 |
let |
595 |
val {sign, ...} = rep_thm state; |
|
596 |
||
597 |
val prt_term = pretty_term sign; |
|
598 |
val prt_typ = Sign.pretty_typ sign; |
|
599 |
val prt_sort = Sign.pretty_sort sign; |
|
600 |
||
601 |
fun prt_atoms prt prtT (X, xs) = Pretty.block |
|
602 |
[Pretty.block (Pretty.commas (map prt xs)), Pretty.str " ::", |
|
603 |
Pretty.brk 1, prtT X]; |
|
604 |
||
605 |
fun prt_var (x, ~1) = prt_term (Syntax.free x) |
|
606 |
| prt_var xi = prt_term (Syntax.var xi); |
|
607 |
||
608 |
fun prt_varT (x, ~1) = prt_typ (TFree (x, [])) |
|
609 |
| prt_varT xi = prt_typ (TVar (xi, [])); |
|
610 |
||
611 |
val prt_consts = prt_atoms (prt_term o Const) prt_typ; |
|
612 |
val prt_vars = prt_atoms prt_var prt_typ; |
|
613 |
val prt_varsT = prt_atoms prt_varT prt_sort; |
|
614 |
||
615 |
||
8459 | 616 |
fun pretty_list _ _ [] = [] |
617 |
| pretty_list name prt lst = [Pretty.big_list name (map prt lst)]; |
|
5244 | 618 |
|
8459 | 619 |
fun pretty_subgoal (n, A) = |
620 |
Pretty.blk (0, [Pretty.str (begin_goal ^ " " ^ string_of_int n ^ ". "), prt_term A]); |
|
621 |
fun pretty_subgoals As = map pretty_subgoal (1 upto length As ~~ As); |
|
5244 | 622 |
|
8459 | 623 |
val pretty_ffpairs = pretty_list "Flex-flex pairs:" (prt_term o Logic.mk_flexpair); |
5244 | 624 |
|
8459 | 625 |
val pretty_consts = pretty_list "Constants:" prt_consts o consts_of; |
626 |
val pretty_vars = pretty_list "Variables:" prt_vars o vars_of; |
|
627 |
val pretty_varsT = pretty_list "Type variables:" prt_varsT o varsT_of; |
|
5244 | 628 |
|
629 |
||
630 |
val {prop, ...} = rep_thm state; |
|
631 |
val (tpairs, As, B) = Logic.strip_horn prop; |
|
632 |
val ngoals = length As; |
|
633 |
||
8459 | 634 |
fun pretty_gs (types, sorts) = |
635 |
[prt_term B] @ |
|
636 |
(if ngoals = 0 then [Pretty.str "No subgoals!"] |
|
637 |
else if ngoals > maxgoals then |
|
638 |
pretty_subgoals (take (maxgoals, As)) @ |
|
639 |
[Pretty.str ("A total of " ^ string_of_int ngoals ^ " subgoals...")] |
|
640 |
else pretty_subgoals As) @ |
|
641 |
pretty_ffpairs tpairs @ |
|
8908 | 642 |
(if ! show_consts then pretty_consts prop else []) @ |
8459 | 643 |
(if types then pretty_vars prop else []) @ |
644 |
(if sorts then pretty_varsT prop else []); |
|
5244 | 645 |
in |
646 |
setmp show_no_free_types true |
|
647 |
(setmp show_types (! show_types orelse ! show_sorts) |
|
8459 | 648 |
(setmp show_sorts false pretty_gs)) |
5244 | 649 |
(! show_types orelse ! show_sorts, ! show_sorts) |
650 |
end; |
|
651 |
||
8459 | 652 |
val print_goals_marker = (Pretty.writeln o Pretty.chunks) ooo pretty_goals_marker; |
653 |
||
654 |
val pretty_goals = pretty_goals_marker ""; |
|
5729
5d66410cceae
support current_goals_markers ref variable for print_current_goals;
wenzelm
parents:
5244
diff
changeset
|
655 |
val print_goals = print_goals_marker ""; |
5d66410cceae
support current_goals_markers ref variable for print_current_goals;
wenzelm
parents:
5244
diff
changeset
|
656 |
|
5244 | 657 |
end; |
658 |
||
659 |
||
660 |
||
661 |
(** locale theory setup **) |
|
662 |
||
663 |
val setup = |
|
664 |
[LocalesData.init]; |
|
665 |
||
666 |
end; |
|
667 |
||
668 |
structure BasicLocale: BASIC_LOCALE = Locale; |
|
669 |
open BasicLocale; |