| author | haftmann | 
| Thu, 16 Sep 2010 16:51:34 +0200 | |
| changeset 39475 | 9cc1ba3c5706 | 
| parent 38838 | 62f6ba39b3d4 | 
| child 41959 | b460124855b8 | 
| permissions | -rw-r--r-- | 
| 25171 | 1 | (* Title: StateSpaceLocale.thy | 
| 2 | Author: Norbert Schirmer, TU Muenchen | |
| 3 | *) | |
| 4 | ||
| 5 | header {* Setup for State Space Locales \label{sec:StateSpaceLocale}*}
 | |
| 6 | ||
| 7 | theory StateSpaceLocale imports StateFun | |
| 25174 | 8 | uses "state_space.ML" "state_fun.ML" | 
| 25171 | 9 | begin | 
| 10 | ||
| 11 | setup StateFun.setup | |
| 12 | ||
| 13 | text {* For every type that is to be stored in a state space, an
 | |
| 14 | instance of this locale is imported in order convert the abstract and | |
| 15 | concrete values.*} | |
| 16 | ||
| 17 | ||
| 29247 
95d3a82857e5
adapted statespace module to new locales;
 Norbert Schirmer <norbert.schirmer@web.de> parents: 
29235diff
changeset | 18 | locale project_inject = | 
| 25171 | 19 | fixes project :: "'value \<Rightarrow> 'a" | 
| 38838 | 20 | and inject :: "'a \<Rightarrow> 'value" | 
| 25171 | 21 | assumes project_inject_cancel [statefun_simp]: "project (inject x) = x" | 
| 22 | ||
| 23 | lemma (in project_inject) | |
| 24 | ex_project [statefun_simp]: "\<exists>v. project v = x" | |
| 25 | apply (rule_tac x= "inject x" in exI) | |
| 26 | apply (simp add: project_inject_cancel) | |
| 27 | done | |
| 28 | ||
| 29 | lemma (in project_inject) | |
| 30 | project_inject_comp_id [statefun_simp]: "project \<circ> inject = id" | |
| 31 | by (rule ext) (simp add: project_inject_cancel) | |
| 32 | ||
| 33 | lemma (in project_inject) | |
| 34 | project_inject_comp_cancel[statefun_simp]: "f \<circ> project \<circ> inject = f" | |
| 35 | by (rule ext) (simp add: project_inject_cancel) | |
| 36 | ||
| 37 | ||
| 38 | ||
| 39 | end |