| author | blanchet | 
| Tue, 10 Sep 2013 15:56:51 +0200 | |
| changeset 53503 | d2f21e305d0c | 
| parent 53015 | a1119cf551e8 | 
| permissions | -rw-r--r-- | 
| 43158 | 1 | (* Author: Tobias Nipkow *) | 
| 2 | ||
| 50161 | 3 | theory Def_Init_Exp | 
| 4 | imports Vars | |
| 43158 | 5 | begin | 
| 6 | ||
| 7 | subsection "Initialization-Sensitive Expressions Evaluation" | |
| 8 | ||
| 45212 | 9 | type_synonym state = "vname \<Rightarrow> val option" | 
| 43158 | 10 | |
| 11 | ||
| 12 | fun aval :: "aexp \<Rightarrow> state \<Rightarrow> val option" where | |
| 13 | "aval (N i) s = Some i" | | |
| 14 | "aval (V x) s = s x" | | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50161diff
changeset | 15 | "aval (Plus a\<^sub>1 a\<^sub>2) s = | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50161diff
changeset | 16 | (case (aval a\<^sub>1 s, aval a\<^sub>2 s) of | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50161diff
changeset | 17 | (Some i\<^sub>1,Some i\<^sub>2) \<Rightarrow> Some(i\<^sub>1+i\<^sub>2) | _ \<Rightarrow> None)" | 
| 43158 | 18 | |
| 19 | ||
| 20 | fun bval :: "bexp \<Rightarrow> state \<Rightarrow> bool option" where | |
| 45200 | 21 | "bval (Bc v) s = Some v" | | 
| 43158 | 22 | "bval (Not b) s = (case bval b s of None \<Rightarrow> None | Some bv \<Rightarrow> Some(\<not> bv))" | | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50161diff
changeset | 23 | "bval (And b\<^sub>1 b\<^sub>2) s = (case (bval b\<^sub>1 s, bval b\<^sub>2 s) of | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50161diff
changeset | 24 | (Some bv\<^sub>1, Some bv\<^sub>2) \<Rightarrow> Some(bv\<^sub>1 & bv\<^sub>2) | _ \<Rightarrow> None)" | | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50161diff
changeset | 25 | "bval (Less a\<^sub>1 a\<^sub>2) s = (case (aval a\<^sub>1 s, aval a\<^sub>2 s) of | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
50161diff
changeset | 26 | (Some i\<^sub>1, Some i\<^sub>2) \<Rightarrow> Some(i\<^sub>1 < i\<^sub>2) | _ \<Rightarrow> None)" | 
| 43158 | 27 | |
| 28 | ||
| 29 | lemma aval_Some: "vars a \<subseteq> dom s \<Longrightarrow> \<exists> i. aval a s = Some i" | |
| 30 | by (induct a) auto | |
| 31 | ||
| 32 | lemma bval_Some: "vars b \<subseteq> dom s \<Longrightarrow> \<exists> bv. bval b s = Some bv" | |
| 33 | by (induct b) (auto dest!: aval_Some) | |
| 34 | ||
| 35 | end |