author | wenzelm |
Sun, 16 Feb 2014 15:38:08 +0100 | |
changeset 55512 | 75c68e05f9ea |
parent 48891 | c0eafbd55de3 |
child 58825 | 2065f49da190 |
permissions | -rw-r--r-- |
41959 | 1 |
(* Title: HOL/Statespace/StateSpaceLocale.thy |
25171 | 2 |
Author: Norbert Schirmer, TU Muenchen |
3 |
*) |
|
4 |
||
5 |
header {* Setup for State Space Locales \label{sec:StateSpaceLocale}*} |
|
6 |
||
7 |
theory StateSpaceLocale imports StateFun |
|
46950
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents:
45358
diff
changeset
|
8 |
keywords "statespace" :: thy_decl |
25171 | 9 |
begin |
10 |
||
48891 | 11 |
ML_file "state_space.ML" |
12 |
ML_file "state_fun.ML" |
|
25171 | 13 |
setup StateFun.setup |
14 |
||
15 |
text {* For every type that is to be stored in a state space, an |
|
16 |
instance of this locale is imported in order convert the abstract and |
|
17 |
concrete values.*} |
|
18 |
||
19 |
||
29247
95d3a82857e5
adapted statespace module to new locales;
Norbert Schirmer <norbert.schirmer@web.de>
parents:
29235
diff
changeset
|
20 |
locale project_inject = |
25171 | 21 |
fixes project :: "'value \<Rightarrow> 'a" |
38838 | 22 |
and inject :: "'a \<Rightarrow> 'value" |
25171 | 23 |
assumes project_inject_cancel [statefun_simp]: "project (inject x) = x" |
45358 | 24 |
begin |
25171 | 25 |
|
45358 | 26 |
lemma ex_project [statefun_simp]: "\<exists>v. project v = x" |
27 |
proof |
|
28 |
show "project (inject x) = x" |
|
29 |
by (rule project_inject_cancel) |
|
30 |
qed |
|
25171 | 31 |
|
45358 | 32 |
lemma project_inject_comp_id [statefun_simp]: "project \<circ> inject = id" |
25171 | 33 |
by (rule ext) (simp add: project_inject_cancel) |
34 |
||
45358 | 35 |
lemma project_inject_comp_cancel[statefun_simp]: "f \<circ> project \<circ> inject = f" |
25171 | 36 |
by (rule ext) (simp add: project_inject_cancel) |
37 |
||
45358 | 38 |
end |
25171 | 39 |
|
40 |
end |