src/HOL/IMP/BExp.thy
author nipkow
Wed, 19 Oct 2011 16:32:12 +0200
changeset 45200 1f1897ac7877
parent 45015 fdac1e9880eb
child 45216 a51a70687517
permissions -rw-r--r--
renamed B to Bc
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     1
theory BExp imports AExp begin
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     2
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     3
subsection "Boolean Expressions"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     4
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
     5
datatype bexp = Bc bool | Not bexp | And bexp bexp | Less aexp aexp
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     6
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     7
fun bval :: "bexp \<Rightarrow> state \<Rightarrow> bool" where
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
     8
"bval (Bc v) _ = v" |
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     9
"bval (Not b) s = (\<not> bval b s)" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    10
"bval (And b1 b2) s = (if bval b1 s then bval b2 s else False)" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    11
"bval (Less a1 a2) s = (aval a1 s < aval a2 s)"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    12
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    13
value "bval (Less (V ''x'') (Plus (N 3) (V ''y'')))
44036
d03f9f28d01d new state syntax with less conflicts
kleing
parents: 43158
diff changeset
    14
            <''x'' := 3, ''y'' := 1>"
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    15
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    16
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    17
subsection "Optimization"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    18
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    19
text{* Optimized constructors: *}
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    20
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    21
fun less :: "aexp \<Rightarrow> aexp \<Rightarrow> bexp" where
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
    22
"less (N n1) (N n2) = Bc(n1 < n2)" |
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    23
"less a1 a2 = Less a1 a2"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    24
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    25
lemma [simp]: "bval (less a1 a2) s = (aval a1 s < aval a2 s)"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44036
diff changeset
    26
apply(induction a1 a2 rule: less.induct)
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    27
apply simp_all
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    28
done
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    29
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    30
fun "and" :: "bexp \<Rightarrow> bexp \<Rightarrow> bexp" where
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
    31
"and (Bc True) b = b" |
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
    32
"and b (Bc True) = b" |
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
    33
"and (Bc False) b = Bc False" |
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
    34
"and b (Bc False) = Bc False" |
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    35
"and b1 b2 = And b1 b2"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    36
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    37
lemma bval_and[simp]: "bval (and b1 b2) s = (bval b1 s \<and> bval b2 s)"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44036
diff changeset
    38
apply(induction b1 b2 rule: and.induct)
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    39
apply simp_all
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    40
done
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    41
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    42
fun not :: "bexp \<Rightarrow> bexp" where
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
    43
"not (Bc True) = Bc False" |
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
    44
"not (Bc False) = Bc True" |
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    45
"not b = Not b"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    46
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    47
lemma bval_not[simp]: "bval (not b) s = (~bval b s)"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44036
diff changeset
    48
apply(induction b rule: not.induct)
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    49
apply simp_all
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    50
done
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    51
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    52
text{* Now the overall optimizer: *}
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    53
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    54
fun bsimp :: "bexp \<Rightarrow> bexp" where
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    55
"bsimp (Less a1 a2) = less (asimp a1) (asimp a2)" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    56
"bsimp (And b1 b2) = and (bsimp b1) (bsimp b2)" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    57
"bsimp (Not b) = not(bsimp b)" |
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
    58
"bsimp (Bc v) = Bc v"
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    59
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    60
value "bsimp (And (Less (N 0) (N 1)) b)"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    61
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    62
value "bsimp (And (Less (N 1) (N 0)) (B True))"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    63
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    64
theorem "bval (bsimp b) s = bval b s"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44036
diff changeset
    65
apply(induction b)
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    66
apply simp_all
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    67
done
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    68
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    69
end