author | hoelzl |
Thu Sep 02 10:14:32 2010 +0200 (2010-09-02) | |
changeset 39072 | 1030b1a166ef |
parent 38838 | 62f6ba39b3d4 |
child 41959 | b460124855b8 |
permissions | -rw-r--r-- |
schirmer@25171 | 1 |
(* Title: StateFun.thy |
schirmer@25171 | 2 |
Author: Norbert Schirmer, TU Muenchen |
schirmer@25171 | 3 |
*) |
schirmer@25171 | 4 |
|
schirmer@25171 | 5 |
header {* State Space Representation as Function \label{sec:StateFun}*} |
schirmer@25171 | 6 |
|
schirmer@25171 | 7 |
theory StateFun imports DistinctTreeProver |
schirmer@25171 | 8 |
begin |
schirmer@25171 | 9 |
|
schirmer@25171 | 10 |
|
schirmer@25171 | 11 |
text {* The state space is represented as a function from names to |
schirmer@25171 | 12 |
values. We neither fix the type of names nor the type of values. We |
schirmer@25171 | 13 |
define lookup and update functions and provide simprocs that simplify |
schirmer@25171 | 14 |
expressions containing these, similar to HOL-records. |
schirmer@25171 | 15 |
|
schirmer@25171 | 16 |
The lookup and update function get constructor/destructor functions as |
schirmer@25171 | 17 |
parameters. These are used to embed various HOL-types into the |
schirmer@25171 | 18 |
abstract value type. Conceptually the abstract value type is a sum of |
schirmer@25171 | 19 |
all types that we attempt to store in the state space. |
schirmer@25171 | 20 |
|
schirmer@25171 | 21 |
The update is actually generalized to a map function. The map supplies |
schirmer@25171 | 22 |
better compositionality, especially if you think of nested state |
schirmer@25171 | 23 |
spaces. *} |
schirmer@25171 | 24 |
|
haftmann@35416 | 25 |
definition K_statefun :: "'a \<Rightarrow> 'b \<Rightarrow> 'a" where "K_statefun c x \<equiv> c" |
schirmer@25171 | 26 |
|
schirmer@25171 | 27 |
lemma K_statefun_apply [simp]: "K_statefun c x = c" |
schirmer@25171 | 28 |
by (simp add: K_statefun_def) |
schirmer@25171 | 29 |
|
schirmer@25171 | 30 |
lemma K_statefun_comp [simp]: "(K_statefun c \<circ> f) = K_statefun c" |
schirmer@25171 | 31 |
by (rule ext) (simp add: K_statefun_apply comp_def) |
schirmer@25171 | 32 |
|
schirmer@25171 | 33 |
lemma K_statefun_cong [cong]: "K_statefun c x = K_statefun c x" |
schirmer@25171 | 34 |
by (rule refl) |
schirmer@25171 | 35 |
|
wenzelm@38838 | 36 |
definition lookup :: "('v \<Rightarrow> 'a) \<Rightarrow> 'n \<Rightarrow> ('n \<Rightarrow> 'v) \<Rightarrow> 'a" |
wenzelm@35114 | 37 |
where "lookup destr n s = destr (s n)" |
schirmer@25171 | 38 |
|
wenzelm@38838 | 39 |
definition update :: |
schirmer@25171 | 40 |
"('v \<Rightarrow> 'a1) \<Rightarrow> ('a2 \<Rightarrow> 'v) \<Rightarrow> 'n \<Rightarrow> ('a1 \<Rightarrow> 'a2) \<Rightarrow> ('n \<Rightarrow> 'v) \<Rightarrow> ('n \<Rightarrow> 'v)" |
wenzelm@35114 | 41 |
where "update destr constr n f s = s(n := constr (f (destr (s n))))" |
schirmer@25171 | 42 |
|
schirmer@25171 | 43 |
lemma lookup_update_same: |
schirmer@25171 | 44 |
"(\<And>v. destr (constr v) = v) \<Longrightarrow> lookup destr n (update destr constr n f s) = |
schirmer@25171 | 45 |
f (destr (s n))" |
schirmer@25171 | 46 |
by (simp add: lookup_def update_def) |
schirmer@25171 | 47 |
|
schirmer@25171 | 48 |
lemma lookup_update_id_same: |
schirmer@25171 | 49 |
"lookup destr n (update destr' id n (K_statefun (lookup id n s')) s) = |
schirmer@25171 | 50 |
lookup destr n s'" |
schirmer@25171 | 51 |
by (simp add: lookup_def update_def) |
schirmer@25171 | 52 |
|
schirmer@25171 | 53 |
lemma lookup_update_other: |
schirmer@25171 | 54 |
"n\<noteq>m \<Longrightarrow> lookup destr n (update destr' constr m f s) = lookup destr n s" |
schirmer@25171 | 55 |
by (simp add: lookup_def update_def) |
schirmer@25171 | 56 |
|
schirmer@25171 | 57 |
|
schirmer@25171 | 58 |
lemma id_id_cancel: "id (id x) = x" |
schirmer@25171 | 59 |
by (simp add: id_def) |
schirmer@25171 | 60 |
|
schirmer@25171 | 61 |
lemma destr_contstr_comp_id: |
schirmer@25171 | 62 |
"(\<And>v. destr (constr v) = v) \<Longrightarrow> destr \<circ> constr = id" |
schirmer@25171 | 63 |
by (rule ext) simp |
schirmer@25171 | 64 |
|
schirmer@25171 | 65 |
|
schirmer@25171 | 66 |
|
schirmer@25171 | 67 |
lemma block_conj_cong: "(P \<and> Q) = (P \<and> Q)" |
schirmer@25171 | 68 |
by simp |
schirmer@25171 | 69 |
|
schirmer@25171 | 70 |
lemma conj1_False: "(P\<equiv>False) \<Longrightarrow> (P \<and> Q) \<equiv> False" |
schirmer@25171 | 71 |
by simp |
schirmer@25171 | 72 |
|
schirmer@25171 | 73 |
lemma conj2_False: "\<lbrakk>Q\<equiv>False\<rbrakk> \<Longrightarrow> (P \<and> Q) \<equiv> False" |
schirmer@25171 | 74 |
by simp |
schirmer@25171 | 75 |
|
schirmer@25171 | 76 |
lemma conj_True: "\<lbrakk>P\<equiv>True; Q\<equiv>True\<rbrakk> \<Longrightarrow> (P \<and> Q) \<equiv> True" |
schirmer@25171 | 77 |
by simp |
schirmer@25171 | 78 |
|
schirmer@25171 | 79 |
lemma conj_cong: "\<lbrakk>P\<equiv>P'; Q\<equiv>Q'\<rbrakk> \<Longrightarrow> (P \<and> Q) \<equiv> (P' \<and> Q')" |
schirmer@25171 | 80 |
by simp |
schirmer@25171 | 81 |
|
schirmer@25171 | 82 |
|
schirmer@25171 | 83 |
lemma update_apply: "(update destr constr n f s x) = |
schirmer@25171 | 84 |
(if x=n then constr (f (destr (s n))) else s x)" |
schirmer@25171 | 85 |
by (simp add: update_def) |
schirmer@25171 | 86 |
|
schirmer@25171 | 87 |
lemma ex_id: "\<exists>x. id x = y" |
schirmer@25171 | 88 |
by (simp add: id_def) |
schirmer@25171 | 89 |
|
schirmer@25171 | 90 |
lemma swap_ex_eq: |
schirmer@25171 | 91 |
"\<exists>s. f s = x \<equiv> True \<Longrightarrow> |
schirmer@25171 | 92 |
\<exists>s. x = f s \<equiv> True" |
schirmer@25171 | 93 |
apply (rule eq_reflection) |
schirmer@25171 | 94 |
apply auto |
schirmer@25171 | 95 |
done |
schirmer@25171 | 96 |
|
schirmer@25171 | 97 |
lemmas meta_ext = eq_reflection [OF ext] |
schirmer@25171 | 98 |
|
schirmer@25171 | 99 |
(* This lemma only works if the store is welltyped: |
schirmer@25171 | 100 |
"\<exists>x. s ''n'' = (c x)" |
schirmer@25171 | 101 |
or in general when c (d x) = x, |
schirmer@25171 | 102 |
(for example: c=id and d=id) |
schirmer@25171 | 103 |
*) |
schirmer@25171 | 104 |
lemma "update d c n (K_statespace (lookup d n s)) s = s" |
schirmer@25171 | 105 |
apply (simp add: update_def lookup_def) |
schirmer@25171 | 106 |
apply (rule ext) |
schirmer@25171 | 107 |
apply simp |
schirmer@25171 | 108 |
oops |
schirmer@25171 | 109 |
|
schirmer@25171 | 110 |
end |