author | haftmann |
Wed, 20 Feb 2008 14:52:38 +0100 | |
changeset 26101 | a657683e902a |
parent 25999 | f8bcd311d501 |
child 26336 | a0e2b706ce73 |
permissions | -rw-r--r-- |
25171 | 1 |
(* Title: state_space.ML |
2 |
ID: $Id$ |
|
3 |
Author: Norbert Schirmer, TU Muenchen |
|
4 |
*) |
|
5 |
||
25408 | 6 |
signature STATE_SPACE = |
7 |
sig |
|
8 |
val KN : string |
|
9 |
val distinct_compsN : string |
|
10 |
val getN : string |
|
11 |
val putN : string |
|
12 |
val injectN : string |
|
13 |
val namespaceN : string |
|
14 |
val projectN : string |
|
15 |
val valuetypesN : string |
|
16 |
||
17 |
val quiet_mode : bool ref |
|
18 |
||
19 |
val namespace_definition : |
|
20 |
bstring -> |
|
21 |
Term.typ -> |
|
22 |
Locale.expr -> |
|
23 |
string list -> string list -> Context.theory -> Context.theory |
|
24 |
||
25 |
val define_statespace : |
|
26 |
string list -> |
|
27 |
string -> |
|
28 |
(string list * bstring * (string * string) list) list -> |
|
29 |
(string * string) list -> Context.theory -> Context.theory |
|
30 |
val define_statespace_i : |
|
31 |
string option -> |
|
32 |
string list -> |
|
33 |
string -> |
|
34 |
(Term.typ list * bstring * (string * string) list) list -> |
|
35 |
(string * Term.typ) list -> Context.theory -> Context.theory |
|
36 |
||
37 |
val statespace_decl : |
|
38 |
OuterParse.token list -> |
|
39 |
((string list * bstring) * |
|
40 |
((string list * xstring * (bstring * bstring) list) list * |
|
41 |
(bstring * string) list)) * OuterParse.token list |
|
42 |
||
43 |
||
44 |
val neq_x_y : Context.proof -> Term.term -> Term.term -> Thm.thm option |
|
45 |
val distinctNameSolver : MetaSimplifier.solver |
|
46 |
val distinctTree_tac : |
|
47 |
Context.proof -> Term.term * int -> Tactical.tactic |
|
48 |
val distinct_simproc : MetaSimplifier.simproc |
|
49 |
||
50 |
||
51 |
val change_simpset : (MetaSimplifier.simpset -> MetaSimplifier.simpset) -> |
|
52 |
Context.generic -> Context.generic |
|
53 |
||
54 |
val get_comp : Context.generic -> string -> (Term.typ * string) Option.option |
|
55 |
val get_silent : Context.generic -> bool |
|
56 |
val set_silent : bool -> Context.generic -> Context.generic |
|
57 |
||
58 |
val read_typ : |
|
59 |
Context.theory -> |
|
60 |
string -> |
|
61 |
(string * Term.sort) list -> Term.typ * (string * Term.sort) list |
|
62 |
||
63 |
val gen_lookup_tr : Context.proof -> Term.term -> string -> Term.term |
|
64 |
val lookup_swap_tr : Context.proof -> Term.term list -> Term.term |
|
65 |
val lookup_tr : Context.proof -> Term.term list -> Term.term |
|
66 |
val lookup_tr' : Context.proof -> Term.term list -> Term.term |
|
67 |
||
68 |
val gen_update_tr : |
|
69 |
bool -> Context.proof -> string -> Term.term -> Term.term -> Term.term |
|
70 |
val update_tr : Context.proof -> Term.term list -> Term.term |
|
71 |
val update_tr' : Context.proof -> Term.term list -> Term.term |
|
72 |
end; |
|
73 |
||
74 |
||
75 |
structure StateSpace: STATE_SPACE = |
|
25171 | 76 |
struct |
77 |
||
78 |
(* Theorems *) |
|
79 |
||
80 |
(* Names *) |
|
81 |
val distinct_compsN = "distinct_names" |
|
82 |
val namespaceN = "_namespace" |
|
83 |
val valuetypesN = "_valuetypes" |
|
84 |
val projectN = "project" |
|
85 |
val injectN = "inject" |
|
86 |
val getN = "get" |
|
87 |
val putN = "put" |
|
88 |
val project_injectL = "StateSpaceLocale.project_inject"; |
|
89 |
val KN = "StateFun.K_statefun" |
|
90 |
||
91 |
||
92 |
(* messages *) |
|
93 |
||
94 |
val quiet_mode = ref false; |
|
95 |
fun message s = if ! quiet_mode then () else writeln s; |
|
96 |
||
97 |
(* Library *) |
|
98 |
||
99 |
fun fold1 f xs = fold f (tl xs) (hd xs) |
|
100 |
fun fold1' f [] x = x |
|
101 |
| fold1' f xs _ = fold1 f xs |
|
102 |
||
103 |
fun sublist_idx eq xs ys = |
|
104 |
let |
|
105 |
fun sublist n xs ys = |
|
106 |
if is_prefix eq xs ys then SOME n |
|
107 |
else (case ys of [] => NONE |
|
108 |
| (y::ys') => sublist (n+1) xs ys') |
|
109 |
in sublist 0 xs ys end; |
|
110 |
||
111 |
fun is_sublist eq xs ys = is_some (sublist_idx eq xs ys); |
|
112 |
||
113 |
fun sorted_subset eq [] ys = true |
|
114 |
| sorted_subset eq (x::xs) [] = false |
|
115 |
| sorted_subset eq (x::xs) (y::ys) = if eq (x,y) then sorted_subset eq xs ys |
|
116 |
else sorted_subset eq (x::xs) ys; |
|
117 |
||
118 |
||
119 |
||
120 |
type namespace_info = |
|
121 |
{declinfo: (typ*string) Termtab.table, (* type, name of statespace *) |
|
122 |
distinctthm: thm Symtab.table, |
|
123 |
silent: bool |
|
124 |
}; |
|
125 |
||
126 |
structure NameSpaceArgs = |
|
127 |
struct |
|
128 |
type T = namespace_info; |
|
129 |
val empty = {declinfo = Termtab.empty, distinctthm = Symtab.empty, silent = false}; |
|
130 |
val extend = I; |
|
131 |
fun merge pp ({declinfo=declinfo1, distinctthm=distinctthm1, silent=silent1}, |
|
132 |
{declinfo=declinfo2, distinctthm=distinctthm2, silent=silent2}) = |
|
133 |
{declinfo = Termtab.merge (K true) (declinfo1, declinfo2), |
|
134 |
distinctthm = Symtab.merge (K true) (distinctthm1, distinctthm2), |
|
135 |
silent = silent1 andalso silent2} |
|
136 |
end; |
|
137 |
||
138 |
structure NameSpaceData = GenericDataFun(NameSpaceArgs); |
|
139 |
||
140 |
fun make_namespace_data declinfo distinctthm silent = |
|
141 |
{declinfo=declinfo,distinctthm=distinctthm,silent=silent}; |
|
142 |
||
143 |
||
144 |
fun delete_declinfo n ctxt = |
|
145 |
let val {declinfo,distinctthm,silent} = NameSpaceData.get ctxt; |
|
146 |
in NameSpaceData.put |
|
147 |
(make_namespace_data (Termtab.delete_safe n declinfo) distinctthm silent) ctxt |
|
148 |
end; |
|
149 |
||
150 |
||
151 |
fun update_declinfo (n,v) ctxt = |
|
152 |
let val {declinfo,distinctthm,silent} = NameSpaceData.get ctxt; |
|
153 |
in NameSpaceData.put |
|
154 |
(make_namespace_data (Termtab.update (n,v) declinfo) distinctthm silent) ctxt |
|
155 |
end; |
|
156 |
||
157 |
fun set_silent silent ctxt = |
|
158 |
let val {declinfo,distinctthm,...} = NameSpaceData.get ctxt; |
|
159 |
in NameSpaceData.put |
|
160 |
(make_namespace_data declinfo distinctthm silent) ctxt |
|
161 |
end; |
|
162 |
||
163 |
val get_silent = #silent o NameSpaceData.get; |
|
164 |
||
165 |
fun prove_interpretation_in ctxt_tac (name, expr) thy = |
|
166 |
thy |
|
167 |
|> Locale.interpretation_in_locale I (name, expr) |
|
168 |
|> Proof.global_terminal_proof |
|
169 |
(Method.Basic (fn ctxt => Method.SIMPLE_METHOD (ctxt_tac ctxt),Position.none), NONE) |
|
170 |
|> ProofContext.theory_of |
|
171 |
||
172 |
type statespace_info = |
|
173 |
{args: (string * sort) list, (* type arguments *) |
|
174 |
parents: (typ list * string * string option list) list, |
|
175 |
(* type instantiation, state-space name, component renamings *) |
|
176 |
components: (string * typ) list, |
|
177 |
types: typ list (* range types of state space *) |
|
178 |
}; |
|
179 |
||
180 |
structure StateSpaceArgs = |
|
181 |
struct |
|
182 |
val name = "HOL/StateSpace"; |
|
183 |
type T = statespace_info Symtab.table; |
|
184 |
val empty = Symtab.empty; |
|
185 |
val extend = I; |
|
186 |
||
187 |
fun merge pp (nt1,nt2) = Symtab.merge (K true) (nt1, nt2); |
|
188 |
end; |
|
189 |
||
190 |
structure StateSpaceData = GenericDataFun(StateSpaceArgs); |
|
191 |
||
192 |
fun add_statespace name args parents components types ctxt = |
|
193 |
StateSpaceData.put |
|
194 |
(Symtab.update_new (name, {args=args,parents=parents, |
|
195 |
components=components,types=types}) (StateSpaceData.get ctxt)) |
|
196 |
ctxt; |
|
197 |
||
198 |
fun get_statespace ctxt name = |
|
199 |
Symtab.lookup (StateSpaceData.get ctxt) name; |
|
200 |
||
201 |
||
202 |
fun lookupI eq xs n = |
|
203 |
(case AList.lookup eq xs n of |
|
204 |
SOME v => v |
|
205 |
| NONE => n); |
|
206 |
||
207 |
fun mk_free ctxt name = |
|
208 |
if Variable.is_fixed ctxt name orelse Variable.is_declared ctxt name |
|
209 |
then |
|
210 |
let val n' = lookupI (op =) (Variable.fixes_of ctxt) name |
|
211 |
in SOME (Free (n',ProofContext.infer_type ctxt n')) end |
|
212 |
else NONE |
|
213 |
||
214 |
||
215 |
fun get_dist_thm ctxt name = Symtab.lookup (#distinctthm (NameSpaceData.get ctxt)) name; |
|
216 |
fun get_comp ctxt name = |
|
217 |
Option.mapPartial |
|
218 |
(Termtab.lookup (#declinfo (NameSpaceData.get ctxt))) |
|
219 |
(mk_free (Context.proof_of ctxt) name); |
|
220 |
||
221 |
||
222 |
(*** Tactics ***) |
|
223 |
||
224 |
fun neq_x_y ctxt x y = |
|
225 |
(let |
|
226 |
val dist_thm = the (get_dist_thm (Context.Proof ctxt) (#1 (dest_Free x))); |
|
227 |
val ctree = cprop_of dist_thm |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2; |
|
228 |
val tree = term_of ctree; |
|
229 |
val x_path = the (DistinctTreeProver.find_tree x tree); |
|
230 |
val y_path = the (DistinctTreeProver.find_tree y tree); |
|
231 |
val thm = DistinctTreeProver.distinctTreeProver dist_thm x_path y_path; |
|
232 |
in SOME thm |
|
233 |
end handle Option => NONE) |
|
234 |
||
235 |
fun distinctTree_tac ctxt |
|
236 |
(Const ("Trueprop",_) $ |
|
237 |
(Const ("Not", _) $ (Const ("op =", _) $ (x as Free _)$ (y as Free _))), i) = |
|
238 |
(case (neq_x_y ctxt x y) of |
|
239 |
SOME neq => rtac neq i |
|
240 |
| NONE => no_tac) |
|
241 |
| distinctTree_tac _ _ = no_tac; |
|
242 |
||
243 |
val distinctNameSolver = mk_solver' "distinctNameSolver" |
|
244 |
(fn ss => case #context (#1 (rep_ss ss)) of |
|
245 |
SOME ctxt => SUBGOAL (distinctTree_tac ctxt) |
|
246 |
| NONE => fn i => no_tac) |
|
247 |
||
248 |
val distinct_simproc = |
|
249 |
Simplifier.simproc HOL.thy "StateSpace.distinct_simproc" ["x = y"] |
|
250 |
(fn thy => fn ss => (fn (Const ("op =",_)$(x as Free _)$(y as Free _)) => |
|
251 |
(case #context (#1 (rep_ss ss)) of |
|
252 |
SOME ctxt => Option.map (fn neq => DistinctTreeProver.neq_to_eq_False OF [neq]) |
|
253 |
(neq_x_y ctxt x y) |
|
254 |
| NONE => NONE) |
|
255 |
| _ => NONE)) |
|
256 |
||
257 |
fun change_simpset f = |
|
258 |
Context.mapping |
|
259 |
(fn thy => (change_simpset_of thy f; thy)) |
|
260 |
(fn ctxt => Simplifier.put_local_simpset (f (Simplifier.get_local_simpset ctxt)) ctxt); |
|
261 |
||
262 |
fun read_typ thy s = |
|
263 |
Sign.read_typ thy s; |
|
264 |
||
265 |
local |
|
266 |
val ss = HOL_basic_ss |
|
267 |
in |
|
268 |
fun interprete_parent name dist_thm_name parent_expr thy = |
|
269 |
let |
|
270 |
||
271 |
fun solve_tac ctxt (_,i) st = |
|
272 |
let |
|
273 |
val distinct_thm = ProofContext.get_thm ctxt (Name dist_thm_name); |
|
274 |
val goal = List.nth (cprems_of st,i-1); |
|
275 |
val rule = DistinctTreeProver.distinct_implProver distinct_thm goal; |
|
276 |
in EVERY [rtac rule i] st |
|
277 |
end |
|
278 |
||
279 |
fun tac ctxt = EVERY [Locale.intro_locales_tac true ctxt [], |
|
280 |
ALLGOALS (SUBGOAL (solve_tac ctxt))] |
|
281 |
||
282 |
in thy |
|
283 |
|> prove_interpretation_in tac (name,parent_expr) |
|
284 |
end; |
|
285 |
||
286 |
end; |
|
287 |
||
288 |
fun namespace_definition name nameT parent_expr parent_comps new_comps thy = |
|
289 |
let |
|
290 |
val all_comps = parent_comps @ new_comps; |
|
291 |
val vars = Locale.Merge |
|
292 |
(map (fn n => Locale.Rename (Locale.Locale (Locale.intern thy "var") |
|
293 |
,[SOME (n,NONE)])) all_comps); |
|
294 |
||
295 |
val full_name = Sign.full_name thy name; |
|
296 |
val dist_thm_name = distinct_compsN; |
|
297 |
val dist_thm_full_name = |
|
298 |
let val prefix = fold1' (fn name => fn prfx => prfx ^ "_" ^ name) all_comps ""; |
|
299 |
in if prefix = "" then dist_thm_name else prefix ^ "." ^ dist_thm_name end; |
|
300 |
||
301 |
fun comps_of_thm thm = prop_of thm |
|
302 |
|> (fn (_$(_$t)) => DistinctTreeProver.dest_tree t) |> map (fst o dest_Free); |
|
303 |
||
304 |
fun type_attr phi (ctxt,thm) = |
|
305 |
(case ctxt of Context.Theory _ => (ctxt,thm) |
|
306 |
| _ => |
|
307 |
let |
|
308 |
val {declinfo,distinctthm=tt,silent} = (NameSpaceData.get ctxt); |
|
309 |
val all_names = comps_of_thm thm; |
|
310 |
fun upd name tt = |
|
311 |
(case (Symtab.lookup tt name) of |
|
312 |
SOME dthm => if sorted_subset (op =) (comps_of_thm dthm) all_names |
|
313 |
then Symtab.update (name,thm) tt else tt |
|
314 |
| NONE => Symtab.update (name,thm) tt) |
|
315 |
||
316 |
val tt' = tt |> fold upd all_names; |
|
317 |
val activate_simproc = |
|
318 |
Output.no_warnings |
|
319 |
(change_simpset (fn ss => ss addsimprocs [distinct_simproc])); |
|
320 |
val ctxt' = |
|
321 |
ctxt |
|
322 |
|> NameSpaceData.put {declinfo=declinfo,distinctthm=tt',silent=silent} |
|
323 |
|> activate_simproc |
|
324 |
in (ctxt',thm) |
|
325 |
end) |
|
326 |
||
327 |
val attr = Attrib.internal type_attr; |
|
328 |
||
329 |
val assumes = Element.Assumes [((dist_thm_name,[attr]), |
|
330 |
[(HOLogic.Trueprop $ |
|
331 |
(Const ("DistinctTreeProver.all_distinct", |
|
332 |
Type ("DistinctTreeProver.tree",[nameT]) --> HOLogic.boolT) $ |
|
333 |
DistinctTreeProver.mk_tree (fn n => Free (n,nameT)) nameT |
|
334 |
(sort fast_string_ord all_comps)), |
|
335 |
([]))])]; |
|
336 |
||
337 |
in thy |
|
338 |
|> Locale.add_locale_i (SOME "") name vars [assumes] |
|
339 |
||> ProofContext.theory_of |
|
340 |
||> interprete_parent name dist_thm_full_name parent_expr |
|
341 |
|> #2 |
|
342 |
end; |
|
343 |
||
344 |
structure Typetab = TableFun(type key=typ val ord = Term.typ_ord); |
|
345 |
||
346 |
fun encode_dot x = if x= #"." then #"_" else x; |
|
347 |
||
348 |
fun encode_type (TFree (s, _)) = s |
|
349 |
| encode_type (TVar ((s,i),_)) = "?" ^ s ^ string_of_int i |
|
350 |
| encode_type (Type (n,Ts)) = |
|
351 |
let |
|
352 |
val Ts' = fold1' (fn x => fn y => x ^ "_" ^ y) (map encode_type Ts) ""; |
|
353 |
val n' = String.map encode_dot n; |
|
354 |
in if Ts'="" then n' else Ts' ^ "_" ^ n' end; |
|
355 |
||
356 |
fun project_name T = projectN ^"_"^encode_type T; |
|
357 |
fun inject_name T = injectN ^"_"^encode_type T; |
|
358 |
||
359 |
fun project_free T pT V = Free (project_name T, V --> pT); |
|
360 |
fun inject_free T pT V = Free (inject_name T, pT --> V); |
|
361 |
||
362 |
fun get_name n = getN ^ "_" ^ n; |
|
363 |
fun put_name n = putN ^ "_" ^ n; |
|
364 |
fun get_const n T nT V = Free (get_name n, (nT --> V) --> T); |
|
365 |
fun put_const n T nT V = Free (put_name n, T --> (nT --> V) --> (nT --> V)); |
|
366 |
||
367 |
fun lookup_const T nT V = Const ("StateFun.lookup",(V --> T) --> nT --> (nT --> V) --> T); |
|
368 |
fun update_const T nT V = |
|
369 |
Const ("StateFun.update", |
|
370 |
(V --> T) --> (T --> V) --> nT --> (T --> T) --> (nT --> V) --> (nT --> V)); |
|
371 |
||
372 |
fun K_const T = Const ("StateFun.K_statefun",T --> T --> T); |
|
373 |
||
374 |
val no_syn = #3 (Syntax.no_syn ((),())); |
|
375 |
||
376 |
||
377 |
fun add_declaration name decl thy = |
|
378 |
thy |
|
379 |
|> TheoryTarget.init name |
|
380 |
|> (fn lthy => LocalTheory.declaration (decl lthy) lthy) |
|
381 |
|> LocalTheory.exit |
|
382 |
|> ProofContext.theory_of; |
|
383 |
||
384 |
fun parent_components thy (Ts, pname, renaming) = |
|
385 |
let |
|
386 |
val ctxt = Context.Theory thy; |
|
387 |
fun rename [] xs = xs |
|
388 |
| rename (NONE::rs) (x::xs) = x::rename rs xs |
|
389 |
| rename (SOME r::rs) ((x,T)::xs) = (r,T)::rename rs xs; |
|
390 |
val {args,parents,components,...} = |
|
391 |
the (Symtab.lookup (StateSpaceData.get ctxt) pname); |
|
392 |
val inst = map fst args ~~ Ts; |
|
393 |
val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst); |
|
394 |
val parent_comps = |
|
395 |
List.concat (map (fn (Ts',n,rs) => parent_components thy (map subst Ts',n,rs)) parents); |
|
396 |
val all_comps = rename renaming (parent_comps @ map (apsnd subst) components); |
|
397 |
in all_comps end; |
|
398 |
||
399 |
fun take_upto i xs = List.take(xs,i) handle Subscript => xs; |
|
400 |
||
401 |
fun statespace_definition state_type args name parents parent_comps components thy = |
|
402 |
let |
|
403 |
val full_name = Sign.full_name thy name; |
|
404 |
val all_comps = parent_comps @ components; |
|
405 |
||
406 |
val components' = map (fn (n,T) => (n,(T,full_name))) components; |
|
407 |
val all_comps' = map (fn (n,T) => (n,(T,full_name))) all_comps; |
|
408 |
fun parent_expr (_,n,rs) = Locale.Rename |
|
409 |
(Locale.Locale (suffix namespaceN n), |
|
410 |
map (Option.map (fn s => (s,NONE))) rs); |
|
411 |
val parents_expr = Locale.Merge (fold (fn p => fn es => parent_expr p::es) parents []); |
|
412 |
||
413 |
fun distinct_types Ts = |
|
414 |
let val tab = fold (fn T => fn tab => Typetab.update (T,()) tab) Ts Typetab.empty; |
|
415 |
in map fst (Typetab.dest tab) end; |
|
416 |
||
417 |
val Ts = distinct_types (map snd all_comps); |
|
418 |
val arg_names = map fst args; |
|
419 |
val valueN = Name.variant arg_names "'value"; |
|
420 |
val nameN = Name.variant (valueN::arg_names) "'name"; |
|
421 |
val valueT = TFree (valueN, Sign.defaultS thy); |
|
422 |
val nameT = TFree (nameN, Sign.defaultS thy); |
|
423 |
val stateT = nameT --> valueT; |
|
424 |
fun projectT T = valueT --> T; |
|
425 |
fun injectT T = T --> valueT; |
|
426 |
||
427 |
val locs = map (fn T => Locale.Rename (Locale.Locale project_injectL, |
|
428 |
[SOME (project_name T,NONE), |
|
429 |
SOME (inject_name T ,NONE)])) Ts; |
|
430 |
val constrains = List.concat |
|
431 |
(map (fn T => [(project_name T,projectT T),(inject_name T,injectT T)]) Ts); |
|
432 |
||
433 |
fun interprete_parent_valuetypes (Ts, pname, _) = |
|
434 |
let |
|
435 |
val {args,types,...} = |
|
436 |
the (Symtab.lookup (StateSpaceData.get (Context.Theory thy)) pname); |
|
437 |
val inst = map fst args ~~ Ts; |
|
438 |
val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst); |
|
439 |
val pars = List.concat (map ((fn T => [project_name T,inject_name T]) o subst) types); |
|
440 |
val expr = Locale.Rename (Locale.Locale (suffix valuetypesN name), |
|
441 |
map (fn n => SOME (n,NONE)) pars); |
|
442 |
in prove_interpretation_in (K all_tac) |
|
443 |
(suffix valuetypesN name, expr) end; |
|
444 |
||
445 |
fun interprete_parent (_, pname, rs) = |
|
446 |
let |
|
447 |
val expr = Locale.Rename (Locale.Locale pname, map (Option.map (fn n => (n,NONE))) rs) |
|
448 |
in prove_interpretation_in |
|
449 |
(fn ctxt => Locale.intro_locales_tac false ctxt []) |
|
450 |
(full_name, expr) end; |
|
451 |
||
452 |
fun declare_declinfo updates lthy phi ctxt = |
|
453 |
let |
|
454 |
fun upd_prf ctxt = |
|
455 |
let |
|
456 |
fun upd (n,v) = |
|
457 |
let |
|
458 |
val nT = ProofContext.infer_type (LocalTheory.target_of lthy) n |
|
459 |
in Context.proof_map |
|
460 |
(update_declinfo (Morphism.term phi (Free (n,nT)),v)) |
|
461 |
end; |
|
462 |
in ctxt |> fold upd updates end; |
|
463 |
||
464 |
in Context.mapping I upd_prf ctxt end; |
|
465 |
||
466 |
fun string_of_typ T = |
|
467 |
setmp show_sorts true |
|
25696
c2058af6d9bc
PrintMode.setmp (avoid direct access to print_mode ref);
wenzelm
parents:
25408
diff
changeset
|
468 |
(PrintMode.setmp [] (Syntax.string_of_typ (ProofContext.init thy))) T; |
25171 | 469 |
val fixestate = (case state_type of |
470 |
NONE => [] |
|
471 |
| SOME s => |
|
472 |
let |
|
473 |
val fx = Element.Fixes [(s,SOME (string_of_typ stateT),NoSyn)]; |
|
474 |
val cs = Element.Constrains |
|
475 |
(map (fn (n,T) => (n,string_of_typ T)) |
|
476 |
((map (fn (n,_) => (n,nameT)) all_comps) @ |
|
477 |
constrains)) |
|
478 |
in [fx,cs] end |
|
479 |
) |
|
480 |
||
481 |
||
482 |
in thy |
|
483 |
|> namespace_definition |
|
484 |
(suffix namespaceN name) nameT parents_expr |
|
485 |
(map fst parent_comps) (map fst components) |
|
486 |
|> Context.theory_map (add_statespace full_name args parents components []) |
|
487 |
|> Locale.add_locale_i (SOME "") (suffix valuetypesN name) (Locale.Merge locs) |
|
488 |
[Element.Constrains constrains] |
|
489 |
|> ProofContext.theory_of o #2 |
|
490 |
|> fold interprete_parent_valuetypes parents |
|
491 |
|> Locale.add_locale (SOME "") name |
|
492 |
(Locale.Merge [Locale.Locale (suffix namespaceN full_name) |
|
493 |
,Locale.Locale (suffix valuetypesN full_name)]) fixestate |
|
494 |
|> ProofContext.theory_of o #2 |
|
495 |
|> fold interprete_parent parents |
|
496 |
|> add_declaration (SOME full_name) (declare_declinfo components') |
|
497 |
end; |
|
498 |
||
499 |
||
500 |
(* prepare arguments *) |
|
501 |
||
502 |
fun read_raw_parent sg s = |
|
503 |
(case Sign.read_typ_abbrev sg s handle TYPE (msg, _, _) => error msg of |
|
504 |
Type (name, Ts) => (Ts, name) |
|
505 |
| _ => error ("Bad parent statespace specification: " ^ quote s)); |
|
506 |
||
507 |
fun read_typ sg s env = |
|
508 |
let |
|
509 |
fun def_sort (x, ~1) = AList.lookup (op =) env x |
|
510 |
| def_sort _ = NONE; |
|
511 |
val T = Type.no_tvars (Sign.read_def_typ (sg, def_sort) s) handle TYPE (msg, _, _) => error msg; |
|
512 |
in (T, Term.add_typ_tfrees (T, env)) end; |
|
513 |
||
514 |
fun cert_typ sg raw_T env = |
|
515 |
let val T = Type.no_tvars (Sign.certify_typ sg raw_T) handle TYPE (msg, _, _) => error msg |
|
516 |
in (T, Term.add_typ_tfrees (T, env)) end; |
|
517 |
||
518 |
||
519 |
||
520 |
||
521 |
fun gen_define_statespace prep_typ state_space args name parents comps thy = |
|
522 |
let (* - args distinct |
|
523 |
- only args may occur in comps and parent-instantiations |
|
524 |
- number of insts must match parent args |
|
525 |
- no duplicate renamings |
|
526 |
- renaming should occur in namespace |
|
527 |
*) |
|
528 |
val _ = message ("Defining statespace " ^ quote name ^ " ..."); |
|
529 |
||
530 |
fun add_parent (Ts,pname,rs) env = |
|
531 |
let |
|
532 |
val full_pname = Sign.full_name thy pname; |
|
533 |
val {args,components,...} = |
|
534 |
(case get_statespace (Context.Theory thy) full_pname of |
|
535 |
SOME r => r |
|
536 |
| NONE => error ("Undefined statespace " ^ quote pname)); |
|
537 |
||
538 |
||
539 |
val (Ts',env') = fold_map (prep_typ thy) Ts env |
|
540 |
handle ERROR msg => cat_error msg |
|
541 |
("The error(s) above occured in parent statespace specification " |
|
542 |
^ quote pname); |
|
543 |
val err_insts = if length args <> length Ts' then |
|
544 |
["number of type instantiation(s) does not match arguments of parent statespace " |
|
545 |
^ quote pname] |
|
546 |
else []; |
|
547 |
||
548 |
val rnames = map fst rs |
|
549 |
val err_dup_renamings = (case duplicates (op =) rnames of |
|
550 |
[] => [] |
|
551 |
| dups => ["Duplicate renaming(s) for " ^ commas dups]) |
|
552 |
||
553 |
val cnames = map fst components; |
|
554 |
val err_rename_unknowns = (case (filter (fn n => not (n mem cnames))) rnames of |
|
555 |
[] => [] |
|
556 |
| rs => ["Unknown components " ^ commas rs]); |
|
557 |
||
558 |
||
559 |
val rs' = map (AList.lookup (op =) rs o fst) components; |
|
560 |
val errs =err_insts @ err_dup_renamings @ err_rename_unknowns |
|
561 |
in if null errs then ((Ts',full_pname,rs'),env') |
|
562 |
else error (cat_lines (errs @ ["in parent statespace " ^ quote pname])) |
|
563 |
end; |
|
564 |
||
565 |
val (parents',env) = fold_map add_parent parents []; |
|
566 |
||
567 |
val err_dup_args = |
|
568 |
(case duplicates (op =) args of |
|
569 |
[] => [] |
|
570 |
| dups => ["Duplicate type argument(s) " ^ commas dups]); |
|
571 |
||
572 |
||
573 |
val err_dup_components = |
|
574 |
(case duplicates (op =) (map fst comps) of |
|
575 |
[] => [] |
|
576 |
| dups => ["Duplicate state-space components " ^ commas dups]); |
|
577 |
||
578 |
fun prep_comp (n,T) env = |
|
579 |
let val (T', env') = prep_typ thy T env handle ERROR msg => |
|
580 |
cat_error msg ("The error(s) above occured in component " ^ quote n) |
|
581 |
in ((n,T'), env') end; |
|
582 |
||
583 |
val (comps',env') = fold_map prep_comp comps env; |
|
584 |
||
585 |
val err_extra_frees = |
|
586 |
(case subtract (op =) args (map fst env') of |
|
587 |
[] => [] |
|
588 |
| extras => ["Extra free type variable(s) " ^ commas extras]); |
|
589 |
||
590 |
val defaultS = Sign.defaultS thy; |
|
591 |
val args' = map (fn x => (x, AList.lookup (op =) env x |> the_default defaultS)) args; |
|
592 |
||
593 |
||
594 |
fun fst_eq ((x:string,_),(y,_)) = x = y; |
|
595 |
fun snd_eq ((_,t:typ),(_,u)) = t = u; |
|
596 |
||
597 |
val raw_parent_comps = (List.concat (map (parent_components thy) parents')); |
|
598 |
fun check_type (n,T) = |
|
599 |
(case distinct (snd_eq) (filter (curry fst_eq (n,T)) raw_parent_comps) of |
|
600 |
[] => [] |
|
601 |
| [_] => [] |
|
602 |
| rs => ["Different types for component " ^ n ^": " ^ commas |
|
603 |
(map (Pretty.string_of o Display.pretty_ctyp o ctyp_of thy o snd) rs)]) |
|
604 |
||
605 |
val err_dup_types = List.concat (map check_type (duplicates fst_eq raw_parent_comps)) |
|
606 |
||
607 |
val parent_comps = distinct (fst_eq) raw_parent_comps; |
|
608 |
val all_comps = parent_comps @ comps'; |
|
609 |
val err_comp_in_parent = (case duplicates (op =) (map fst all_comps) of |
|
610 |
[] => [] |
|
611 |
| xs => ["Components already defined in parents: " ^ commas xs]); |
|
612 |
val errs = err_dup_args @ err_dup_components @ err_extra_frees @ |
|
613 |
err_dup_types @ err_comp_in_parent; |
|
614 |
||
615 |
in if null errs |
|
616 |
then thy |> statespace_definition state_space args' name parents' parent_comps comps' |
|
617 |
else error (cat_lines errs) |
|
618 |
end |
|
619 |
handle ERROR msg => cat_error msg ("Failed to define statespace " ^ quote name); |
|
620 |
||
621 |
||
622 |
val define_statespace = gen_define_statespace read_typ NONE; |
|
623 |
val define_statespace_i = gen_define_statespace cert_typ; |
|
624 |
||
625 |
||
626 |
(*** parse/print - translations ***) |
|
627 |
||
628 |
||
629 |
local |
|
630 |
fun map_get_comp f ctxt (Free (name,_)) = |
|
631 |
(case (get_comp ctxt name) of |
|
632 |
SOME (T,_) => f T T dummyT |
|
633 |
| NONE => (Syntax.free "arbitrary"(*; error "context not ready"*))) |
|
634 |
| map_get_comp _ _ _ = Syntax.free "arbitrary"; |
|
635 |
||
636 |
val get_comp_projection = map_get_comp project_free; |
|
637 |
val get_comp_injection = map_get_comp inject_free; |
|
638 |
||
639 |
fun name_of (Free (n,_)) = n; |
|
640 |
in |
|
641 |
||
642 |
fun gen_lookup_tr ctxt s n = |
|
643 |
(case get_comp (Context.Proof ctxt) n of |
|
644 |
SOME (T,_) => |
|
645 |
Syntax.const "StateFun.lookup"$Syntax.free (project_name T)$Syntax.free n$s |
|
646 |
| NONE => |
|
647 |
if get_silent (Context.Proof ctxt) |
|
648 |
then Syntax.const "StateFun.lookup"$Syntax.const "arbitrary"$Syntax.free n$s |
|
649 |
else raise TERM ("StateSpace.gen_lookup_tr: component " ^ n ^ " not defined",[])); |
|
650 |
||
651 |
fun lookup_tr ctxt [s,Free (n,_)] = gen_lookup_tr ctxt s n; |
|
652 |
fun lookup_swap_tr ctxt [Free (n,_),s] = gen_lookup_tr ctxt s n; |
|
653 |
||
654 |
fun lookup_tr' ctxt [_$Free (prj,_),n as (_$Free (name,_)),s] = |
|
655 |
( case get_comp (Context.Proof ctxt) name of |
|
656 |
SOME (T,_) => if prj=project_name T then |
|
657 |
Syntax.const "_statespace_lookup" $ s $ n |
|
658 |
else raise Match |
|
659 |
| NONE => raise Match) |
|
660 |
| lookup_tr' _ ts = raise Match; |
|
661 |
||
662 |
fun gen_update_tr id ctxt n v s = |
|
663 |
let |
|
664 |
fun pname T = if id then "Fun.id" else project_name T |
|
665 |
fun iname T = if id then "Fun.id" else inject_name T |
|
666 |
in |
|
667 |
(case get_comp (Context.Proof ctxt) n of |
|
668 |
SOME (T,_) => Syntax.const "StateFun.update"$ |
|
669 |
Syntax.free (pname T)$Syntax.free (iname T)$ |
|
670 |
Syntax.free n$(Syntax.const KN $ v)$s |
|
671 |
| NONE => |
|
672 |
if get_silent (Context.Proof ctxt) |
|
673 |
then Syntax.const "StateFun.update"$ |
|
674 |
Syntax.const "arbitrary"$Syntax.const "arbitrary"$ |
|
675 |
Syntax.free n$(Syntax.const KN $ v)$s |
|
676 |
else raise TERM ("StateSpace.gen_update_tr: component " ^ n ^ " not defined",[])) |
|
677 |
end; |
|
678 |
||
679 |
fun update_tr ctxt [s,Free (n,_),v] = gen_update_tr false ctxt n v s; |
|
680 |
||
681 |
fun update_tr' ctxt [_$Free (prj,_),_$Free (inj,_),n as (_$Free (name,_)),(Const (k,_)$v),s] = |
|
682 |
if NameSpace.base k = NameSpace.base KN then |
|
683 |
(case get_comp (Context.Proof ctxt) name of |
|
684 |
SOME (T,_) => if inj=inject_name T andalso prj=project_name T then |
|
685 |
Syntax.const "_statespace_update" $ s $ n $ v |
|
686 |
else raise Match |
|
687 |
| NONE => raise Match) |
|
688 |
else raise Match |
|
689 |
| update_tr' _ _ = raise Match; |
|
690 |
||
691 |
end; |
|
692 |
||
693 |
||
694 |
(*** outer syntax *) |
|
695 |
||
696 |
local structure P = OuterParse and K = OuterKeyword in |
|
697 |
||
698 |
val type_insts = |
|
699 |
P.typ >> single || |
|
700 |
P.$$$ "(" |-- P.!!! (P.list1 P.typ --| P.$$$ ")") |
|
701 |
||
702 |
val comp = P.name -- (P.$$$ "::" |-- P.!!! P.typ); |
|
703 |
fun plus1_unless test scan = |
|
25999 | 704 |
scan ::: Scan.repeat (P.$$$ "+" |-- Scan.unless test (P.!!! scan)); |
25171 | 705 |
|
706 |
val mapsto = P.$$$ "="; |
|
707 |
val rename = P.name -- (mapsto |-- P.name); |
|
708 |
val renames = Scan.optional (P.$$$ "[" |-- P.!!! (P.list1 rename --| P.$$$ "]")) []; |
|
709 |
||
710 |
||
711 |
val parent = ((type_insts -- P.xname) || (P.xname >> pair [])) -- renames |
|
712 |
>> (fn ((insts,name),renames) => (insts,name,renames)) |
|
713 |
||
714 |
||
715 |
val statespace_decl = |
|
716 |
P.type_args -- P.name -- |
|
717 |
(P.$$$ "=" |-- |
|
718 |
((Scan.repeat1 comp >> pair []) || |
|
719 |
(plus1_unless comp parent -- |
|
720 |
Scan.optional (P.$$$ "+" |-- P.!!! (Scan.repeat1 comp)) []))) |
|
721 |
||
722 |
val statespace_command = |
|
723 |
OuterSyntax.command "statespace" "define state space" K.thy_decl |
|
724 |
(statespace_decl >> (fn ((args,name),(parents,comps)) => |
|
725 |
Toplevel.theory (define_statespace args name parents comps))) |
|
726 |
||
727 |
end; |
|
728 |
||
729 |
end; |