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