| author | wenzelm | 
| Tue, 07 Jun 2022 19:15:08 +0200 | |
| changeset 75536 | 7cdeed5dc96d | 
| parent 69913 | ca515cf61651 | 
| permissions | -rw-r--r-- | 
| 41959 | 1 | (* Title: HOL/Statespace/StateSpaceLocale.thy | 
| 25171 | 2 | Author: Norbert Schirmer, TU Muenchen | 
| 3 | *) | |
| 4 | ||
| 63167 | 5 | section \<open>Setup for State Space Locales \label{sec:StateSpaceLocale}\<close>
 | 
| 25171 | 6 | |
| 7 | theory StateSpaceLocale imports StateFun | |
| 69913 | 8 | keywords "statespace" :: thy_defn | 
| 25171 | 9 | begin | 
| 10 | ||
| 69605 | 11 | ML_file \<open>state_space.ML\<close> | 
| 12 | ML_file \<open>state_fun.ML\<close> | |
| 25171 | 13 | |
| 63167 | 14 | text \<open>For every type that is to be stored in a state space, an | 
| 25171 | 15 | instance of this locale is imported in order convert the abstract and | 
| 63167 | 16 | concrete values.\<close> | 
| 25171 | 17 | |
| 18 | ||
| 29247 
95d3a82857e5
adapted statespace module to new locales;
 Norbert Schirmer <norbert.schirmer@web.de> parents: 
29235diff
changeset | 19 | locale project_inject = | 
| 25171 | 20 | fixes project :: "'value \<Rightarrow> 'a" | 
| 38838 | 21 | and inject :: "'a \<Rightarrow> 'value" | 
| 25171 | 22 | assumes project_inject_cancel [statefun_simp]: "project (inject x) = x" | 
| 45358 | 23 | begin | 
| 25171 | 24 | |
| 45358 | 25 | lemma ex_project [statefun_simp]: "\<exists>v. project v = x" | 
| 26 | proof | |
| 27 | show "project (inject x) = x" | |
| 28 | by (rule project_inject_cancel) | |
| 29 | qed | |
| 25171 | 30 | |
| 45358 | 31 | lemma project_inject_comp_id [statefun_simp]: "project \<circ> inject = id" | 
| 25171 | 32 | by (rule ext) (simp add: project_inject_cancel) | 
| 33 | ||
| 45358 | 34 | lemma project_inject_comp_cancel[statefun_simp]: "f \<circ> project \<circ> inject = f" | 
| 25171 | 35 | by (rule ext) (simp add: project_inject_cancel) | 
| 36 | ||
| 45358 | 37 | end | 
| 25171 | 38 | |
| 62390 | 39 | end |