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" |
|
|
15 |
"aval (Plus a\<^isub>1 a\<^isub>2) s =
|
|
16 |
(case (aval a\<^isub>1 s, aval a\<^isub>2 s) of
|
|
17 |
(Some i\<^isub>1,Some i\<^isub>2) \<Rightarrow> Some(i\<^isub>1+i\<^isub>2) | _ \<Rightarrow> None)"
|
|
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))" |
|
|
23 |
"bval (And b\<^isub>1 b\<^isub>2) s = (case (bval b\<^isub>1 s, bval b\<^isub>2 s) of
|
|
24 |
(Some bv\<^isub>1, Some bv\<^isub>2) \<Rightarrow> Some(bv\<^isub>1 & bv\<^isub>2) | _ \<Rightarrow> None)" |
|
|
25 |
"bval (Less a\<^isub>1 a\<^isub>2) s = (case (aval a\<^isub>1 s, aval a\<^isub>2 s) of
|
|
26 |
(Some i\<^isub>1, Some i\<^isub>2) \<Rightarrow> Some(i\<^isub>1 < i\<^isub>2) | _ \<Rightarrow> None)"
|
|
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
|