src/HOL/IMP/Def_Ass_Exp.thy
author blanchet
Thu, 26 Apr 2012 20:09:38 +0200
changeset 47782 1678955ca991
parent 45212 e87feee00a4c
permissions -rw-r--r--
fixed bug in handling of new numerals -- the left-hand side of "Numeral1 = 1" should be left alone and not translated to a built-in Kodkod numeral in the specification of the "numeral" function
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
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     3
theory Def_Ass_Exp imports Vars
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     4
begin
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     5
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     6
subsection "Initialization-Sensitive Expressions Evaluation"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     7
45212
e87feee00a4c renamed name -> vname
nipkow
parents: 45200
diff changeset
     8
type_synonym state = "vname \<Rightarrow> val option"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     9
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    10
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    11
fun aval :: "aexp \<Rightarrow> state \<Rightarrow> val option" where
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    12
"aval (N i) s = Some i" |
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    13
"aval (V x) s = s x" |
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    14
"aval (Plus a\<^isub>1 a\<^isub>2) s =
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    15
  (case (aval a\<^isub>1 s, aval a\<^isub>2 s) of
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    16
     (Some i\<^isub>1,Some i\<^isub>2) \<Rightarrow> Some(i\<^isub>1+i\<^isub>2) | _ \<Rightarrow> None)"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    17
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    18
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    19
fun bval :: "bexp \<Rightarrow> state \<Rightarrow> bool option" where
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 43158
diff changeset
    20
"bval (Bc v) s = Some v" |
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    21
"bval (Not b) s = (case bval b s of None \<Rightarrow> None | Some bv \<Rightarrow> Some(\<not> bv))" |
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    22
"bval (And b\<^isub>1 b\<^isub>2) s = (case (bval b\<^isub>1 s, bval b\<^isub>2 s) of
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    23
  (Some bv\<^isub>1, Some bv\<^isub>2) \<Rightarrow> Some(bv\<^isub>1 & bv\<^isub>2) | _ \<Rightarrow> None)" |
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    24
"bval (Less a\<^isub>1 a\<^isub>2) s = (case (aval a\<^isub>1 s, aval a\<^isub>2 s) of
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    25
 (Some i\<^isub>1, Some i\<^isub>2) \<Rightarrow> Some(i\<^isub>1 < i\<^isub>2) | _ \<Rightarrow> None)"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    26
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    27
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    28
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
    29
by (induct a) auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    30
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    31
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
    32
by (induct b) (auto dest!: aval_Some)
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    33
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    34
end