src/HOL/IMP/BExp.thy
author nipkow
Mon, 20 Aug 2018 20:54:26 +0200
changeset 68776 403dd13cf6e9
parent 67406 23307fd33906
permissions -rw-r--r--
avoid session qualification because no tex is generated when used; tuned sectioning
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
68776
403dd13cf6e9 avoid session qualification because no tex is generated when used;
nipkow
parents: 67406
diff changeset
     1
subsection "Boolean Expressions"
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     2
68776
403dd13cf6e9 avoid session qualification because no tex is generated when used;
nipkow
parents: 67406
diff changeset
     3
theory BExp imports AExp begin
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     4
58310
91ea607a34d8 updated news
blanchet
parents: 58249
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
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 60170
diff changeset
     7
text_raw\<open>\snip{BExpbvaldef}{1}{2}{%\<close>
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     8
fun bval :: "bexp \<Rightarrow> state \<Rightarrow> bool" where
45216
nipkow
parents: 45200
diff changeset
     9
"bval (Bc v) s = v" |
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    10
"bval (Not b) s = (\<not> bval b s)" |
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 49920
diff changeset
    11
"bval (And b\<^sub>1 b\<^sub>2) s = (bval b\<^sub>1 s \<and> bval b\<^sub>2 s)" |
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 49920
diff changeset
    12
"bval (Less a\<^sub>1 a\<^sub>2) s = (aval a\<^sub>1 s < aval a\<^sub>2 s)"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 60170
diff changeset
    13
text_raw\<open>}%endsnip\<close>
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    14
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    15
value "bval (Less (V ''x'') (Plus (N 3) (V ''y'')))
44036
d03f9f28d01d new state syntax with less conflicts
kleing
parents: 43158
diff changeset
    16
            <''x'' := 3, ''y'' := 1>"
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    17
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    18
45255
nipkow
parents: 45216
diff changeset
    19
subsection "Constant Folding"
nipkow
parents: 45216
diff changeset
    20
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 60170
diff changeset
    21
text\<open>Optimizing constructors:\<close>
45255
nipkow
parents: 45216
diff changeset
    22
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 60170
diff changeset
    23
text_raw\<open>\snip{BExplessdef}{0}{2}{%\<close>
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    24
fun less :: "aexp \<Rightarrow> aexp \<Rightarrow> bexp" where
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 49920
diff changeset
    25
"less (N n\<^sub>1) (N n\<^sub>2) = Bc(n\<^sub>1 < n\<^sub>2)" |
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 49920
diff changeset
    26
"less a\<^sub>1 a\<^sub>2 = Less a\<^sub>1 a\<^sub>2"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 60170
diff changeset
    27
text_raw\<open>}%endsnip\<close>
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    28
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    29
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
    30
apply(induction a1 a2 rule: less.induct)
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    31
apply simp_all
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    32
done
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    33
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 60170
diff changeset
    34
text_raw\<open>\snip{BExpanddef}{2}{2}{%\<close>
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    35
fun "and" :: "bexp \<Rightarrow> bexp \<Rightarrow> bexp" where
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
    36
"and (Bc True) b = b" |
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
    37
"and b (Bc True) = b" |
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
    38
"and (Bc False) b = Bc False" |
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
    39
"and b (Bc False) = Bc False" |
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 49920
diff changeset
    40
"and b\<^sub>1 b\<^sub>2 = And b\<^sub>1 b\<^sub>2"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 60170
diff changeset
    41
text_raw\<open>}%endsnip\<close>
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    42
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    43
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
    44
apply(induction b1 b2 rule: and.induct)
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    45
apply simp_all
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    46
done
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    47
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 60170
diff changeset
    48
text_raw\<open>\snip{BExpnotdef}{2}{2}{%\<close>
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    49
fun not :: "bexp \<Rightarrow> bexp" where
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
    50
"not (Bc True) = Bc False" |
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
    51
"not (Bc False) = Bc True" |
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    52
"not b = Not b"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 60170
diff changeset
    53
text_raw\<open>}%endsnip\<close>
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    54
45255
nipkow
parents: 45216
diff changeset
    55
lemma bval_not[simp]: "bval (not b) s = (\<not> bval b s)"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44036
diff changeset
    56
apply(induction b rule: not.induct)
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    57
apply simp_all
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    58
done
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    59
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 60170
diff changeset
    60
text\<open>Now the overall optimizer:\<close>
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    61
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 60170
diff changeset
    62
text_raw\<open>\snip{BExpbsimpdef}{0}{2}{%\<close>
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    63
fun bsimp :: "bexp \<Rightarrow> bexp" where
45256
62b025f434e9 tuned order of eqns
nipkow
parents: 45255
diff changeset
    64
"bsimp (Bc v) = Bc v" |
62b025f434e9 tuned order of eqns
nipkow
parents: 45255
diff changeset
    65
"bsimp (Not b) = not(bsimp b)" |
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 49920
diff changeset
    66
"bsimp (And b\<^sub>1 b\<^sub>2) = and (bsimp b\<^sub>1) (bsimp b\<^sub>2)" |
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 49920
diff changeset
    67
"bsimp (Less a\<^sub>1 a\<^sub>2) = less (asimp a\<^sub>1) (asimp a\<^sub>2)"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 60170
diff changeset
    68
text_raw\<open>}%endsnip\<close>
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    69
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    70
value "bsimp (And (Less (N 0) (N 1)) b)"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    71
49920
nipkow
parents: 49191
diff changeset
    72
value "bsimp (And (Less (N 1) (N 0)) (Bc True))"
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    73
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    74
theorem "bval (bsimp b) s = bval b s"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44036
diff changeset
    75
apply(induction b)
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    76
apply simp_all
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    77
done
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    78
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    79
end