src/HOL/IMP/Def_Init_Exp.thy
author wenzelm
Tue, 13 Aug 2013 16:25:47 +0200
changeset 53015 a1119cf551e8
parent 50161 4fc4237488ab
permissions -rw-r--r--
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     2
50161
4fc4237488ab tuned names
nipkow
parents: 45212
diff changeset
     3
theory Def_Init_Exp
4fc4237488ab tuned names
nipkow
parents: 45212
diff changeset
     4
imports Vars
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     5
begin
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     6
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     7
subsection "Initialization-Sensitive Expressions Evaluation"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     8
45212
e87feee00a4c renamed name -> vname
nipkow
parents: 45200
diff changeset
     9
type_synonym state = "vname \<Rightarrow> val option"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    10
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    11
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    12
fun aval :: "aexp \<Rightarrow> state \<Rightarrow> val option" where
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    13
"aval (N i) s = Some i" |
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    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: 50161
diff 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: 50161
diff 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: 50161
diff changeset
    17
     (Some i\<^sub>1,Some i\<^sub>2) \<Rightarrow> Some(i\<^sub>1+i\<^sub>2) | _ \<Rightarrow> None)"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    18
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    19
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    20
fun bval :: "bexp \<Rightarrow> state \<Rightarrow> bool option" where
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 43158
diff changeset
    21
"bval (Bc v) s = Some v" |
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    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: 50161
diff 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: 50161
diff 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: 50161
diff 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: 50161
diff changeset
    26
 (Some i\<^sub>1, Some i\<^sub>2) \<Rightarrow> Some(i\<^sub>1 < i\<^sub>2) | _ \<Rightarrow> None)"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    27
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    28
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    29
lemma aval_Some: "vars a \<subseteq> dom s \<Longrightarrow> \<exists> i. aval a s = Some i"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    30
by (induct a) auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    31
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    32
lemma bval_Some: "vars b \<subseteq> dom s \<Longrightarrow> \<exists> bv. bval b s = Some bv"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    33
by (induct b) (auto dest!: aval_Some)
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    34
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    35
end