| author | blanchet | 
| Sun, 16 Feb 2014 21:33:28 +0100 | |
| changeset 55525 | 70b7e91fa1f9 | 
| parent 51737 | 718866dda2fa | 
| child 55972 | 51b342baecda | 
| 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 =  | 
| 45362 | 6  | 
sig  | 
7  | 
val distinct_compsN : string  | 
|
8  | 
val getN : string  | 
|
9  | 
val putN : string  | 
|
10  | 
val injectN : string  | 
|
11  | 
val namespaceN : string  | 
|
12  | 
val projectN : string  | 
|
13  | 
val valuetypesN : string  | 
|
| 25408 | 14  | 
|
| 45362 | 15  | 
val namespace_definition :  | 
16  | 
bstring ->  | 
|
17  | 
typ ->  | 
|
| 46925 | 18  | 
(xstring, string) Expression.expr * (binding * string option * mixfix) list ->  | 
| 45362 | 19  | 
string list -> string list -> theory -> theory  | 
| 25408 | 20  | 
|
| 45362 | 21  | 
val define_statespace :  | 
22  | 
string list ->  | 
|
23  | 
string ->  | 
|
| 
49754
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
24  | 
((string * bool) * (string list * bstring * (string * string) list)) list ->  | 
| 45362 | 25  | 
(string * string) list -> theory -> theory  | 
26  | 
val define_statespace_i :  | 
|
27  | 
string option ->  | 
|
28  | 
string list ->  | 
|
29  | 
string ->  | 
|
| 
49754
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
30  | 
((string * bool) * (typ list * bstring * (string * string) list)) list ->  | 
| 45362 | 31  | 
(string * typ) list -> theory -> theory  | 
| 25408 | 32  | 
|
| 45362 | 33  | 
val statespace_decl :  | 
34  | 
((string list * bstring) *  | 
|
| 
49754
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
35  | 
(((string * bool) * (string list * xstring * (bstring * bstring) list)) list *  | 
| 45362 | 36  | 
(bstring * string) list)) parser  | 
| 25408 | 37  | 
|
38  | 
||
| 45362 | 39  | 
val neq_x_y : Proof.context -> term -> term -> thm option  | 
40  | 
val distinctNameSolver : Simplifier.solver  | 
|
41  | 
val distinctTree_tac : Proof.context -> int -> tactic  | 
|
42  | 
val distinct_simproc : Simplifier.simproc  | 
|
| 25408 | 43  | 
|
44  | 
||
| 48741 | 45  | 
val get_comp : Context.generic -> string -> (typ * string) option  | 
| 45362 | 46  | 
val get_silent : Context.generic -> bool  | 
47  | 
val set_silent : bool -> Context.generic -> Context.generic  | 
|
| 25408 | 48  | 
|
| 45362 | 49  | 
val gen_lookup_tr : Proof.context -> term -> string -> term  | 
50  | 
val lookup_swap_tr : Proof.context -> term list -> term  | 
|
51  | 
val lookup_tr : Proof.context -> term list -> term  | 
|
52  | 
val lookup_tr' : Proof.context -> term list -> term  | 
|
| 25408 | 53  | 
|
| 45362 | 54  | 
val gen_update_tr :  | 
55  | 
bool -> Proof.context -> string -> term -> term -> term  | 
|
56  | 
val update_tr : Proof.context -> term list -> term  | 
|
57  | 
val update_tr' : Proof.context -> term list -> term  | 
|
58  | 
end;  | 
|
| 25408 | 59  | 
|
| 
29247
 
95d3a82857e5
adapted statespace module to new locales;
 
Norbert Schirmer <norbert.schirmer@web.de> 
parents: 
29064 
diff
changeset
 | 
60  | 
structure StateSpace : STATE_SPACE =  | 
| 25171 | 61  | 
struct  | 
62  | 
||
| 45362 | 63  | 
(* Names *)  | 
| 25171 | 64  | 
|
65  | 
val distinct_compsN = "distinct_names"  | 
|
66  | 
val namespaceN = "_namespace"  | 
|
67  | 
val valuetypesN = "_valuetypes"  | 
|
68  | 
val projectN = "project"  | 
|
69  | 
val injectN = "inject"  | 
|
70  | 
val getN = "get"  | 
|
71  | 
val putN = "put"  | 
|
72  | 
val project_injectL = "StateSpaceLocale.project_inject";  | 
|
73  | 
||
74  | 
||
75  | 
(* Library *)  | 
|
76  | 
||
77  | 
fun fold1 f xs = fold f (tl xs) (hd xs)  | 
|
78  | 
fun fold1' f [] x = x  | 
|
79  | 
| fold1' f xs _ = fold1 f xs  | 
|
| 26478 | 80  | 
|
81  | 
fun sublist_idx eq xs ys =  | 
|
| 25171 | 82  | 
let  | 
| 26478 | 83  | 
fun sublist n xs ys =  | 
| 25171 | 84  | 
if is_prefix eq xs ys then SOME n  | 
85  | 
else (case ys of [] => NONE  | 
|
86  | 
| (y::ys') => sublist (n+1) xs ys')  | 
|
87  | 
in sublist 0 xs ys end;  | 
|
88  | 
||
89  | 
fun is_sublist eq xs ys = is_some (sublist_idx eq xs ys);  | 
|
90  | 
||
91  | 
fun sorted_subset eq [] ys = true  | 
|
92  | 
| sorted_subset eq (x::xs) [] = false  | 
|
93  | 
| sorted_subset eq (x::xs) (y::ys) = if eq (x,y) then sorted_subset eq xs ys  | 
|
94  | 
else sorted_subset eq (x::xs) ys;  | 
|
95  | 
||
96  | 
||
97  | 
||
98  | 
type namespace_info =  | 
|
99  | 
 {declinfo: (typ*string) Termtab.table, (* type, name of statespace *)
 | 
|
100  | 
distinctthm: thm Symtab.table,  | 
|
101  | 
silent: bool  | 
|
102  | 
};  | 
|
| 26478 | 103  | 
|
| 33519 | 104  | 
structure NameSpaceData = Generic_Data  | 
105  | 
(  | 
|
| 26478 | 106  | 
type T = namespace_info;  | 
| 25171 | 107  | 
  val empty = {declinfo = Termtab.empty, distinctthm = Symtab.empty, silent = false};
 | 
108  | 
val extend = I;  | 
|
| 33519 | 109  | 
fun merge  | 
110  | 
    ({declinfo=declinfo1, distinctthm=distinctthm1, silent=silent1},
 | 
|
111  | 
      {declinfo=declinfo2, distinctthm=distinctthm2, silent=silent2}) : T =
 | 
|
112  | 
    {declinfo = Termtab.merge (K true) (declinfo1, declinfo2),
 | 
|
113  | 
distinctthm = Symtab.merge (K true) (distinctthm1, distinctthm2),  | 
|
| 
41472
 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 
wenzelm 
parents: 
41270 
diff
changeset
 | 
114  | 
silent = silent1 andalso silent2 (* FIXME odd merge *)}  | 
| 33519 | 115  | 
);  | 
| 25171 | 116  | 
|
117  | 
fun make_namespace_data declinfo distinctthm silent =  | 
|
118  | 
     {declinfo=declinfo,distinctthm=distinctthm,silent=silent};
 | 
|
119  | 
||
120  | 
||
121  | 
fun delete_declinfo n ctxt =  | 
|
122  | 
  let val {declinfo,distinctthm,silent} = NameSpaceData.get ctxt;
 | 
|
| 26478 | 123  | 
in NameSpaceData.put  | 
| 25171 | 124  | 
(make_namespace_data (Termtab.delete_safe n declinfo) distinctthm silent) ctxt  | 
125  | 
end;  | 
|
126  | 
||
127  | 
||
128  | 
fun update_declinfo (n,v) ctxt =  | 
|
129  | 
  let val {declinfo,distinctthm,silent} = NameSpaceData.get ctxt;
 | 
|
| 26478 | 130  | 
in NameSpaceData.put  | 
| 25171 | 131  | 
(make_namespace_data (Termtab.update (n,v) declinfo) distinctthm silent) ctxt  | 
132  | 
end;  | 
|
133  | 
||
134  | 
fun set_silent silent ctxt =  | 
|
135  | 
  let val {declinfo,distinctthm,...} = NameSpaceData.get ctxt;
 | 
|
| 26478 | 136  | 
in NameSpaceData.put  | 
| 25171 | 137  | 
(make_namespace_data declinfo distinctthm silent) ctxt  | 
138  | 
end;  | 
|
139  | 
||
140  | 
val get_silent = #silent o NameSpaceData.get;  | 
|
| 26478 | 141  | 
|
| 46925 | 142  | 
fun expression_no_pos (expr, fixes) : Expression.expression =  | 
143  | 
(map (fn (name, inst) => ((name, Position.none), inst)) expr, fixes);  | 
|
144  | 
||
| 25171 | 145  | 
fun prove_interpretation_in ctxt_tac (name, expr) thy =  | 
146  | 
thy  | 
|
| 
51737
 
718866dda2fa
target-sensitive user-level commands interpretation and sublocale
 
haftmann 
parents: 
51717 
diff
changeset
 | 
147  | 
|> Expression.sublocale_global_cmd I (name, Position.none) (expression_no_pos expr) []  | 
| 26478 | 148  | 
|> Proof.global_terminal_proof  | 
| 
49889
 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 
wenzelm 
parents: 
49866 
diff
changeset
 | 
149  | 
((Method.Basic (fn ctxt => SIMPLE_METHOD (ctxt_tac ctxt)), Position.no_range), NONE)  | 
| 42361 | 150  | 
|> Proof_Context.theory_of  | 
| 25171 | 151  | 
|
| 
29247
 
95d3a82857e5
adapted statespace module to new locales;
 
Norbert Schirmer <norbert.schirmer@web.de> 
parents: 
29064 
diff
changeset
 | 
152  | 
fun add_locale name expr elems thy =  | 
| 45362 | 153  | 
thy  | 
| 
41585
 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 
wenzelm 
parents: 
41472 
diff
changeset
 | 
154  | 
|> Expression.add_locale I (Binding.name name) (Binding.name name) expr elems  | 
| 29362 | 155  | 
|> snd  | 
| 33671 | 156  | 
|> Local_Theory.exit;  | 
| 
29247
 
95d3a82857e5
adapted statespace module to new locales;
 
Norbert Schirmer <norbert.schirmer@web.de> 
parents: 
29064 
diff
changeset
 | 
157  | 
|
| 
 
95d3a82857e5
adapted statespace module to new locales;
 
Norbert Schirmer <norbert.schirmer@web.de> 
parents: 
29064 
diff
changeset
 | 
158  | 
fun add_locale_cmd name expr elems thy =  | 
| 45362 | 159  | 
thy  | 
| 46925 | 160  | 
|> Expression.add_locale_cmd I (Binding.name name) Binding.empty (expression_no_pos expr) elems  | 
| 29362 | 161  | 
|> snd  | 
| 33671 | 162  | 
|> Local_Theory.exit;  | 
| 
29247
 
95d3a82857e5
adapted statespace module to new locales;
 
Norbert Schirmer <norbert.schirmer@web.de> 
parents: 
29064 
diff
changeset
 | 
163  | 
|
| 25171 | 164  | 
type statespace_info =  | 
165  | 
 {args: (string * sort) list, (* type arguments *)
 | 
|
166  | 
parents: (typ list * string * string option list) list,  | 
|
167  | 
(* type instantiation, state-space name, component renamings *)  | 
|
168  | 
components: (string * typ) list,  | 
|
169  | 
types: typ list (* range types of state space *)  | 
|
170  | 
};  | 
|
171  | 
||
| 33519 | 172  | 
structure StateSpaceData = Generic_Data  | 
173  | 
(  | 
|
| 25171 | 174  | 
type T = statespace_info Symtab.table;  | 
175  | 
val empty = Symtab.empty;  | 
|
176  | 
val extend = I;  | 
|
| 33519 | 177  | 
fun merge data : T = Symtab.merge (K true) data;  | 
178  | 
);  | 
|
| 25171 | 179  | 
|
180  | 
fun add_statespace name args parents components types ctxt =  | 
|
| 26478 | 181  | 
StateSpaceData.put  | 
| 25171 | 182  | 
      (Symtab.update_new (name, {args=args,parents=parents,
 | 
183  | 
components=components,types=types}) (StateSpaceData.get ctxt))  | 
|
| 26478 | 184  | 
ctxt;  | 
| 25171 | 185  | 
|
186  | 
fun get_statespace ctxt name =  | 
|
187  | 
Symtab.lookup (StateSpaceData.get ctxt) name;  | 
|
188  | 
||
189  | 
||
190  | 
fun mk_free ctxt name =  | 
|
| 26478 | 191  | 
if Variable.is_fixed ctxt name orelse Variable.is_declared ctxt name  | 
| 25171 | 192  | 
then  | 
| 
42488
 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 
wenzelm 
parents: 
42368 
diff
changeset
 | 
193  | 
let val n' = Variable.intern_fixed ctxt name  | 
| 
 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 
wenzelm 
parents: 
42368 
diff
changeset
 | 
194  | 
in SOME (Free (n', Proof_Context.infer_type ctxt (n', dummyT))) end  | 
| 25171 | 195  | 
else NONE  | 
| 26478 | 196  | 
|
197  | 
||
| 25171 | 198  | 
fun get_dist_thm ctxt name = Symtab.lookup (#distinctthm (NameSpaceData.get ctxt)) name;  | 
| 26478 | 199  | 
fun get_comp ctxt name =  | 
200  | 
Option.mapPartial  | 
|
201  | 
(Termtab.lookup (#declinfo (NameSpaceData.get ctxt)))  | 
|
| 25171 | 202  | 
(mk_free (Context.proof_of ctxt) name);  | 
203  | 
||
204  | 
||
205  | 
(*** Tactics ***)  | 
|
206  | 
||
207  | 
fun neq_x_y ctxt x y =  | 
|
208  | 
(let  | 
|
| 26478 | 209  | 
val dist_thm = the (get_dist_thm (Context.Proof ctxt) (#1 (dest_Free x)));  | 
| 25171 | 210  | 
val ctree = cprop_of dist_thm |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2;  | 
211  | 
val tree = term_of ctree;  | 
|
212  | 
val x_path = the (DistinctTreeProver.find_tree x tree);  | 
|
213  | 
val y_path = the (DistinctTreeProver.find_tree y tree);  | 
|
214  | 
val thm = DistinctTreeProver.distinctTreeProver dist_thm x_path y_path;  | 
|
| 26478 | 215  | 
in SOME thm  | 
| 45361 | 216  | 
end handle Option.Option => NONE)  | 
| 25171 | 217  | 
|
| 
42368
 
3b8498ac2314
proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
218  | 
fun distinctTree_tac ctxt = SUBGOAL (fn (goal, i) =>  | 
| 
 
3b8498ac2314
proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
219  | 
(case goal of  | 
| 
 
3b8498ac2314
proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
220  | 
    Const (@{const_name Trueprop}, _) $
 | 
| 
 
3b8498ac2314
proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
221  | 
      (Const (@{const_name Not}, _) $
 | 
| 
 
3b8498ac2314
proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
222  | 
        (Const (@{const_name HOL.eq}, _) $ (x as Free _) $ (y as Free _))) =>
 | 
| 
 
3b8498ac2314
proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
223  | 
(case neq_x_y ctxt x y of  | 
| 
 
3b8498ac2314
proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
224  | 
SOME neq => rtac neq i  | 
| 
 
3b8498ac2314
proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
225  | 
| NONE => no_tac)  | 
| 
 
3b8498ac2314
proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
226  | 
| _ => no_tac));  | 
| 25171 | 227  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
50214 
diff
changeset
 | 
228  | 
val distinctNameSolver = mk_solver "distinctNameSolver" distinctTree_tac;  | 
| 25171 | 229  | 
|
230  | 
val distinct_simproc =  | 
|
| 
38715
 
6513ea67d95d
renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
 
wenzelm 
parents: 
38558 
diff
changeset
 | 
231  | 
  Simplifier.simproc_global @{theory HOL} "StateSpace.distinct_simproc" ["x = y"]
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
50214 
diff
changeset
 | 
232  | 
    (fn ctxt => (fn (Const (@{const_name HOL.eq},_)$(x as Free _)$(y as Free _)) =>
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
50214 
diff
changeset
 | 
233  | 
Option.map (fn neq => DistinctTreeProver.neq_to_eq_False OF [neq]) (neq_x_y ctxt x y)  | 
| 45362 | 234  | 
| _ => NONE));  | 
| 25171 | 235  | 
|
| 26478 | 236  | 
fun interprete_parent name dist_thm_name parent_expr thy =  | 
| 25171 | 237  | 
let  | 
| 
42368
 
3b8498ac2314
proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
238  | 
fun solve_tac ctxt = CSUBGOAL (fn (goal, i) =>  | 
| 25171 | 239  | 
let  | 
| 42361 | 240  | 
val distinct_thm = Proof_Context.get_thm ctxt dist_thm_name;  | 
| 25171 | 241  | 
val rule = DistinctTreeProver.distinct_implProver distinct_thm goal;  | 
| 
42368
 
3b8498ac2314
proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
242  | 
in rtac rule i end);  | 
| 26478 | 243  | 
|
| 
42368
 
3b8498ac2314
proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
244  | 
fun tac ctxt =  | 
| 
 
3b8498ac2314
proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
 
wenzelm 
parents: 
42364 
diff
changeset
 | 
245  | 
Locale.intro_locales_tac true ctxt [] THEN ALLGOALS (solve_tac ctxt);  | 
| 26478 | 246  | 
|
| 45362 | 247  | 
in  | 
248  | 
thy |> prove_interpretation_in tac (name, parent_expr)  | 
|
| 26478 | 249  | 
end;  | 
| 25171 | 250  | 
|
| 26478 | 251  | 
fun namespace_definition name nameT parent_expr parent_comps new_comps thy =  | 
| 25171 | 252  | 
let  | 
253  | 
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
 | 
254  | 
val vars = (map (fn n => (Binding.name n, NONE, NoSyn)) all_comps);  | 
| 28965 | 255  | 
val full_name = Sign.full_bname thy name;  | 
| 25171 | 256  | 
val dist_thm_name = distinct_compsN;  | 
257  | 
||
| 
29247
 
95d3a82857e5
adapted statespace module to new locales;
 
Norbert Schirmer <norbert.schirmer@web.de> 
parents: 
29064 
diff
changeset
 | 
258  | 
val dist_thm_full_name = dist_thm_name;  | 
| 26478 | 259  | 
fun comps_of_thm thm = prop_of thm  | 
| 25171 | 260  | 
|> (fn (_$(_$t)) => DistinctTreeProver.dest_tree t) |> map (fst o dest_Free);  | 
| 26478 | 261  | 
|
| 
38835
 
088502dfd89f
eliminated broken Output.no_warnings_CRITICAL -- context visibility does the job;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
262  | 
fun type_attr phi = Thm.declaration_attribute (fn thm => fn context =>  | 
| 
 
088502dfd89f
eliminated broken Output.no_warnings_CRITICAL -- context visibility does the job;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
263  | 
(case context of  | 
| 
 
088502dfd89f
eliminated broken Output.no_warnings_CRITICAL -- context visibility does the job;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
264  | 
Context.Theory _ => context  | 
| 
 
088502dfd89f
eliminated broken Output.no_warnings_CRITICAL -- context visibility does the job;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
265  | 
| Context.Proof ctxt =>  | 
| 26478 | 266  | 
let  | 
| 
38835
 
088502dfd89f
eliminated broken Output.no_warnings_CRITICAL -- context visibility does the job;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
267  | 
          val {declinfo,distinctthm=tt,silent} = NameSpaceData.get context;
 | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32952 
diff
changeset
 | 
268  | 
val all_names = comps_of_thm thm;  | 
| 25171 | 269  | 
fun upd name tt =  | 
| 
38835
 
088502dfd89f
eliminated broken Output.no_warnings_CRITICAL -- context visibility does the job;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
270  | 
(case Symtab.lookup tt name of  | 
| 25171 | 271  | 
SOME dthm => if sorted_subset (op =) (comps_of_thm dthm) all_names  | 
272  | 
then Symtab.update (name,thm) tt else tt  | 
|
| 
38835
 
088502dfd89f
eliminated broken Output.no_warnings_CRITICAL -- context visibility does the job;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
273  | 
| NONE => Symtab.update (name,thm) tt)  | 
| 25171 | 274  | 
|
275  | 
val tt' = tt |> fold upd all_names;  | 
|
| 
38835
 
088502dfd89f
eliminated broken Output.no_warnings_CRITICAL -- context visibility does the job;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
276  | 
val context' =  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
50214 
diff
changeset
 | 
277  | 
Context_Position.set_visible false ctxt  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
50214 
diff
changeset
 | 
278  | 
addsimprocs [distinct_simproc]  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
50214 
diff
changeset
 | 
279  | 
|> Context_Position.restore_visible ctxt  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
50214 
diff
changeset
 | 
280  | 
|> Context.Proof  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
50214 
diff
changeset
 | 
281  | 
              |> NameSpaceData.put {declinfo=declinfo,distinctthm=tt',silent=silent};
 | 
| 
38835
 
088502dfd89f
eliminated broken Output.no_warnings_CRITICAL -- context visibility does the job;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
282  | 
in context' end));  | 
| 26478 | 283  | 
|
| 25171 | 284  | 
val attr = Attrib.internal type_attr;  | 
285  | 
||
| 45362 | 286  | 
val assume =  | 
287  | 
((Binding.name dist_thm_name, [attr]),  | 
|
288  | 
[(HOLogic.Trueprop $  | 
|
289  | 
          (Const (@{const_name all_distinct}, Type (@{type_name tree}, [nameT]) --> HOLogic.boolT) $
 | 
|
290  | 
DistinctTreeProver.mk_tree (fn n => Free (n, nameT)) nameT  | 
|
291  | 
(sort fast_string_ord all_comps)), [])]);  | 
|
292  | 
in  | 
|
293  | 
thy  | 
|
294  | 
|> add_locale name ([], vars) [Element.Assumes [assume]]  | 
|
295  | 
|> Proof_Context.theory_of  | 
|
296  | 
|> interprete_parent name dist_thm_full_name parent_expr  | 
|
| 25171 | 297  | 
end;  | 
298  | 
||
| 32651 | 299  | 
fun encode_dot x = if x= #"." then #"_" else x;  | 
| 25171 | 300  | 
|
301  | 
fun encode_type (TFree (s, _)) = s  | 
|
302  | 
| encode_type (TVar ((s,i),_)) = "?" ^ s ^ string_of_int i  | 
|
| 26478 | 303  | 
| encode_type (Type (n,Ts)) =  | 
| 25171 | 304  | 
let  | 
305  | 
val Ts' = fold1' (fn x => fn y => x ^ "_" ^ y) (map encode_type Ts) "";  | 
|
| 32651 | 306  | 
val n' = String.map encode_dot n;  | 
| 25171 | 307  | 
in if Ts'="" then n' else Ts' ^ "_" ^ n' end;  | 
308  | 
||
309  | 
fun project_name T = projectN ^"_"^encode_type T;  | 
|
310  | 
fun inject_name T = injectN ^"_"^encode_type T;  | 
|
311  | 
||
312  | 
fun project_free T pT V = Free (project_name T, V --> pT);  | 
|
313  | 
fun inject_free T pT V = Free (inject_name T, pT --> V);  | 
|
314  | 
||
315  | 
fun get_name n = getN ^ "_" ^ n;  | 
|
316  | 
fun put_name n = putN ^ "_" ^ n;  | 
|
317  | 
fun get_const n T nT V = Free (get_name n, (nT --> V) --> T);  | 
|
318  | 
fun put_const n T nT V = Free (put_name n, T --> (nT --> V) --> (nT --> V));  | 
|
319  | 
||
| 45362 | 320  | 
fun lookup_const T nT V =  | 
321  | 
  Const (@{const_name StateFun.lookup}, (V --> T) --> nT --> (nT --> V) --> T);
 | 
|
322  | 
||
| 26478 | 323  | 
fun update_const T nT V =  | 
| 45362 | 324  | 
  Const (@{const_name StateFun.update},
 | 
325  | 
(V --> T) --> (T --> V) --> nT --> (T --> T) --> (nT --> V) --> (nT --> V));  | 
|
| 25171 | 326  | 
|
| 45362 | 327  | 
fun K_const T = Const (@{const_name K_statefun}, T --> T --> T);
 | 
| 25171 | 328  | 
|
329  | 
||
330  | 
fun add_declaration name decl thy =  | 
|
331  | 
thy  | 
|
| 
41585
 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 
wenzelm 
parents: 
41472 
diff
changeset
 | 
332  | 
|> Named_Target.init I name  | 
| 
45291
 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 
wenzelm 
parents: 
44121 
diff
changeset
 | 
333  | 
  |> (fn lthy => Local_Theory.declaration {syntax = false, pervasive = false} (decl lthy) lthy)
 | 
| 33671 | 334  | 
|> Local_Theory.exit_global;  | 
| 25171 | 335  | 
|
336  | 
fun parent_components thy (Ts, pname, renaming) =  | 
|
337  | 
let  | 
|
338  | 
val ctxt = Context.Theory thy;  | 
|
339  | 
fun rename [] xs = xs  | 
|
340  | 
| rename (NONE::rs) (x::xs) = x::rename rs xs  | 
|
341  | 
| rename (SOME r::rs) ((x,T)::xs) = (r,T)::rename rs xs;  | 
|
| 45362 | 342  | 
    val {args, parents, components, ...} = the (Symtab.lookup (StateSpaceData.get ctxt) pname);
 | 
| 25171 | 343  | 
val inst = map fst args ~~ Ts;  | 
344  | 
val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst);  | 
|
| 26478 | 345  | 
val parent_comps =  | 
| 
49754
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
346  | 
maps (fn (Ts',n,rs) => parent_components thy (map subst Ts', n, rs)) parents;  | 
| 25171 | 347  | 
val all_comps = rename renaming (parent_comps @ map (apsnd subst) components);  | 
348  | 
in all_comps end;  | 
|
349  | 
||
| 43278 | 350  | 
fun take_upto i xs = List.take(xs,i) handle General.Subscript => xs;  | 
| 25171 | 351  | 
|
352  | 
fun statespace_definition state_type args name parents parent_comps components thy =  | 
|
| 26478 | 353  | 
let  | 
| 28965 | 354  | 
val full_name = Sign.full_bname thy name;  | 
| 25171 | 355  | 
val all_comps = parent_comps @ components;  | 
356  | 
||
357  | 
val components' = map (fn (n,T) => (n,(T,full_name))) components;  | 
|
| 26478 | 358  | 
val all_comps' = map (fn (n,T) => (n,(T,full_name))) all_comps;  | 
| 25171 | 359  | 
|
| 
49754
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
360  | 
fun parent_expr (prefix, (_, n, rs)) =  | 
| 
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
361  | 
(suffix namespaceN n, (prefix, Expression.Positional rs));  | 
| 
29247
 
95d3a82857e5
adapted statespace module to new locales;
 
Norbert Schirmer <norbert.schirmer@web.de> 
parents: 
29064 
diff
changeset
 | 
362  | 
val parents_expr = map parent_expr parents;  | 
| 25171 | 363  | 
fun distinct_types Ts =  | 
| 27276 | 364  | 
let val tab = fold (fn T => fn tab => Typtab.update (T,()) tab) Ts Typtab.empty;  | 
365  | 
in map fst (Typtab.dest tab) end;  | 
|
| 25171 | 366  | 
|
367  | 
val Ts = distinct_types (map snd all_comps);  | 
|
368  | 
val arg_names = map fst args;  | 
|
| 
43324
 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 
wenzelm 
parents: 
43278 
diff
changeset
 | 
369  | 
val valueN = singleton (Name.variant_list arg_names) "'value";  | 
| 
 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 
wenzelm 
parents: 
43278 
diff
changeset
 | 
370  | 
val nameN = singleton (Name.variant_list (valueN :: arg_names)) "'name";  | 
| 25171 | 371  | 
val valueT = TFree (valueN, Sign.defaultS thy);  | 
372  | 
val nameT = TFree (nameN, Sign.defaultS thy);  | 
|
373  | 
val stateT = nameT --> valueT;  | 
|
374  | 
fun projectT T = valueT --> T;  | 
|
| 26478 | 375  | 
fun injectT T = T --> valueT;  | 
| 
29247
 
95d3a82857e5
adapted statespace module to new locales;
 
Norbert Schirmer <norbert.schirmer@web.de> 
parents: 
29064 
diff
changeset
 | 
376  | 
val locinsts = map (fn T => (project_injectL,  | 
| 
49754
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
377  | 
((encode_type T,false),Expression.Positional  | 
| 45362 | 378  | 
[SOME (Free (project_name T,projectT T)),  | 
| 
29247
 
95d3a82857e5
adapted statespace module to new locales;
 
Norbert Schirmer <norbert.schirmer@web.de> 
parents: 
29064 
diff
changeset
 | 
379  | 
SOME (Free ((inject_name T,injectT T)))]))) Ts;  | 
| 32952 | 380  | 
val locs = maps (fn T => [(Binding.name (project_name T),NONE,NoSyn),  | 
381  | 
(Binding.name (inject_name T),NONE,NoSyn)]) Ts;  | 
|
382  | 
val constrains = maps (fn T => [(project_name T,projectT T),(inject_name T,injectT T)]) Ts;  | 
|
| 25171 | 383  | 
|
| 
49754
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
384  | 
fun interprete_parent_valuetypes (prefix, (Ts, pname, _)) thy =  | 
| 25171 | 385  | 
let  | 
| 26478 | 386  | 
        val {args,types,...} =
 | 
| 25171 | 387  | 
the (Symtab.lookup (StateSpaceData.get (Context.Theory thy)) pname);  | 
388  | 
val inst = map fst args ~~ Ts;  | 
|
389  | 
val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst);  | 
|
| 32952 | 390  | 
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
 | 
391  | 
|
| 45362 | 392  | 
val expr = ([(suffix valuetypesN name,  | 
| 
49754
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
393  | 
(prefix, Expression.Positional (map SOME pars)))],[]);  | 
| 29291 | 394  | 
in  | 
| 
30473
 
e0b66c11e7e4
Assumption.all_prems_of, Assumption.all_assms_of;
 
wenzelm 
parents: 
30364 
diff
changeset
 | 
395  | 
prove_interpretation_in (ALLGOALS o solve_tac o Assumption.all_prems_of)  | 
| 29291 | 396  | 
(suffix valuetypesN name, expr) thy  | 
397  | 
end;  | 
|
| 25171 | 398  | 
|
| 
49754
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
399  | 
fun interprete_parent (prefix, (_, pname, rs)) =  | 
| 25171 | 400  | 
let  | 
| 
49754
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
401  | 
val expr = ([(pname, (prefix, Expression.Positional rs))],[])  | 
| 26478 | 402  | 
in prove_interpretation_in  | 
| 29360 | 403  | 
(fn ctxt => Locale.intro_locales_tac false ctxt [])  | 
| 25171 | 404  | 
(full_name, expr) end;  | 
405  | 
||
406  | 
fun declare_declinfo updates lthy phi ctxt =  | 
|
407  | 
let  | 
|
| 26478 | 408  | 
fun upd_prf ctxt =  | 
| 25171 | 409  | 
let  | 
410  | 
fun upd (n,v) =  | 
|
411  | 
let  | 
|
| 42361 | 412  | 
val nT = Proof_Context.infer_type (Local_Theory.target_of lthy) (n, dummyT)  | 
| 26478 | 413  | 
in Context.proof_map  | 
414  | 
(update_declinfo (Morphism.term phi (Free (n,nT)),v))  | 
|
| 25171 | 415  | 
end;  | 
416  | 
in ctxt |> fold upd updates end;  | 
|
417  | 
||
418  | 
in Context.mapping I upd_prf ctxt end;  | 
|
419  | 
||
| 26478 | 420  | 
fun string_of_typ T =  | 
| 
39134
 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
421  | 
Print_Mode.setmp []  | 
| 
 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
422  | 
(Syntax.string_of_typ (Config.put show_sorts true (Syntax.init_pretty_global thy))) T;  | 
| 25171 | 423  | 
val fixestate = (case state_type of  | 
424  | 
NONE => []  | 
|
| 26478 | 425  | 
| SOME s =>  | 
426  | 
let  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32952 
diff
changeset
 | 
427  | 
val fx = Element.Fixes [(Binding.name s,SOME (string_of_typ stateT),NoSyn)];  | 
| 26478 | 428  | 
val cs = Element.Constrains  | 
| 25171 | 429  | 
(map (fn (n,T) => (n,string_of_typ T))  | 
430  | 
((map (fn (n,_) => (n,nameT)) all_comps) @  | 
|
431  | 
constrains))  | 
|
432  | 
in [fx,cs] end  | 
|
433  | 
)  | 
|
434  | 
||
| 26478 | 435  | 
|
436  | 
in thy  | 
|
437  | 
|> namespace_definition  | 
|
| 
29247
 
95d3a82857e5
adapted statespace module to new locales;
 
Norbert Schirmer <norbert.schirmer@web.de> 
parents: 
29064 
diff
changeset
 | 
438  | 
(suffix namespaceN name) nameT (parents_expr,[])  | 
| 25171 | 439  | 
(map fst parent_comps) (map fst components)  | 
| 
49754
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
440  | 
|> Context.theory_map (add_statespace full_name args (map snd parents) components [])  | 
| 
29247
 
95d3a82857e5
adapted statespace module to new locales;
 
Norbert Schirmer <norbert.schirmer@web.de> 
parents: 
29064 
diff
changeset
 | 
441  | 
|> add_locale (suffix valuetypesN name) (locinsts,locs) []  | 
| 45362 | 442  | 
|> Proof_Context.theory_of  | 
| 26478 | 443  | 
|> fold interprete_parent_valuetypes parents  | 
| 
29247
 
95d3a82857e5
adapted statespace module to new locales;
 
Norbert Schirmer <norbert.schirmer@web.de> 
parents: 
29064 
diff
changeset
 | 
444  | 
|> add_locale_cmd name  | 
| 
 
95d3a82857e5
adapted statespace module to new locales;
 
Norbert Schirmer <norbert.schirmer@web.de> 
parents: 
29064 
diff
changeset
 | 
445  | 
              ([(suffix namespaceN full_name ,(("",false),Expression.Named [])),
 | 
| 
 
95d3a82857e5
adapted statespace module to new locales;
 
Norbert Schirmer <norbert.schirmer@web.de> 
parents: 
29064 
diff
changeset
 | 
446  | 
                (suffix valuetypesN full_name,(("",false),Expression.Named  []))],[]) fixestate
 | 
| 45362 | 447  | 
|> Proof_Context.theory_of  | 
| 25171 | 448  | 
|> fold interprete_parent parents  | 
| 
38389
 
d7d915bae307
Named_Target.init: empty string represents theory target
 
haftmann 
parents: 
38350 
diff
changeset
 | 
449  | 
|> add_declaration full_name (declare_declinfo components')  | 
| 25171 | 450  | 
end;  | 
451  | 
||
452  | 
||
453  | 
(* prepare arguments *)  | 
|
454  | 
||
| 27283 | 455  | 
fun read_raw_parent ctxt raw_T =  | 
| 42361 | 456  | 
(case Proof_Context.read_typ_abbrev ctxt raw_T of  | 
| 25171 | 457  | 
Type (name, Ts) => (Ts, name)  | 
| 27283 | 458  | 
  | T => error ("Bad parent statespace specification: " ^ Syntax.string_of_typ ctxt T));
 | 
| 25171 | 459  | 
|
| 36149 | 460  | 
fun read_typ ctxt raw_T env =  | 
461  | 
let  | 
|
462  | 
val ctxt' = fold (Variable.declare_typ o TFree) env ctxt;  | 
|
463  | 
val T = Syntax.read_typ ctxt' raw_T;  | 
|
| 45741 | 464  | 
val env' = Term.add_tfreesT T env;  | 
| 36149 | 465  | 
in (T, env') end;  | 
466  | 
||
467  | 
fun cert_typ ctxt raw_T env =  | 
|
468  | 
let  | 
|
| 42361 | 469  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 36149 | 470  | 
val T = Type.no_tvars (Sign.certify_typ thy raw_T)  | 
471  | 
handle TYPE (msg, _, _) => error msg;  | 
|
| 45741 | 472  | 
val env' = Term.add_tfreesT T env;  | 
| 36149 | 473  | 
in (T, env') end;  | 
474  | 
||
| 25171 | 475  | 
fun gen_define_statespace prep_typ state_space args name parents comps thy =  | 
476  | 
let (* - args distinct  | 
|
477  | 
- only args may occur in comps and parent-instantiations  | 
|
| 26478 | 478  | 
- number of insts must match parent args  | 
| 25171 | 479  | 
- no duplicate renamings  | 
480  | 
- renaming should occur in namespace  | 
|
481  | 
*)  | 
|
| 26478 | 482  | 
    val _ = writeln ("Defining statespace " ^ quote name ^ " ...");
 | 
| 25171 | 483  | 
|
| 42361 | 484  | 
val ctxt = Proof_Context.init_global thy;  | 
| 27283 | 485  | 
|
| 
49754
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
486  | 
fun add_parent (prefix, (Ts, pname, rs)) env =  | 
| 25171 | 487  | 
let  | 
| 
49754
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
488  | 
val prefix' =  | 
| 
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
489  | 
(case prefix of  | 
| 
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
490  | 
            ("", mandatory) => (pname, mandatory)
 | 
| 
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
491  | 
| _ => prefix);  | 
| 
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
492  | 
|
| 28965 | 493  | 
val full_pname = Sign.full_bname thy pname;  | 
| 26478 | 494  | 
        val {args,components,...} =
 | 
| 25171 | 495  | 
(case get_statespace (Context.Theory thy) full_pname of  | 
496  | 
SOME r => r  | 
|
497  | 
               | NONE => error ("Undefined statespace " ^ quote pname));
 | 
|
498  | 
||
499  | 
||
| 27283 | 500  | 
val (Ts',env') = fold_map (prep_typ ctxt) Ts env  | 
| 26478 | 501  | 
handle ERROR msg => cat_error msg  | 
| 36149 | 502  | 
                    ("The error(s) above occurred in parent statespace specification "
 | 
| 25171 | 503  | 
^ quote pname);  | 
504  | 
val err_insts = if length args <> length Ts' then  | 
|
505  | 
["number of type instantiation(s) does not match arguments of parent statespace "  | 
|
506  | 
^ quote pname]  | 
|
507  | 
else [];  | 
|
| 26478 | 508  | 
|
| 25171 | 509  | 
val rnames = map fst rs  | 
510  | 
val err_dup_renamings = (case duplicates (op =) rnames of  | 
|
511  | 
[] => []  | 
|
512  | 
| dups => ["Duplicate renaming(s) for " ^ commas dups])  | 
|
513  | 
||
514  | 
val cnames = map fst components;  | 
|
| 
36692
 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 
haftmann 
parents: 
36610 
diff
changeset
 | 
515  | 
val err_rename_unknowns = (case subtract (op =) cnames rnames of  | 
| 25171 | 516  | 
[] => []  | 
517  | 
| rs => ["Unknown components " ^ commas rs]);  | 
|
| 26478 | 518  | 
|
| 25171 | 519  | 
|
520  | 
val rs' = map (AList.lookup (op =) rs o fst) components;  | 
|
521  | 
val errs =err_insts @ err_dup_renamings @ err_rename_unknowns  | 
|
| 
49754
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
522  | 
in  | 
| 
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
523  | 
if null errs then ((prefix', (Ts', full_pname, rs')), env')  | 
| 
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
524  | 
else error (cat_lines (errs @ ["in parent statespace " ^ quote pname]))  | 
| 25171 | 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 =>  | 
| 36149 | 542  | 
       cat_error msg ("The error(s) above occurred in component " ^ quote n)
 | 
| 25171 | 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  | 
|
| 
49754
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
559  | 
val raw_parent_comps = maps (parent_components thy o snd) 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  | 
| [_] => []  | 
|
| 45660 | 564  | 
| rs => ["Different types for component " ^ quote n ^ ": " ^  | 
| 
32432
 
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  | 
[] => []  | 
|
| 45660 | 573  | 
| xs => ["Components already defined in parents: " ^ commas_quote xs]);  | 
| 25171 | 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  | 
|
| 36149 | 582  | 
val define_statespace = gen_define_statespace read_typ NONE;  | 
583  | 
val define_statespace_i = gen_define_statespace cert_typ;  | 
|
| 26478 | 584  | 
|
| 25171 | 585  | 
|
| 45362 | 586  | 
|
| 25171 | 587  | 
(*** parse/print - translations ***)  | 
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 =  | 
| 45362 | 603  | 
(case get_comp (Context.Proof ctxt) n of  | 
604  | 
SOME (T, _) =>  | 
|
605  | 
      Syntax.const @{const_name StateFun.lookup} $
 | 
|
606  | 
Syntax.free (project_name T) $ Syntax.free n $ s  | 
|
607  | 
| NONE =>  | 
|
608  | 
if get_silent (Context.Proof ctxt)  | 
|
609  | 
      then Syntax.const @{const_name StateFun.lookup} $
 | 
|
610  | 
        Syntax.const @{const_syntax undefined} $ Syntax.free n $ s
 | 
|
| 45660 | 611  | 
      else raise TERM ("StateSpace.gen_lookup_tr: component " ^ quote n ^ " not defined", []));
 | 
| 25171 | 612  | 
|
| 
42052
 
34f1d2d81284
statespace syntax: strip positions -- type constraints are unexpected here;
 
wenzelm 
parents: 
41585 
diff
changeset
 | 
613  | 
fun lookup_tr ctxt [s, x] =  | 
| 42264 | 614  | 
(case Term_Position.strip_positions x of  | 
| 
42052
 
34f1d2d81284
statespace syntax: strip positions -- type constraints are unexpected here;
 
wenzelm 
parents: 
41585 
diff
changeset
 | 
615  | 
Free (n,_) => gen_lookup_tr ctxt s n  | 
| 
 
34f1d2d81284
statespace syntax: strip positions -- type constraints are unexpected here;
 
wenzelm 
parents: 
41585 
diff
changeset
 | 
616  | 
| _ => raise Match);  | 
| 
 
34f1d2d81284
statespace syntax: strip positions -- type constraints are unexpected here;
 
wenzelm 
parents: 
41585 
diff
changeset
 | 
617  | 
|
| 25171 | 618  | 
fun lookup_swap_tr ctxt [Free (n,_),s] = gen_lookup_tr ctxt s n;  | 
619  | 
||
| 45362 | 620  | 
fun lookup_tr' ctxt [_ $ Free (prj, _), n as (_ $ Free (name, _)), s] =  | 
621  | 
(case get_comp (Context.Proof ctxt) name of  | 
|
622  | 
SOME (T, _) =>  | 
|
623  | 
if prj = project_name T  | 
|
624  | 
then Syntax.const "_statespace_lookup" $ s $ n  | 
|
625  | 
else raise Match  | 
|
| 25171 | 626  | 
| NONE => raise Match)  | 
627  | 
| lookup_tr' _ ts = raise Match;  | 
|
628  | 
||
| 26478 | 629  | 
fun gen_update_tr id ctxt n v s =  | 
| 25171 | 630  | 
let  | 
| 45362 | 631  | 
    fun pname T = if id then @{const_name Fun.id} else project_name T;
 | 
632  | 
    fun iname T = if id then @{const_name Fun.id} else inject_name T;
 | 
|
633  | 
in  | 
|
634  | 
(case get_comp (Context.Proof ctxt) n of  | 
|
635  | 
SOME (T, _) =>  | 
|
636  | 
        Syntax.const @{const_name StateFun.update} $
 | 
|
637  | 
Syntax.free (pname T) $ Syntax.free (iname T) $  | 
|
638  | 
          Syntax.free n $ (Syntax.const @{const_name K_statefun} $ v) $ s
 | 
|
639  | 
| NONE =>  | 
|
640  | 
if get_silent (Context.Proof ctxt) then  | 
|
641  | 
          Syntax.const @{const_name StateFun.update} $
 | 
|
642  | 
            Syntax.const @{const_syntax undefined} $ Syntax.const @{const_syntax undefined} $
 | 
|
643  | 
            Syntax.free n $ (Syntax.const @{const_name K_statefun} $ v) $ s
 | 
|
644  | 
       else raise TERM ("StateSpace.gen_update_tr: component " ^ n ^ " not defined", []))
 | 
|
| 25171 | 645  | 
end;  | 
646  | 
||
| 
42052
 
34f1d2d81284
statespace syntax: strip positions -- type constraints are unexpected here;
 
wenzelm 
parents: 
41585 
diff
changeset
 | 
647  | 
fun update_tr ctxt [s, x, v] =  | 
| 42264 | 648  | 
(case Term_Position.strip_positions x of  | 
| 
42052
 
34f1d2d81284
statespace syntax: strip positions -- type constraints are unexpected here;
 
wenzelm 
parents: 
41585 
diff
changeset
 | 
649  | 
Free (n, _) => gen_update_tr false ctxt n v s  | 
| 
 
34f1d2d81284
statespace syntax: strip positions -- type constraints are unexpected here;
 
wenzelm 
parents: 
41585 
diff
changeset
 | 
650  | 
| _ => raise Match);  | 
| 25171 | 651  | 
|
| 45362 | 652  | 
fun update_tr' ctxt  | 
653  | 
[_ $ Free (prj, _), _ $ Free (inj, _), n as (_ $ Free (name, _)), (Const (k, _) $ v), s] =  | 
|
654  | 
      if Long_Name.base_name k = Long_Name.base_name @{const_name K_statefun} then
 | 
|
| 25171 | 655  | 
(case get_comp (Context.Proof ctxt) name of  | 
| 45362 | 656  | 
SOME (T, _) =>  | 
657  | 
if inj = inject_name T andalso prj = project_name T then  | 
|
658  | 
Syntax.const "_statespace_update" $ s $ n $ v  | 
|
659  | 
else raise Match  | 
|
660  | 
| NONE => raise Match)  | 
|
| 25171 | 661  | 
else raise Match  | 
662  | 
| update_tr' _ _ = raise Match;  | 
|
663  | 
||
664  | 
end;  | 
|
665  | 
||
666  | 
||
667  | 
(*** outer syntax *)  | 
|
668  | 
||
| 45362 | 669  | 
local  | 
670  | 
||
| 25171 | 671  | 
val type_insts =  | 
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36958 
diff
changeset
 | 
672  | 
Parse.typ >> single ||  | 
| 46949 | 673  | 
  @{keyword "("} |-- Parse.!!! (Parse.list1 Parse.typ --| @{keyword ")"})
 | 
| 25171 | 674  | 
|
| 46949 | 675  | 
val comp = Parse.name -- (@{keyword "::"} |-- Parse.!!! Parse.typ);
 | 
| 25171 | 676  | 
fun plus1_unless test scan =  | 
| 46949 | 677  | 
  scan ::: Scan.repeat (@{keyword "+"} |-- Scan.unless test (Parse.!!! scan));
 | 
| 25171 | 678  | 
|
| 46949 | 679  | 
val mapsto = @{keyword "="};
 | 
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36958 
diff
changeset
 | 
680  | 
val rename = Parse.name -- (mapsto |-- Parse.name);  | 
| 46949 | 681  | 
val renames = Scan.optional (@{keyword "["} |-- Parse.!!! (Parse.list1 rename --| @{keyword "]"})) [];
 | 
| 25171 | 682  | 
|
| 45362 | 683  | 
val parent =  | 
| 
49754
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
684  | 
Parse_Spec.locale_prefix false --  | 
| 45362 | 685  | 
((type_insts -- Parse.xname) || (Parse.xname >> pair [])) -- renames  | 
| 
49754
 
acafcac41690
more explicit namespace prefix for 'statespace' -- duplicate facts;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
686  | 
>> (fn ((prefix, (insts, name)), renames) => (prefix, (insts, name, renames)));  | 
| 25171 | 687  | 
|
| 45362 | 688  | 
in  | 
| 25171 | 689  | 
|
690  | 
val statespace_decl =  | 
|
| 45362 | 691  | 
Parse.type_args -- Parse.name --  | 
| 46949 | 692  | 
    (@{keyword "="} |--
 | 
| 45362 | 693  | 
((Scan.repeat1 comp >> pair []) ||  | 
694  | 
(plus1_unless comp parent --  | 
|
| 46949 | 695  | 
          Scan.optional (@{keyword "+"} |-- Parse.!!! (Scan.repeat1 comp)) [])));
 | 
| 45362 | 696  | 
val _ =  | 
| 50214 | 697  | 
  Outer_Syntax.command @{command_spec "statespace"} "define state-space as locale context"
 | 
| 45362 | 698  | 
(statespace_decl >> (fn ((args, name), (parents, comps)) =>  | 
699  | 
Toplevel.theory (define_statespace args name parents comps)));  | 
|
| 25171 | 700  | 
|
| 45362 | 701  | 
end;  | 
| 25171 | 702  | 
|
| 45362 | 703  | 
end;  |